diff --git a/src/HOL/Analysis/Abstract_Topology.thy b/src/HOL/Analysis/Abstract_Topology.thy --- a/src/HOL/Analysis/Abstract_Topology.thy +++ b/src/HOL/Analysis/Abstract_Topology.thy @@ -1,4601 +1,4607 @@ (* Author: L C Paulson, University of Cambridge [ported from HOL Light] *) section \Operators involving abstract topology\ theory Abstract_Topology imports Complex_Main "HOL-Library.Set_Idioms" "HOL-Library.FuncSet" begin subsection \General notion of a topology as a value\ definition\<^marker>\tag important\ istopology :: "('a set \ bool) \ bool" where "istopology L \ (\S T. L S \ L T \ L (S \ T)) \ (\\. (\K\\. L K) \ L (\\))" typedef\<^marker>\tag important\ 'a topology = "{L::('a set) \ bool. istopology L}" morphisms "openin" "topology" unfolding istopology_def by blast lemma istopology_openin[intro]: "istopology(openin U)" using openin[of U] by blast lemma istopology_open: "istopology open" by (auto simp: istopology_def) lemma topology_inverse': "istopology U \ openin (topology U) = U" using topology_inverse[unfolded mem_Collect_eq] . lemma topology_inverse_iff: "istopology U \ openin (topology U) = U" using topology_inverse[of U] istopology_openin[of "topology U"] by auto lemma topology_eq: "T1 = T2 \ (\S. openin T1 S \ openin T2 S)" proof assume "T1 = T2" then show "\S. openin T1 S \ openin T2 S" by simp next assume H: "\S. openin T1 S \ openin T2 S" then have "openin T1 = openin T2" by (simp add: fun_eq_iff) then have "topology (openin T1) = topology (openin T2)" by simp then show "T1 = T2" unfolding openin_inverse . qed text\The "universe": the union of all sets in the topology.\ definition "topspace T = \{S. openin T S}" subsubsection \Main properties of open sets\ proposition openin_clauses: fixes U :: "'a topology" shows "openin U {}" "\S T. openin U S \ openin U T \ openin U (S\T)" "\K. (\S \ K. openin U S) \ openin U (\K)" using openin[of U] unfolding istopology_def by auto lemma openin_subset[intro]: "openin U S \ S \ topspace U" unfolding topspace_def by blast lemma openin_empty[simp]: "openin U {}" by (rule openin_clauses) lemma openin_Int[intro]: "openin U S \ openin U T \ openin U (S \ T)" by (rule openin_clauses) lemma openin_Union[intro]: "(\S. S \ K \ openin U S) \ openin U (\K)" using openin_clauses by blast lemma openin_Un[intro]: "openin U S \ openin U T \ openin U (S \ T)" using openin_Union[of "{S,T}" U] by auto lemma openin_topspace[intro, simp]: "openin U (topspace U)" by (force simp: openin_Union topspace_def) lemma openin_subopen: "openin U S \ (\x \ S. \T. openin U T \ x \ T \ T \ S)" (is "?lhs \ ?rhs") proof assume ?lhs then show ?rhs by auto next assume H: ?rhs let ?t = "\{T. openin U T \ T \ S}" have "openin U ?t" by (force simp: openin_Union) also have "?t = S" using H by auto finally show "openin U S" . qed lemma openin_INT [intro]: assumes "finite I" "\i. i \ I \ openin T (U i)" shows "openin T ((\i \ I. U i) \ topspace T)" using assms by (induct, auto simp: inf_sup_aci(2) openin_Int) lemma openin_INT2 [intro]: assumes "finite I" "I \ {}" "\i. i \ I \ openin T (U i)" shows "openin T (\i \ I. U i)" proof - have "(\i \ I. U i) \ topspace T" using \I \ {}\ openin_subset[OF assms(3)] by auto then show ?thesis using openin_INT[of _ _ U, OF assms(1) assms(3)] by (simp add: inf.absorb2 inf_commute) qed lemma openin_Inter [intro]: assumes "finite \" "\ \ {}" "\X. X \ \ \ openin T X" shows "openin T (\\)" by (metis (full_types) assms openin_INT2 image_ident) lemma openin_Int_Inter: assumes "finite \" "openin T U" "\X. X \ \ \ openin T X" shows "openin T (U \ \\)" using openin_Inter [of "insert U \"] assms by auto subsubsection \Closed sets\ definition\<^marker>\tag important\ closedin :: "'a topology \ 'a set \ bool" where "closedin U S \ S \ topspace U \ openin U (topspace U - S)" lemma closedin_subset: "closedin U S \ S \ topspace U" by (metis closedin_def) lemma closedin_empty[simp]: "closedin U {}" by (simp add: closedin_def) lemma closedin_topspace[intro, simp]: "closedin U (topspace U)" by (simp add: closedin_def) lemma closedin_Un[intro]: "closedin U S \ closedin U T \ closedin U (S \ T)" by (auto simp: Diff_Un closedin_def) lemma Diff_Inter[intro]: "A - \S = \{A - s|s. s\S}" by auto lemma closedin_Union: assumes "finite S" "\T. T \ S \ closedin U T" shows "closedin U (\S)" using assms by induction auto lemma closedin_Inter[intro]: assumes Ke: "K \ {}" and Kc: "\S. S \K \ closedin U S" shows "closedin U (\K)" using Ke Kc unfolding closedin_def Diff_Inter by auto lemma closedin_INT[intro]: assumes "A \ {}" "\x. x \ A \ closedin U (B x)" shows "closedin U (\x\A. B x)" apply (rule closedin_Inter) using assms apply auto done lemma closedin_Int[intro]: "closedin U S \ closedin U T \ closedin U (S \ T)" using closedin_Inter[of "{S,T}" U] by auto lemma openin_closedin_eq: "openin U S \ S \ topspace U \ closedin U (topspace U - S)" apply (auto simp: closedin_def Diff_Diff_Int inf_absorb2) apply (metis openin_subset subset_eq) done lemma topology_finer_closedin: "topspace X = topspace Y \ (\S. openin Y S \ openin X S) \ (\S. closedin Y S \ closedin X S)" apply safe apply (simp add: closedin_def) by (simp add: openin_closedin_eq) lemma openin_closedin: "S \ topspace U \ (openin U S \ closedin U (topspace U - S))" by (simp add: openin_closedin_eq) lemma openin_diff[intro]: assumes oS: "openin U S" and cT: "closedin U T" shows "openin U (S - T)" proof - have "S - T = S \ (topspace U - T)" using openin_subset[of U S] oS cT by (auto simp: topspace_def openin_subset) then show ?thesis using oS cT by (auto simp: closedin_def) qed lemma closedin_diff[intro]: assumes oS: "closedin U S" and cT: "openin U T" shows "closedin U (S - T)" proof - have "S - T = S \ (topspace U - T)" using closedin_subset[of U S] oS cT by (auto simp: topspace_def) then show ?thesis using oS cT by (auto simp: openin_closedin_eq) qed subsection\The discrete topology\ definition discrete_topology where "discrete_topology U \ topology (\S. S \ U)" lemma openin_discrete_topology [simp]: "openin (discrete_topology U) S \ S \ U" proof - have "istopology (\S. S \ U)" by (auto simp: istopology_def) then show ?thesis by (simp add: discrete_topology_def topology_inverse') qed lemma topspace_discrete_topology [simp]: "topspace(discrete_topology U) = U" by (meson openin_discrete_topology openin_subset openin_topspace order_refl subset_antisym) lemma closedin_discrete_topology [simp]: "closedin (discrete_topology U) S \ S \ U" by (simp add: closedin_def) lemma discrete_topology_unique: "discrete_topology U = X \ topspace X = U \ (\x \ U. openin X {x})" (is "?lhs = ?rhs") proof assume R: ?rhs then have "openin X S" if "S \ U" for S using openin_subopen subsetD that by fastforce moreover have "x \ topspace X" if "openin X S" and "x \ S" for x S using openin_subset that by blast ultimately show ?lhs using R by (auto simp: topology_eq) qed auto lemma discrete_topology_unique_alt: "discrete_topology U = X \ topspace X \ U \ (\x \ U. openin X {x})" using openin_subset by (auto simp: discrete_topology_unique) lemma subtopology_eq_discrete_topology_empty: "X = discrete_topology {} \ topspace X = {}" using discrete_topology_unique [of "{}" X] by auto lemma subtopology_eq_discrete_topology_sing: "X = discrete_topology {a} \ topspace X = {a}" by (metis discrete_topology_unique openin_topspace singletonD) subsection \Subspace topology\ definition\<^marker>\tag important\ subtopology :: "'a topology \ 'a set \ 'a topology" where "subtopology U V = topology (\T. \S. T = S \ V \ openin U S)" lemma istopology_subtopology: "istopology (\T. \S. T = S \ V \ openin U S)" (is "istopology ?L") proof - have "?L {}" by blast { fix A B assume A: "?L A" and B: "?L B" from A B obtain Sa and Sb where Sa: "openin U Sa" "A = Sa \ V" and Sb: "openin U Sb" "B = Sb \ V" by blast have "A \ B = (Sa \ Sb) \ V" "openin U (Sa \ Sb)" using Sa Sb by blast+ then have "?L (A \ B)" by blast } moreover { fix K assume K: "K \ Collect ?L" have th0: "Collect ?L = (\S. S \ V) ` Collect (openin U)" by blast from K[unfolded th0 subset_image_iff] obtain Sk where Sk: "Sk \ Collect (openin U)" "K = (\S. S \ V) ` Sk" by blast have "\K = (\Sk) \ V" using Sk by auto moreover have "openin U (\Sk)" using Sk by (auto simp: subset_eq) ultimately have "?L (\K)" by blast } ultimately show ?thesis unfolding subset_eq mem_Collect_eq istopology_def by auto qed lemma openin_subtopology: "openin (subtopology U V) S \ (\T. openin U T \ S = T \ V)" unfolding subtopology_def topology_inverse'[OF istopology_subtopology] by auto lemma openin_subtopology_Int: "openin X S \ openin (subtopology X T) (S \ T)" using openin_subtopology by auto lemma openin_subtopology_Int2: "openin X T \ openin (subtopology X S) (S \ T)" using openin_subtopology by auto lemma openin_subtopology_diff_closed: "\S \ topspace X; closedin X T\ \ openin (subtopology X S) (S - T)" unfolding closedin_def openin_subtopology by (rule_tac x="topspace X - T" in exI) auto lemma openin_relative_to: "(openin X relative_to S) = openin (subtopology X S)" by (force simp: relative_to_def openin_subtopology) lemma topspace_subtopology [simp]: "topspace (subtopology U V) = topspace U \ V" by (auto simp: topspace_def openin_subtopology) lemma topspace_subtopology_subset: "S \ topspace X \ topspace(subtopology X S) = S" by (simp add: inf.absorb_iff2) lemma closedin_subtopology: "closedin (subtopology U V) S \ (\T. closedin U T \ S = T \ V)" unfolding closedin_def topspace_subtopology by (auto simp: openin_subtopology) lemma openin_subtopology_refl: "openin (subtopology U V) V \ V \ topspace U" unfolding openin_subtopology by auto (metis IntD1 in_mono openin_subset) lemma subtopology_subtopology: "subtopology (subtopology X S) T = subtopology X (S \ T)" proof - have eq: "\T'. (\S'. T' = S' \ T \ (\T. openin X T \ S' = T \ S)) = (\Sa. T' = Sa \ (S \ T) \ openin X Sa)" by (metis inf_assoc) have "subtopology (subtopology X S) T = topology (\Ta. \Sa. Ta = Sa \ T \ openin (subtopology X S) Sa)" by (simp add: subtopology_def) also have "\ = subtopology X (S \ T)" by (simp add: openin_subtopology eq) (simp add: subtopology_def) finally show ?thesis . qed lemma openin_subtopology_alt: "openin (subtopology X U) S \ S \ (\T. U \ T) ` Collect (openin X)" by (simp add: image_iff inf_commute openin_subtopology) lemma closedin_subtopology_alt: "closedin (subtopology X U) S \ S \ (\T. U \ T) ` Collect (closedin X)" by (simp add: image_iff inf_commute closedin_subtopology) lemma subtopology_superset: assumes UV: "topspace U \ V" shows "subtopology U V = U" proof - { fix S { fix T assume T: "openin U T" "S = T \ V" from T openin_subset[OF T(1)] UV have eq: "S = T" by blast have "openin U S" unfolding eq using T by blast } moreover { assume S: "openin U S" then have "\T. openin U T \ S = T \ V" using openin_subset[OF S] UV by auto } ultimately have "(\T. openin U T \ S = T \ V) \ openin U S" by blast } then show ?thesis unfolding topology_eq openin_subtopology by blast qed lemma subtopology_topspace[simp]: "subtopology U (topspace U) = U" by (simp add: subtopology_superset) lemma subtopology_UNIV[simp]: "subtopology U UNIV = U" by (simp add: subtopology_superset) lemma subtopology_restrict: "subtopology X (topspace X \ S) = subtopology X S" by (metis subtopology_subtopology subtopology_topspace) lemma openin_subtopology_empty: "openin (subtopology U {}) S \ S = {}" by (metis Int_empty_right openin_empty openin_subtopology) lemma closedin_subtopology_empty: "closedin (subtopology U {}) S \ S = {}" by (metis Int_empty_right closedin_empty closedin_subtopology) lemma closedin_subtopology_refl [simp]: "closedin (subtopology U X) X \ X \ topspace U" by (metis closedin_def closedin_topspace inf.absorb_iff2 le_inf_iff topspace_subtopology) lemma closedin_topspace_empty: "topspace T = {} \ (closedin T S \ S = {})" by (simp add: closedin_def) lemma open_in_topspace_empty: "topspace X = {} \ openin X S \ S = {}" by (simp add: openin_closedin_eq) lemma openin_imp_subset: "openin (subtopology U S) T \ T \ S" by (metis Int_iff openin_subtopology subsetI) lemma closedin_imp_subset: "closedin (subtopology U S) T \ T \ S" by (simp add: closedin_def) lemma openin_open_subtopology: "openin X S \ openin (subtopology X S) T \ openin X T \ T \ S" by (metis inf.orderE openin_Int openin_imp_subset openin_subtopology) lemma closedin_closed_subtopology: "closedin X S \ (closedin (subtopology X S) T \ closedin X T \ T \ S)" by (metis closedin_Int closedin_imp_subset closedin_subtopology inf.orderE) lemma openin_subtopology_Un: "\openin (subtopology X T) S; openin (subtopology X U) S\ \ openin (subtopology X (T \ U)) S" by (simp add: openin_subtopology) blast lemma closedin_subtopology_Un: "\closedin (subtopology X T) S; closedin (subtopology X U) S\ \ closedin (subtopology X (T \ U)) S" by (simp add: closedin_subtopology) blast lemma openin_trans_full: "\openin (subtopology X U) S; openin X U\ \ openin X S" by (simp add: openin_open_subtopology) subsection \The canonical topology from the underlying type class\ abbreviation\<^marker>\tag important\ euclidean :: "'a::topological_space topology" where "euclidean \ topology open" abbreviation top_of_set :: "'a::topological_space set \ 'a topology" where "top_of_set \ subtopology (topology open)" lemma open_openin: "open S \ openin euclidean S" apply (rule cong[where x=S and y=S]) apply (rule topology_inverse[symmetric]) apply (auto simp: istopology_def) done declare open_openin [symmetric, simp] lemma topspace_euclidean [simp]: "topspace euclidean = UNIV" by (force simp: topspace_def) lemma topspace_euclidean_subtopology[simp]: "topspace (top_of_set S) = S" by (simp) lemma closed_closedin: "closed S \ closedin euclidean S" by (simp add: closed_def closedin_def Compl_eq_Diff_UNIV) declare closed_closedin [symmetric, simp] lemma openin_subtopology_self [simp]: "openin (top_of_set S) S" by (metis openin_topspace topspace_euclidean_subtopology) subsubsection\The most basic facts about the usual topology and metric on R\ abbreviation euclideanreal :: "real topology" where "euclideanreal \ topology open" subsection \Basic "localization" results are handy for connectedness.\ lemma openin_open: "openin (top_of_set U) S \ (\T. open T \ (S = U \ T))" by (auto simp: openin_subtopology) lemma openin_Int_open: "\openin (top_of_set U) S; open T\ \ openin (top_of_set U) (S \ T)" by (metis open_Int Int_assoc openin_open) lemma openin_open_Int[intro]: "open S \ openin (top_of_set U) (U \ S)" by (auto simp: openin_open) lemma open_openin_trans[trans]: "open S \ open T \ T \ S \ openin (top_of_set S) T" by (metis Int_absorb1 openin_open_Int) lemma open_subset: "S \ T \ open S \ openin (top_of_set T) S" by (auto simp: openin_open) lemma closedin_closed: "closedin (top_of_set U) S \ (\T. closed T \ S = U \ T)" by (simp add: closedin_subtopology Int_ac) lemma closedin_closed_Int: "closed S \ closedin (top_of_set U) (U \ S)" by (metis closedin_closed) lemma closed_subset: "S \ T \ closed S \ closedin (top_of_set T) S" by (auto simp: closedin_closed) lemma closedin_closed_subset: "\closedin (top_of_set U) V; T \ U; S = V \ T\ \ closedin (top_of_set T) S" by (metis (no_types, lifting) Int_assoc Int_commute closedin_closed inf.orderE) lemma finite_imp_closedin: fixes S :: "'a::t1_space set" shows "\finite S; S \ T\ \ closedin (top_of_set T) S" by (simp add: finite_imp_closed closed_subset) lemma closedin_singleton [simp]: fixes a :: "'a::t1_space" shows "closedin (top_of_set U) {a} \ a \ U" using closedin_subset by (force intro: closed_subset) lemma openin_euclidean_subtopology_iff: fixes S U :: "'a::metric_space set" shows "openin (top_of_set U) S \ S \ U \ (\x\S. \e>0. \x'\U. dist x' x < e \ x'\ S)" (is "?lhs \ ?rhs") proof assume ?lhs then show ?rhs unfolding openin_open open_dist by blast next define T where "T = {x. \a\S. \d>0. (\y\U. dist y a < d \ y \ S) \ dist x a < d}" have 1: "\x\T. \e>0. \y. dist y x < e \ y \ T" unfolding T_def apply clarsimp apply (rule_tac x="d - dist x a" in exI) apply (clarsimp simp add: less_diff_eq) by (metis dist_commute dist_triangle_lt) assume ?rhs then have 2: "S = U \ T" unfolding T_def by auto (metis dist_self) from 1 2 show ?lhs unfolding openin_open open_dist by fast qed lemma connected_openin: "connected S \ \(\E1 E2. openin (top_of_set S) E1 \ openin (top_of_set S) E2 \ S \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" - apply (simp add: connected_def openin_open disjoint_iff_not_equal, safe) - by (simp_all, blast+) (* SLOW *) + unfolding connected_def openin_open disjoint_iff_not_equal by blast lemma connected_openin_eq: "connected S \ \(\E1 E2. openin (top_of_set S) E1 \ openin (top_of_set S) E2 \ E1 \ E2 = S \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" apply (simp add: connected_openin, safe, blast) by (metis Int_lower1 Un_subset_iff openin_open subset_antisym) lemma connected_closedin: "connected S \ (\E1 E2. closedin (top_of_set S) E1 \ closedin (top_of_set S) E2 \ S \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (auto simp add: connected_closed closedin_closed) next assume R: ?rhs then show ?lhs proof (clarsimp simp add: connected_closed closedin_closed) fix A B assume s_sub: "S \ A \ B" "B \ S \ {}" and disj: "A \ B \ S = {}" and cl: "closed A" "closed B" have "S \ (A \ B) = S" using s_sub(1) by auto have "S - A = B \ S" using Diff_subset_conv Un_Diff_Int disj s_sub(1) by auto then have "S \ A = {}" by (metis Diff_Diff_Int Diff_disjoint Un_Diff_Int R cl closedin_closed_Int inf_commute order_refl s_sub(2)) then show "A \ S = {}" by blast qed qed lemma connected_closedin_eq: "connected S \ \(\E1 E2. closedin (top_of_set S) E1 \ closedin (top_of_set S) E2 \ E1 \ E2 = S \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" apply (simp add: connected_closedin, safe, blast) by (metis Int_lower1 Un_subset_iff closedin_closed subset_antisym) text \These "transitivity" results are handy too\ lemma openin_trans[trans]: "openin (top_of_set T) S \ openin (top_of_set U) T \ openin (top_of_set U) S" by (metis openin_Int_open openin_open) lemma openin_open_trans: "openin (top_of_set T) S \ open T \ open S" by (auto simp: openin_open intro: openin_trans) lemma closedin_trans[trans]: "closedin (top_of_set T) S \ closedin (top_of_set U) T \ closedin (top_of_set U) S" by (auto simp: closedin_closed closed_Inter Int_assoc) lemma closedin_closed_trans: "closedin (top_of_set T) S \ closed T \ closed S" by (auto simp: closedin_closed intro: closedin_trans) lemma openin_subtopology_Int_subset: "\openin (top_of_set u) (u \ S); v \ u\ \ openin (top_of_set v) (v \ S)" by (auto simp: openin_subtopology) lemma openin_open_eq: "open s \ (openin (top_of_set s) t \ open t \ t \ s)" using open_subset openin_open_trans openin_subset by fastforce subsection\Derived set (set of limit points)\ definition derived_set_of :: "'a topology \ 'a set \ 'a set" (infixl "derived'_set'_of" 80) where "X derived_set_of S \ {x \ topspace X. (\T. x \ T \ openin X T \ (\y\x. y \ S \ y \ T))}" lemma derived_set_of_restrict [simp]: "X derived_set_of (topspace X \ S) = X derived_set_of S" by (simp add: derived_set_of_def) (metis openin_subset subset_iff) lemma in_derived_set_of: "x \ X derived_set_of S \ x \ topspace X \ (\T. x \ T \ openin X T \ (\y\x. y \ S \ y \ T))" by (simp add: derived_set_of_def) lemma derived_set_of_subset_topspace: "X derived_set_of S \ topspace X" by (auto simp add: derived_set_of_def) lemma derived_set_of_subtopology: "(subtopology X U) derived_set_of S = U \ (X derived_set_of (U \ S))" by (simp add: derived_set_of_def openin_subtopology) blast lemma derived_set_of_subset_subtopology: "(subtopology X S) derived_set_of T \ S" by (simp add: derived_set_of_subtopology) lemma derived_set_of_empty [simp]: "X derived_set_of {} = {}" by (auto simp: derived_set_of_def) lemma derived_set_of_mono: "S \ T \ X derived_set_of S \ X derived_set_of T" unfolding derived_set_of_def by blast lemma derived_set_of_Un: "X derived_set_of (S \ T) = X derived_set_of S \ X derived_set_of T" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" apply (clarsimp simp: in_derived_set_of) by (metis IntE IntI openin_Int) show "?rhs \ ?lhs" by (simp add: derived_set_of_mono) qed lemma derived_set_of_Union: "finite \ \ X derived_set_of (\\) = (\S \ \. X derived_set_of S)" proof (induction \ rule: finite_induct) case (insert S \) then show ?case by (simp add: derived_set_of_Un) qed auto lemma derived_set_of_topspace: "X derived_set_of (topspace X) = {x \ topspace X. \ openin X {x}}" apply (auto simp: in_derived_set_of) by (metis Set.set_insert all_not_in_conv insertCI openin_subset subsetCE) lemma discrete_topology_unique_derived_set: "discrete_topology U = X \ topspace X = U \ X derived_set_of U = {}" by (auto simp: discrete_topology_unique derived_set_of_topspace) lemma subtopology_eq_discrete_topology_eq: "subtopology X U = discrete_topology U \ U \ topspace X \ U \ X derived_set_of U = {}" using discrete_topology_unique_derived_set [of U "subtopology X U"] by (auto simp: eq_commute derived_set_of_subtopology) lemma subtopology_eq_discrete_topology: "S \ topspace X \ S \ X derived_set_of S = {} \ subtopology X S = discrete_topology S" by (simp add: subtopology_eq_discrete_topology_eq) lemma subtopology_eq_discrete_topology_gen: "S \ X derived_set_of S = {} \ subtopology X S = discrete_topology(topspace X \ S)" by (metis Int_lower1 derived_set_of_restrict inf_assoc inf_bot_right subtopology_eq_discrete_topology_eq subtopology_subtopology subtopology_topspace) lemma subtopology_discrete_topology [simp]: "subtopology (discrete_topology U) S = discrete_topology(U \ S)" proof - have "(\T. \Sa. T = Sa \ S \ Sa \ U) = (\Sa. Sa \ U \ Sa \ S)" by force then show ?thesis by (simp add: subtopology_def) (simp add: discrete_topology_def) qed lemma openin_Int_derived_set_of_subset: "openin X S \ S \ X derived_set_of T \ X derived_set_of (S \ T)" by (auto simp: derived_set_of_def) lemma openin_Int_derived_set_of_eq: "openin X S \ S \ X derived_set_of T = S \ X derived_set_of (S \ T)" apply auto apply (meson IntI openin_Int_derived_set_of_subset subsetCE) by (meson derived_set_of_mono inf_sup_ord(2) subset_eq) subsection\ Closure with respect to a topological space\ definition closure_of :: "'a topology \ 'a set \ 'a set" (infixr "closure'_of" 80) where "X closure_of S \ {x \ topspace X. \T. x \ T \ openin X T \ (\y \ S. y \ T)}" lemma closure_of_restrict: "X closure_of S = X closure_of (topspace X \ S)" unfolding closure_of_def apply safe apply (meson IntI openin_subset subset_iff) by auto lemma in_closure_of: "x \ X closure_of S \ x \ topspace X \ (\T. x \ T \ openin X T \ (\y. y \ S \ y \ T))" by (auto simp: closure_of_def) lemma closure_of: "X closure_of S = topspace X \ (S \ X derived_set_of S)" by (fastforce simp: in_closure_of in_derived_set_of) lemma closure_of_alt: "X closure_of S = topspace X \ S \ X derived_set_of S" using derived_set_of_subset_topspace [of X S] unfolding closure_of_def in_derived_set_of by safe (auto simp: in_derived_set_of) lemma derived_set_of_subset_closure_of: "X derived_set_of S \ X closure_of S" by (fastforce simp: closure_of_def in_derived_set_of) lemma closure_of_subtopology: "(subtopology X U) closure_of S = U \ (X closure_of (U \ S))" unfolding closure_of_def topspace_subtopology openin_subtopology by safe (metis (full_types) IntI Int_iff inf.commute)+ lemma closure_of_empty [simp]: "X closure_of {} = {}" by (simp add: closure_of_alt) lemma closure_of_topspace [simp]: "X closure_of topspace X = topspace X" by (simp add: closure_of) lemma closure_of_UNIV [simp]: "X closure_of UNIV = topspace X" by (simp add: closure_of) lemma closure_of_subset_topspace: "X closure_of S \ topspace X" by (simp add: closure_of) lemma closure_of_subset_subtopology: "(subtopology X S) closure_of T \ S" by (simp add: closure_of_subtopology) lemma closure_of_mono: "S \ T \ X closure_of S \ X closure_of T" by (fastforce simp add: closure_of_def) lemma closure_of_subtopology_subset: "(subtopology X U) closure_of S \ (X closure_of S)" unfolding closure_of_subtopology by clarsimp (meson closure_of_mono contra_subsetD inf.cobounded2) lemma closure_of_subtopology_mono: "T \ U \ (subtopology X T) closure_of S \ (subtopology X U) closure_of S" unfolding closure_of_subtopology by auto (meson closure_of_mono inf_mono subset_iff) lemma closure_of_Un [simp]: "X closure_of (S \ T) = X closure_of S \ X closure_of T" by (simp add: Un_assoc Un_left_commute closure_of_alt derived_set_of_Un inf_sup_distrib1) lemma closure_of_Union: "finite \ \ X closure_of (\\) = (\S \ \. X closure_of S)" by (induction \ rule: finite_induct) auto lemma closure_of_subset: "S \ topspace X \ S \ X closure_of S" by (auto simp: closure_of_def) lemma closure_of_subset_Int: "topspace X \ S \ X closure_of S" by (auto simp: closure_of_def) lemma closure_of_subset_eq: "S \ topspace X \ X closure_of S \ S \ closedin X S" proof (cases "S \ topspace X") case True then have "\x. x \ topspace X \ (\T. x \ T \ openin X T \ (\y\S. y \ T)) \ x \ S \ openin X (topspace X - S)" apply (subst openin_subopen, safe) by (metis DiffI subset_eq openin_subset [of X]) then show ?thesis by (auto simp: closedin_def closure_of_def) next case False then show ?thesis by (simp add: closedin_def) qed lemma closure_of_eq: "X closure_of S = S \ closedin X S" proof (cases "S \ topspace X") case True then show ?thesis by (metis closure_of_subset closure_of_subset_eq set_eq_subset) next case False then show ?thesis using closure_of closure_of_subset_eq by fastforce qed lemma closedin_contains_derived_set: "closedin X S \ X derived_set_of S \ S \ S \ topspace X" proof (intro iffI conjI) show "closedin X S \ X derived_set_of S \ S" using closure_of_eq derived_set_of_subset_closure_of by fastforce show "closedin X S \ S \ topspace X" using closedin_subset by blast show "X derived_set_of S \ S \ S \ topspace X \ closedin X S" by (metis closure_of closure_of_eq inf.absorb_iff2 sup.orderE) qed lemma derived_set_subset_gen: "X derived_set_of S \ S \ closedin X (topspace X \ S)" by (simp add: closedin_contains_derived_set derived_set_of_restrict derived_set_of_subset_topspace) lemma derived_set_subset: "S \ topspace X \ (X derived_set_of S \ S \ closedin X S)" by (simp add: closedin_contains_derived_set) lemma closedin_derived_set: "closedin (subtopology X T) S \ S \ topspace X \ S \ T \ (\x. x \ X derived_set_of S \ x \ T \ x \ S)" by (auto simp: closedin_contains_derived_set derived_set_of_subtopology Int_absorb1) lemma closedin_Int_closure_of: "closedin (subtopology X S) T \ S \ X closure_of T = T" by (metis Int_left_absorb closure_of_eq closure_of_subtopology) lemma closure_of_closedin: "closedin X S \ X closure_of S = S" by (simp add: closure_of_eq) lemma closure_of_eq_diff: "X closure_of S = topspace X - \{T. openin X T \ disjnt S T}" by (auto simp: closure_of_def disjnt_iff) lemma closedin_closure_of [simp]: "closedin X (X closure_of S)" unfolding closure_of_eq_diff by blast lemma closure_of_closure_of [simp]: "X closure_of (X closure_of S) = X closure_of S" by (simp add: closure_of_eq) lemma closure_of_hull: assumes "S \ topspace X" shows "X closure_of S = (closedin X) hull S" proof (rule hull_unique [THEN sym]) show "S \ X closure_of S" by (simp add: closure_of_subset assms) next show "closedin X (X closure_of S)" by simp show "\T. \S \ T; closedin X T\ \ X closure_of S \ T" by (metis closure_of_eq closure_of_mono) qed lemma closure_of_minimal: "\S \ T; closedin X T\ \ (X closure_of S) \ T" by (metis closure_of_eq closure_of_mono) lemma closure_of_minimal_eq: "\S \ topspace X; closedin X T\ \ (X closure_of S) \ T \ S \ T" by (meson closure_of_minimal closure_of_subset subset_trans) lemma closure_of_unique: "\S \ T; closedin X T; \T'. \S \ T'; closedin X T'\ \ T \ T'\ \ X closure_of S = T" by (meson closedin_closure_of closedin_subset closure_of_minimal closure_of_subset eq_iff order.trans) lemma closure_of_eq_empty_gen: "X closure_of S = {} \ disjnt (topspace X) S" unfolding disjnt_def closure_of_restrict [where S=S] using closure_of by fastforce lemma closure_of_eq_empty: "S \ topspace X \ X closure_of S = {} \ S = {}" using closure_of_subset by fastforce lemma openin_Int_closure_of_subset: assumes "openin X S" shows "S \ X closure_of T \ X closure_of (S \ T)" proof - have "S \ X derived_set_of T = S \ X derived_set_of (S \ T)" by (meson assms openin_Int_derived_set_of_eq) moreover have "S \ (S \ T) = S \ T" by fastforce ultimately show ?thesis by (metis closure_of_alt inf.cobounded2 inf_left_commute inf_sup_distrib1) qed lemma closure_of_openin_Int_closure_of: assumes "openin X S" shows "X closure_of (S \ X closure_of T) = X closure_of (S \ T)" proof show "X closure_of (S \ X closure_of T) \ X closure_of (S \ T)" by (simp add: assms closure_of_minimal openin_Int_closure_of_subset) next show "X closure_of (S \ T) \ X closure_of (S \ X closure_of T)" by (metis Int_lower1 Int_subset_iff assms closedin_closure_of closure_of_minimal_eq closure_of_mono inf_le2 le_infI1 openin_subset) qed lemma openin_Int_closure_of_eq: "openin X S \ S \ X closure_of T = S \ X closure_of (S \ T)" apply (rule equalityI) apply (simp add: openin_Int_closure_of_subset) by (meson closure_of_mono inf.cobounded2 inf_mono subset_refl) lemma openin_Int_closure_of_eq_empty: "openin X S \ S \ X closure_of T = {} \ S \ T = {}" apply (subst openin_Int_closure_of_eq, auto) by (meson IntI closure_of_subset_Int disjoint_iff_not_equal openin_subset subset_eq) lemma closure_of_openin_Int_superset: "openin X S \ S \ X closure_of T \ X closure_of (S \ T) = X closure_of S" by (metis closure_of_openin_Int_closure_of inf.orderE) lemma closure_of_openin_subtopology_Int_closure_of: assumes S: "openin (subtopology X U) S" and "T \ U" shows "X closure_of (S \ X closure_of T) = X closure_of (S \ T)" (is "?lhs = ?rhs") proof obtain S0 where S0: "openin X S0" "S = S0 \ U" using assms by (auto simp: openin_subtopology) show "?lhs \ ?rhs" proof - have "S0 \ X closure_of T = S0 \ X closure_of (S0 \ T)" by (meson S0(1) openin_Int_closure_of_eq) moreover have "S0 \ T = S0 \ U \ T" using \T \ U\ by fastforce ultimately have "S \ X closure_of T \ X closure_of (S \ T)" using S0(2) by auto then show ?thesis by (meson closedin_closure_of closure_of_minimal) qed next show "?rhs \ ?lhs" proof - have "T \ S \ T \ X derived_set_of T" by force then show ?thesis by (metis Int_subset_iff S closure_of closure_of_mono inf.cobounded2 inf.coboundedI2 inf_commute openin_closedin_eq topspace_subtopology) qed qed lemma closure_of_subtopology_open: "openin X U \ S \ U \ (subtopology X U) closure_of S = U \ X closure_of S" by (metis closure_of_subtopology inf_absorb2 openin_Int_closure_of_eq) lemma discrete_topology_closure_of: "(discrete_topology U) closure_of S = U \ S" by (metis closedin_discrete_topology closure_of_restrict closure_of_unique discrete_topology_unique inf_sup_ord(1) order_refl) text\ Interior with respect to a topological space. \ definition interior_of :: "'a topology \ 'a set \ 'a set" (infixr "interior'_of" 80) where "X interior_of S \ {x. \T. openin X T \ x \ T \ T \ S}" lemma interior_of_restrict: "X interior_of S = X interior_of (topspace X \ S)" using openin_subset by (auto simp: interior_of_def) lemma interior_of_eq: "(X interior_of S = S) \ openin X S" unfolding interior_of_def using openin_subopen by blast lemma interior_of_openin: "openin X S \ X interior_of S = S" by (simp add: interior_of_eq) lemma interior_of_empty [simp]: "X interior_of {} = {}" by (simp add: interior_of_eq) lemma interior_of_topspace [simp]: "X interior_of (topspace X) = topspace X" by (simp add: interior_of_eq) lemma openin_interior_of [simp]: "openin X (X interior_of S)" unfolding interior_of_def using openin_subopen by fastforce lemma interior_of_interior_of [simp]: "X interior_of X interior_of S = X interior_of S" by (simp add: interior_of_eq) lemma interior_of_subset: "X interior_of S \ S" by (auto simp: interior_of_def) lemma interior_of_subset_closure_of: "X interior_of S \ X closure_of S" by (metis closure_of_subset_Int dual_order.trans interior_of_restrict interior_of_subset) lemma subset_interior_of_eq: "S \ X interior_of S \ openin X S" by (metis interior_of_eq interior_of_subset subset_antisym) lemma interior_of_mono: "S \ T \ X interior_of S \ X interior_of T" by (auto simp: interior_of_def) lemma interior_of_maximal: "\T \ S; openin X T\ \ T \ X interior_of S" by (auto simp: interior_of_def) lemma interior_of_maximal_eq: "openin X T \ T \ X interior_of S \ T \ S" by (meson interior_of_maximal interior_of_subset order_trans) lemma interior_of_unique: "\T \ S; openin X T; \T'. \T' \ S; openin X T'\ \ T' \ T\ \ X interior_of S = T" by (simp add: interior_of_maximal_eq interior_of_subset subset_antisym) lemma interior_of_subset_topspace: "X interior_of S \ topspace X" by (simp add: openin_subset) lemma interior_of_subset_subtopology: "(subtopology X S) interior_of T \ S" by (meson openin_imp_subset openin_interior_of) lemma interior_of_Int: "X interior_of (S \ T) = X interior_of S \ X interior_of T" apply (rule equalityI) apply (simp add: interior_of_mono) apply (auto simp: interior_of_maximal_eq openin_Int interior_of_subset le_infI1 le_infI2) done lemma interior_of_Inter_subset: "X interior_of (\\) \ (\S \ \. X interior_of S)" by (simp add: INT_greatest Inf_lower interior_of_mono) lemma union_interior_of_subset: "X interior_of S \ X interior_of T \ X interior_of (S \ T)" by (simp add: interior_of_mono) lemma interior_of_eq_empty: "X interior_of S = {} \ (\T. openin X T \ T \ S \ T = {})" by (metis bot.extremum_uniqueI interior_of_maximal interior_of_subset openin_interior_of) lemma interior_of_eq_empty_alt: "X interior_of S = {} \ (\T. openin X T \ T \ {} \ T - S \ {})" by (auto simp: interior_of_eq_empty) lemma interior_of_Union_openin_subsets: "\{T. openin X T \ T \ S} = X interior_of S" by (rule interior_of_unique [symmetric]) auto lemma interior_of_complement: "X interior_of (topspace X - S) = topspace X - X closure_of S" by (auto simp: interior_of_def closure_of_def) lemma interior_of_closure_of: "X interior_of S = topspace X - X closure_of (topspace X - S)" unfolding interior_of_complement [symmetric] by (metis Diff_Diff_Int interior_of_restrict) lemma closure_of_interior_of: "X closure_of S = topspace X - X interior_of (topspace X - S)" by (simp add: interior_of_complement Diff_Diff_Int closure_of) lemma closure_of_complement: "X closure_of (topspace X - S) = topspace X - X interior_of S" unfolding interior_of_def closure_of_def by (blast dest: openin_subset) lemma interior_of_eq_empty_complement: "X interior_of S = {} \ X closure_of (topspace X - S) = topspace X" using interior_of_subset_topspace [of X S] closure_of_complement by fastforce lemma closure_of_eq_topspace: "X closure_of S = topspace X \ X interior_of (topspace X - S) = {}" using closure_of_subset_topspace [of X S] interior_of_complement by fastforce lemma interior_of_subtopology_subset: "U \ X interior_of S \ (subtopology X U) interior_of S" by (auto simp: interior_of_def openin_subtopology) lemma interior_of_subtopology_subsets: "T \ U \ T \ (subtopology X U) interior_of S \ (subtopology X T) interior_of S" by (metis inf.absorb_iff2 interior_of_subtopology_subset subtopology_subtopology) lemma interior_of_subtopology_mono: "\S \ T; T \ U\ \ (subtopology X U) interior_of S \ (subtopology X T) interior_of S" by (metis dual_order.trans inf.orderE inf_commute interior_of_subset interior_of_subtopology_subsets) lemma interior_of_subtopology_open: assumes "openin X U" shows "(subtopology X U) interior_of S = U \ X interior_of S" proof - have "\A. U \ X closure_of (U \ A) = U \ X closure_of A" using assms openin_Int_closure_of_eq by blast then have "topspace X \ U - U \ X closure_of (topspace X \ U - S) = U \ (topspace X - X closure_of (topspace X - S))" by (metis (no_types) Diff_Int_distrib Int_Diff inf_commute) then show ?thesis unfolding interior_of_closure_of closure_of_subtopology_open topspace_subtopology using openin_Int_closure_of_eq [OF assms] by (metis assms closure_of_subtopology_open) qed lemma dense_intersects_open: "X closure_of S = topspace X \ (\T. openin X T \ T \ {} \ S \ T \ {})" proof - have "X closure_of S = topspace X \ (topspace X - X interior_of (topspace X - S) = topspace X)" by (simp add: closure_of_interior_of) also have "\ \ X interior_of (topspace X - S) = {}" by (simp add: closure_of_complement interior_of_eq_empty_complement) also have "\ \ (\T. openin X T \ T \ {} \ S \ T \ {})" unfolding interior_of_eq_empty_alt using openin_subset by fastforce finally show ?thesis . qed lemma interior_of_closedin_union_empty_interior_of: assumes "closedin X S" and disj: "X interior_of T = {}" shows "X interior_of (S \ T) = X interior_of S" proof - have "X closure_of (topspace X - T) = topspace X" by (metis Diff_Diff_Int disj closure_of_eq_topspace closure_of_restrict interior_of_closure_of) then show ?thesis unfolding interior_of_closure_of by (metis Diff_Un Diff_subset assms(1) closedin_def closure_of_openin_Int_superset) qed lemma interior_of_union_eq_empty: "closedin X S \ (X interior_of (S \ T) = {} \ X interior_of S = {} \ X interior_of T = {})" by (metis interior_of_closedin_union_empty_interior_of le_sup_iff subset_empty union_interior_of_subset) lemma discrete_topology_interior_of [simp]: "(discrete_topology U) interior_of S = U \ S" by (simp add: interior_of_restrict [of _ S] interior_of_eq) subsection \Frontier with respect to topological space \ definition frontier_of :: "'a topology \ 'a set \ 'a set" (infixr "frontier'_of" 80) where "X frontier_of S \ X closure_of S - X interior_of S" lemma frontier_of_closures: "X frontier_of S = X closure_of S \ X closure_of (topspace X - S)" by (metis Diff_Diff_Int closure_of_complement closure_of_subset_topspace double_diff frontier_of_def interior_of_subset_closure_of) lemma interior_of_union_frontier_of [simp]: "X interior_of S \ X frontier_of S = X closure_of S" by (simp add: frontier_of_def interior_of_subset_closure_of subset_antisym) lemma frontier_of_restrict: "X frontier_of S = X frontier_of (topspace X \ S)" by (metis closure_of_restrict frontier_of_def interior_of_restrict) lemma closedin_frontier_of: "closedin X (X frontier_of S)" by (simp add: closedin_Int frontier_of_closures) lemma frontier_of_subset_topspace: "X frontier_of S \ topspace X" by (simp add: closedin_frontier_of closedin_subset) lemma frontier_of_subset_subtopology: "(subtopology X S) frontier_of T \ S" by (metis (no_types) closedin_derived_set closedin_frontier_of) lemma frontier_of_subtopology_subset: "U \ (subtopology X U) frontier_of S \ (X frontier_of S)" proof - have "U \ X interior_of S - subtopology X U interior_of S = {}" by (simp add: interior_of_subtopology_subset) moreover have "X closure_of S \ subtopology X U closure_of S = subtopology X U closure_of S" by (meson closure_of_subtopology_subset inf.absorb_iff2) ultimately show ?thesis unfolding frontier_of_def by blast qed lemma frontier_of_subtopology_mono: "\S \ T; T \ U\ \ (subtopology X T) frontier_of S \ (subtopology X U) frontier_of S" by (simp add: frontier_of_def Diff_mono closure_of_subtopology_mono interior_of_subtopology_mono) lemma clopenin_eq_frontier_of: "closedin X S \ openin X S \ S \ topspace X \ X frontier_of S = {}" proof (cases "S \ topspace X") case True then show ?thesis by (metis Diff_eq_empty_iff closure_of_eq closure_of_subset_eq frontier_of_def interior_of_eq interior_of_subset interior_of_union_frontier_of sup_bot_right) next case False then show ?thesis by (simp add: frontier_of_closures openin_closedin_eq) qed lemma frontier_of_eq_empty: "S \ topspace X \ (X frontier_of S = {} \ closedin X S \ openin X S)" by (simp add: clopenin_eq_frontier_of) lemma frontier_of_openin: "openin X S \ X frontier_of S = X closure_of S - S" by (metis (no_types) frontier_of_def interior_of_eq) lemma frontier_of_openin_straddle_Int: assumes "openin X U" "U \ X frontier_of S \ {}" shows "U \ S \ {}" "U - S \ {}" proof - have "U \ (X closure_of S \ X closure_of (topspace X - S)) \ {}" using assms by (simp add: frontier_of_closures) then show "U \ S \ {}" using assms openin_Int_closure_of_eq_empty by fastforce show "U - S \ {}" proof - have "\A. X closure_of (A - S) \ U \ {}" using \U \ (X closure_of S \ X closure_of (topspace X - S)) \ {}\ by blast then have "\ U \ S" by (metis Diff_disjoint Diff_eq_empty_iff Int_Diff assms(1) inf_commute openin_Int_closure_of_eq_empty) then show ?thesis by blast qed qed lemma frontier_of_subset_closedin: "closedin X S \ (X frontier_of S) \ S" using closure_of_eq frontier_of_def by fastforce lemma frontier_of_empty [simp]: "X frontier_of {} = {}" by (simp add: frontier_of_def) lemma frontier_of_topspace [simp]: "X frontier_of topspace X = {}" by (simp add: frontier_of_def) lemma frontier_of_subset_eq: assumes "S \ topspace X" shows "(X frontier_of S) \ S \ closedin X S" proof show "X frontier_of S \ S \ closedin X S" by (metis assms closure_of_subset_eq interior_of_subset interior_of_union_frontier_of le_sup_iff) show "closedin X S \ X frontier_of S \ S" by (simp add: frontier_of_subset_closedin) qed lemma frontier_of_complement: "X frontier_of (topspace X - S) = X frontier_of S" by (metis Diff_Diff_Int closure_of_restrict frontier_of_closures inf_commute) lemma frontier_of_disjoint_eq: assumes "S \ topspace X" shows "((X frontier_of S) \ S = {} \ openin X S)" proof assume "X frontier_of S \ S = {}" then have "closedin X (topspace X - S)" using assms closure_of_subset frontier_of_def interior_of_eq interior_of_subset by fastforce then show "openin X S" using assms by (simp add: openin_closedin) next show "openin X S \ X frontier_of S \ S = {}" by (simp add: Diff_Diff_Int closedin_def frontier_of_openin inf.absorb_iff2 inf_commute) qed lemma frontier_of_disjoint_eq_alt: "S \ (topspace X - X frontier_of S) \ openin X S" proof (cases "S \ topspace X") case True show ?thesis using True frontier_of_disjoint_eq by auto next case False then show ?thesis by (meson Diff_subset openin_subset subset_trans) qed lemma frontier_of_Int: "X frontier_of (S \ T) = X closure_of (S \ T) \ (X frontier_of S \ X frontier_of T)" proof - have *: "U \ S \ U \ T \ U \ (S \ A \ T \ B) = U \ (A \ B)" for U S T A B :: "'a set" by blast show ?thesis by (simp add: frontier_of_closures closure_of_mono Diff_Int * flip: closure_of_Un) qed lemma frontier_of_Int_subset: "X frontier_of (S \ T) \ X frontier_of S \ X frontier_of T" by (simp add: frontier_of_Int) lemma frontier_of_Int_closedin: "\closedin X S; closedin X T\ \ X frontier_of(S \ T) = X frontier_of S \ T \ S \ X frontier_of T" apply (simp add: frontier_of_Int closedin_Int closure_of_closedin) using frontier_of_subset_closedin by blast lemma frontier_of_Un_subset: "X frontier_of(S \ T) \ X frontier_of S \ X frontier_of T" by (metis Diff_Un frontier_of_Int_subset frontier_of_complement) lemma frontier_of_Union_subset: "finite \ \ X frontier_of (\\) \ (\T \ \. X frontier_of T)" proof (induction \ rule: finite_induct) case (insert A \) then show ?case using frontier_of_Un_subset by fastforce qed simp lemma frontier_of_frontier_of_subset: "X frontier_of (X frontier_of S) \ X frontier_of S" by (simp add: closedin_frontier_of frontier_of_subset_closedin) lemma frontier_of_subtopology_open: "openin X U \ (subtopology X U) frontier_of S = U \ X frontier_of S" by (simp add: Diff_Int_distrib closure_of_subtopology_open frontier_of_def interior_of_subtopology_open) lemma discrete_topology_frontier_of [simp]: "(discrete_topology U) frontier_of S = {}" by (simp add: Diff_eq discrete_topology_closure_of frontier_of_closures) subsection\Locally finite collections\ definition locally_finite_in where "locally_finite_in X \ \ (\\ \ topspace X) \ (\x \ topspace X. \V. openin X V \ x \ V \ finite {U \ \. U \ V \ {}})" lemma finite_imp_locally_finite_in: "\finite \; \\ \ topspace X\ \ locally_finite_in X \" by (auto simp: locally_finite_in_def) lemma locally_finite_in_subset: assumes "locally_finite_in X \" "\ \ \" shows "locally_finite_in X \" proof - have "finite {U \ \. U \ V \ {}} \ finite {U \ \. U \ V \ {}}" for V apply (erule rev_finite_subset) using \\ \ \\ by blast then show ?thesis using assms unfolding locally_finite_in_def by (fastforce simp add:) qed lemma locally_finite_in_refinement: assumes \: "locally_finite_in X \" and f: "\S. S \ \ \ f S \ S" shows "locally_finite_in X (f ` \)" proof - show ?thesis unfolding locally_finite_in_def proof safe fix x assume "x \ topspace X" then obtain V where "openin X V" "x \ V" "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def by blast moreover have "{U \ \. f U \ V \ {}} \ {U \ \. U \ V \ {}}" for V using f by blast ultimately have "finite {U \ \. f U \ V \ {}}" using finite_subset by blast moreover have "f ` {U \ \. f U \ V \ {}} = {U \ f ` \. U \ V \ {}}" by blast ultimately have "finite {U \ f ` \. U \ V \ {}}" by (metis (no_types, lifting) finite_imageI) then show "\V. openin X V \ x \ V \ finite {U \ f ` \. U \ V \ {}}" using \openin X V\ \x \ V\ by blast next show "\x xa. \xa \ \; x \ f xa\ \ x \ topspace X" by (meson Sup_upper \ f locally_finite_in_def subset_iff) qed qed lemma locally_finite_in_subtopology: assumes \: "locally_finite_in X \" "\\ \ S" shows "locally_finite_in (subtopology X S) \" unfolding locally_finite_in_def proof safe fix x assume x: "x \ topspace (subtopology X S)" then obtain V where "openin X V" "x \ V" and fin: "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def topspace_subtopology by blast show "\V. openin (subtopology X S) V \ x \ V \ finite {U \ \. U \ V \ {}}" proof (intro exI conjI) show "openin (subtopology X S) (S \ V)" by (simp add: \openin X V\ openin_subtopology_Int2) have "{U \ \. U \ (S \ V) \ {}} \ {U \ \. U \ V \ {}}" by auto with fin show "finite {U \ \. U \ (S \ V) \ {}}" using finite_subset by auto show "x \ S \ V" using x \x \ V\ by (simp) qed next show "\x A. \x \ A; A \ \\ \ x \ topspace (subtopology X S)" using assms unfolding locally_finite_in_def topspace_subtopology by blast qed lemma closedin_locally_finite_Union: assumes clo: "\S. S \ \ \ closedin X S" and \: "locally_finite_in X \" shows "closedin X (\\)" using \ unfolding locally_finite_in_def closedin_def proof clarify show "openin X (topspace X - \\)" proof (subst openin_subopen, clarify) fix x assume "x \ topspace X" and "x \ \\" then obtain V where "openin X V" "x \ V" and fin: "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def by blast let ?T = "V - \{S \ \. S \ V \ {}}" show "\T. openin X T \ x \ T \ T \ topspace X - \\" proof (intro exI conjI) show "openin X ?T" by (metis (no_types, lifting) fin \openin X V\ clo closedin_Union mem_Collect_eq openin_diff) show "x \ ?T" using \x \ \\\ \x \ V\ by auto show "?T \ topspace X - \\" using \openin X V\ openin_subset by auto qed qed qed lemma locally_finite_in_closure: assumes \: "locally_finite_in X \" shows "locally_finite_in X ((\S. X closure_of S) ` \)" using \ unfolding locally_finite_in_def proof (intro conjI; clarsimp) fix x A assume "x \ X closure_of A" then show "x \ topspace X" by (meson in_closure_of) next fix x assume "x \ topspace X" and "\\ \ topspace X" then obtain V where V: "openin X V" "x \ V" and fin: "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def by blast have eq: "{y \ f ` \. Q y} = f ` {x. x \ \ \ Q(f x)}" for f Q by blast have eq2: "{A \ \. X closure_of A \ V \ {}} = {A \ \. A \ V \ {}}" using openin_Int_closure_of_eq_empty V by blast have "finite {U \ (closure_of) X ` \. U \ V \ {}}" by (simp add: eq eq2 fin) with V show "\V. openin X V \ x \ V \ finite {U \ (closure_of) X ` \. U \ V \ {}}" by blast qed lemma closedin_Union_locally_finite_closure: "locally_finite_in X \ \ closedin X (\((\S. X closure_of S) ` \))" by (metis (mono_tags) closedin_closure_of closedin_locally_finite_Union imageE locally_finite_in_closure) lemma closure_of_Union_subset: "\((\S. X closure_of S) ` \) \ X closure_of (\\)" by clarify (meson Union_upper closure_of_mono subsetD) lemma closure_of_locally_finite_Union: "locally_finite_in X \ \ X closure_of (\\) = \((\S. X closure_of S) ` \)" apply (rule closure_of_unique) apply (simp add: SUP_upper2 Sup_le_iff closure_of_subset locally_finite_in_def) apply (simp add: closedin_Union_locally_finite_closure) by (simp add: Sup_le_iff closure_of_minimal) subsection\<^marker>\tag important\ \Continuous maps\ text \We will need to deal with continuous maps in terms of topologies and not in terms of type classes, as defined below.\ definition continuous_map where "continuous_map X Y f \ (\x \ topspace X. f x \ topspace Y) \ (\U. openin Y U \ openin X {x \ topspace X. f x \ U})" lemma continuous_map: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\U. openin Y U \ openin X {x \ topspace X. f x \ U})" by (auto simp: continuous_map_def) lemma continuous_map_image_subset_topspace: "continuous_map X Y f \ f ` (topspace X) \ topspace Y" by (auto simp: continuous_map_def) lemma continuous_map_on_empty: "topspace X = {} \ continuous_map X Y f" by (auto simp: continuous_map_def) lemma continuous_map_closedin: "continuous_map X Y f \ (\x \ topspace X. f x \ topspace Y) \ (\C. closedin Y C \ closedin X {x \ topspace X. f x \ C})" proof - have "(\U. openin Y U \ openin X {x \ topspace X. f x \ U}) = (\C. closedin Y C \ closedin X {x \ topspace X. f x \ C})" if "\x. x \ topspace X \ f x \ topspace Y" proof - have eq: "{x \ topspace X. f x \ topspace Y \ f x \ C} = (topspace X - {x \ topspace X. f x \ C})" for C using that by blast show ?thesis proof (intro iffI allI impI) fix C assume "\U. openin Y U \ openin X {x \ topspace X. f x \ U}" and "closedin Y C" then have "openin X {x \ topspace X. f x \ topspace Y - C}" by blast then show "closedin X {x \ topspace X. f x \ C}" by (auto simp add: closedin_def eq) next fix U assume "\C. closedin Y C \ closedin X {x \ topspace X. f x \ C}" and "openin Y U" then have "closedin X {x \ topspace X. f x \ topspace Y - U}" by blast then show "openin X {x \ topspace X. f x \ U}" by (auto simp add: openin_closedin_eq eq) qed qed then show ?thesis by (auto simp: continuous_map_def) qed lemma openin_continuous_map_preimage: "\continuous_map X Y f; openin Y U\ \ openin X {x \ topspace X. f x \ U}" by (simp add: continuous_map_def) lemma closedin_continuous_map_preimage: "\continuous_map X Y f; closedin Y C\ \ closedin X {x \ topspace X. f x \ C}" by (simp add: continuous_map_closedin) lemma openin_continuous_map_preimage_gen: assumes "continuous_map X Y f" "openin X U" "openin Y V" shows "openin X {x \ U. f x \ V}" proof - have eq: "{x \ U. f x \ V} = U \ {x \ topspace X. f x \ V}" using assms(2) openin_closedin_eq by fastforce show ?thesis unfolding eq using assms openin_continuous_map_preimage by fastforce qed lemma closedin_continuous_map_preimage_gen: assumes "continuous_map X Y f" "closedin X U" "closedin Y V" shows "closedin X {x \ U. f x \ V}" proof - have eq: "{x \ U. f x \ V} = U \ {x \ topspace X. f x \ V}" using assms(2) closedin_def by fastforce show ?thesis unfolding eq using assms closedin_continuous_map_preimage by fastforce qed lemma continuous_map_image_closure_subset: assumes "continuous_map X Y f" shows "f ` (X closure_of S) \ Y closure_of f ` S" proof - have *: "f ` (topspace X) \ topspace Y" by (meson assms continuous_map) have "X closure_of T \ {x \ X closure_of T. f x \ Y closure_of (f ` T)}" if "T \ topspace X" for T proof (rule closure_of_minimal) show "T \ {x \ X closure_of T. f x \ Y closure_of f ` T}" using closure_of_subset * that by (fastforce simp: in_closure_of) next show "closedin X {x \ X closure_of T. f x \ Y closure_of f ` T}" using assms closedin_continuous_map_preimage_gen by fastforce qed then have "f ` (X closure_of (topspace X \ S)) \ Y closure_of (f ` (topspace X \ S))" by blast also have "\ \ Y closure_of (topspace Y \ f ` S)" using * by (blast intro!: closure_of_mono) finally have "f ` (X closure_of (topspace X \ S)) \ Y closure_of (topspace Y \ f ` S)" . then show ?thesis by (metis closure_of_restrict) qed lemma continuous_map_subset_aux1: "continuous_map X Y f \ (\S. f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_image_closure_subset by blast lemma continuous_map_subset_aux2: assumes "\S. S \ topspace X \ f ` (X closure_of S) \ Y closure_of f ` S" shows "continuous_map X Y f" unfolding continuous_map_closedin proof (intro conjI ballI allI impI) fix x assume "x \ topspace X" then show "f x \ topspace Y" using assms closure_of_subset_topspace by fastforce next fix C assume "closedin Y C" then show "closedin X {x \ topspace X. f x \ C}" proof (clarsimp simp flip: closure_of_subset_eq, intro conjI) fix x assume x: "x \ X closure_of {x \ topspace X. f x \ C}" and "C \ topspace Y" and "Y closure_of C \ C" show "x \ topspace X" by (meson x in_closure_of) have "{a \ topspace X. f a \ C} \ topspace X" by simp moreover have "Y closure_of f ` {a \ topspace X. f a \ C} \ C" by (simp add: \closedin Y C\ closure_of_minimal image_subset_iff) ultimately have "f ` (X closure_of {a \ topspace X. f a \ C}) \ C" using assms by blast then show "f x \ C" using x by auto qed qed lemma continuous_map_eq_image_closure_subset: "continuous_map X Y f \ (\S. f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_subset_aux1 continuous_map_subset_aux2 by metis lemma continuous_map_eq_image_closure_subset_alt: "continuous_map X Y f \ (\S. S \ topspace X \ f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_subset_aux1 continuous_map_subset_aux2 by metis lemma continuous_map_eq_image_closure_subset_gen: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\S. f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_subset_aux1 continuous_map_subset_aux2 continuous_map_image_subset_topspace by metis lemma continuous_map_closure_preimage_subset: "continuous_map X Y f \ X closure_of {x \ topspace X. f x \ T} \ {x \ topspace X. f x \ Y closure_of T}" unfolding continuous_map_closedin by (rule closure_of_minimal) (use in_closure_of in \fastforce+\) lemma continuous_map_frontier_frontier_preimage_subset: assumes "continuous_map X Y f" shows "X frontier_of {x \ topspace X. f x \ T} \ {x \ topspace X. f x \ Y frontier_of T}" proof - have eq: "topspace X - {x \ topspace X. f x \ T} = {x \ topspace X. f x \ topspace Y - T}" using assms unfolding continuous_map_def by blast have "X closure_of {x \ topspace X. f x \ T} \ {x \ topspace X. f x \ Y closure_of T}" by (simp add: assms continuous_map_closure_preimage_subset) moreover have "X closure_of (topspace X - {x \ topspace X. f x \ T}) \ {x \ topspace X. f x \ Y closure_of (topspace Y - T)}" using continuous_map_closure_preimage_subset [OF assms] eq by presburger ultimately show ?thesis by (auto simp: frontier_of_closures) qed lemma topology_finer_continuous_id: "topspace X = topspace Y \ ((\S. openin X S \ openin Y S) \ continuous_map Y X id)" unfolding continuous_map_def apply auto using openin_subopen openin_subset apply fastforce using openin_subopen topspace_def by fastforce lemma continuous_map_const [simp]: "continuous_map X Y (\x. C) \ topspace X = {} \ C \ topspace Y" proof (cases "topspace X = {}") case False show ?thesis proof (cases "C \ topspace Y") case True with openin_subopen show ?thesis by (auto simp: continuous_map_def) next case False then show ?thesis unfolding continuous_map_def by fastforce qed qed (auto simp: continuous_map_on_empty) declare continuous_map_const [THEN iffD2, continuous_intros] lemma continuous_map_compose [continuous_intros]: assumes f: "continuous_map X X' f" and g: "continuous_map X' X'' g" shows "continuous_map X X'' (g \ f)" unfolding continuous_map_def proof (intro conjI ballI allI impI) fix x assume "x \ topspace X" then show "(g \ f) x \ topspace X''" using assms unfolding continuous_map_def by force next fix U assume "openin X'' U" have eq: "{x \ topspace X. (g \ f) x \ U} = {x \ topspace X. f x \ {y. y \ topspace X' \ g y \ U}}" by auto (meson f continuous_map_def) show "openin X {x \ topspace X. (g \ f) x \ U}" unfolding eq using assms unfolding continuous_map_def using \openin X'' U\ by blast qed lemma continuous_map_eq: assumes "continuous_map X X' f" and "\x. x \ topspace X \ f x = g x" shows "continuous_map X X' g" proof - have eq: "{x \ topspace X. f x \ U} = {x \ topspace X. g x \ U}" for U using assms by auto show ?thesis using assms by (simp add: continuous_map_def eq) qed lemma restrict_continuous_map [simp]: "topspace X \ S \ continuous_map X X' (restrict f S) \ continuous_map X X' f" by (auto simp: elim!: continuous_map_eq) lemma continuous_map_in_subtopology: "continuous_map X (subtopology X' S) f \ continuous_map X X' f \ f ` (topspace X) \ S" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof - have "\A. f ` (X closure_of A) \ subtopology X' S closure_of f ` A" by (meson L continuous_map_image_closure_subset) then show ?thesis by (metis (no_types) closure_of_subset_subtopology closure_of_subtopology_subset closure_of_topspace continuous_map_eq_image_closure_subset dual_order.trans) qed next assume R: ?rhs then have eq: "{x \ topspace X. f x \ U} = {x \ topspace X. f x \ U \ f x \ S}" for U by auto show ?lhs using R unfolding continuous_map by (auto simp: openin_subtopology eq) qed lemma continuous_map_from_subtopology: "continuous_map X X' f \ continuous_map (subtopology X S) X' f" by (auto simp: continuous_map openin_subtopology) lemma continuous_map_into_fulltopology: "continuous_map X (subtopology X' T) f \ continuous_map X X' f" by (auto simp: continuous_map_in_subtopology) lemma continuous_map_into_subtopology: "\continuous_map X X' f; f ` topspace X \ T\ \ continuous_map X (subtopology X' T) f" by (auto simp: continuous_map_in_subtopology) lemma continuous_map_from_subtopology_mono: "\continuous_map (subtopology X T) X' f; S \ T\ \ continuous_map (subtopology X S) X' f" by (metis inf.absorb_iff2 continuous_map_from_subtopology subtopology_subtopology) lemma continuous_map_from_discrete_topology [simp]: "continuous_map (discrete_topology U) X f \ f ` U \ topspace X" by (auto simp: continuous_map_def) lemma continuous_map_iff_continuous [simp]: "continuous_map (top_of_set S) euclidean g = continuous_on S g" by (fastforce simp add: continuous_map openin_subtopology continuous_on_open_invariant) lemma continuous_map_iff_continuous2 [simp]: "continuous_map euclidean euclidean g = continuous_on UNIV g" by (metis continuous_map_iff_continuous subtopology_UNIV) lemma continuous_map_openin_preimage_eq: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\U. openin Y U \ openin X (topspace X \ f -` U))" by (auto simp: continuous_map_def vimage_def Int_def) lemma continuous_map_closedin_preimage_eq: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\U. closedin Y U \ closedin X (topspace X \ f -` U))" by (auto simp: continuous_map_closedin vimage_def Int_def) lemma continuous_map_square_root: "continuous_map euclideanreal euclideanreal sqrt" by (simp add: continuous_at_imp_continuous_on isCont_real_sqrt) lemma continuous_map_sqrt [continuous_intros]: "continuous_map X euclideanreal f \ continuous_map X euclideanreal (\x. sqrt(f x))" by (meson continuous_map_compose continuous_map_eq continuous_map_square_root o_apply) lemma continuous_map_id [simp, continuous_intros]: "continuous_map X X id" unfolding continuous_map_def using openin_subopen topspace_def by fastforce declare continuous_map_id [unfolded id_def, simp, continuous_intros] lemma continuous_map_id_subt [simp]: "continuous_map (subtopology X S) X id" by (simp add: continuous_map_from_subtopology) declare continuous_map_id_subt [unfolded id_def, simp] lemma\<^marker>\tag important\ continuous_map_alt: "continuous_map T1 T2 f = ((\U. openin T2 U \ openin T1 (f -` U \ topspace T1)) \ f ` topspace T1 \ topspace T2)" by (auto simp: continuous_map_def vimage_def image_def Collect_conj_eq inf_commute) lemma continuous_map_open [intro]: "continuous_map T1 T2 f \ openin T2 U \ openin T1 (f-`U \ topspace(T1))" unfolding continuous_map_alt by auto lemma continuous_map_preimage_topspace [intro]: assumes "continuous_map T1 T2 f" shows "f-`(topspace T2) \ topspace T1 = topspace T1" using assms unfolding continuous_map_def by auto subsection\Open and closed maps (not a priori assumed continuous)\ definition open_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "open_map X1 X2 f \ \U. openin X1 U \ openin X2 (f ` U)" definition closed_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "closed_map X1 X2 f \ \U. closedin X1 U \ closedin X2 (f ` U)" lemma open_map_imp_subset_topspace: "open_map X1 X2 f \ f ` (topspace X1) \ topspace X2" unfolding open_map_def by (simp add: openin_subset) lemma open_map_on_empty: "topspace X = {} \ open_map X Y f" by (metis empty_iff imageE in_mono open_map_def openin_subopen openin_subset) lemma closed_map_on_empty: "topspace X = {} \ closed_map X Y f" by (simp add: closed_map_def closedin_topspace_empty) lemma closed_map_const: "closed_map X Y (\x. c) \ topspace X = {} \ closedin Y {c}" proof (cases "topspace X = {}") case True then show ?thesis by (simp add: closed_map_on_empty) next case False then show ?thesis by (auto simp: closed_map_def image_constant_conv) qed lemma open_map_imp_subset: "\open_map X1 X2 f; S \ topspace X1\ \ f ` S \ topspace X2" by (meson order_trans open_map_imp_subset_topspace subset_image_iff) lemma topology_finer_open_id: "(\S. openin X S \ openin X' S) \ open_map X X' id" unfolding open_map_def by auto lemma open_map_id: "open_map X X id" unfolding open_map_def by auto lemma open_map_eq: "\open_map X X' f; \x. x \ topspace X \ f x = g x\ \ open_map X X' g" unfolding open_map_def by (metis image_cong openin_subset subset_iff) lemma open_map_inclusion_eq: "open_map (subtopology X S) X id \ openin X (topspace X \ S)" proof - have *: "openin X (T \ S)" if "openin X (S \ topspace X)" "openin X T" for T proof - have "T \ topspace X" using that by (simp add: openin_subset) with that show "openin X (T \ S)" by (metis inf.absorb1 inf.left_commute inf_commute openin_Int) qed show ?thesis by (fastforce simp add: open_map_def Int_commute openin_subtopology_alt intro: *) qed lemma open_map_inclusion: "openin X S \ open_map (subtopology X S) X id" by (simp add: open_map_inclusion_eq openin_Int) lemma open_map_compose: "\open_map X X' f; open_map X' X'' g\ \ open_map X X'' (g \ f)" by (metis (no_types, lifting) image_comp open_map_def) lemma closed_map_imp_subset_topspace: "closed_map X1 X2 f \ f ` (topspace X1) \ topspace X2" by (simp add: closed_map_def closedin_subset) lemma closed_map_imp_subset: "\closed_map X1 X2 f; S \ topspace X1\ \ f ` S \ topspace X2" using closed_map_imp_subset_topspace by blast lemma topology_finer_closed_id: "(\S. closedin X S \ closedin X' S) \ closed_map X X' id" by (simp add: closed_map_def) lemma closed_map_id: "closed_map X X id" by (simp add: closed_map_def) lemma closed_map_eq: "\closed_map X X' f; \x. x \ topspace X \ f x = g x\ \ closed_map X X' g" unfolding closed_map_def by (metis image_cong closedin_subset subset_iff) lemma closed_map_compose: "\closed_map X X' f; closed_map X' X'' g\ \ closed_map X X'' (g \ f)" by (metis (no_types, lifting) closed_map_def image_comp) lemma closed_map_inclusion_eq: "closed_map (subtopology X S) X id \ closedin X (topspace X \ S)" proof - have *: "closedin X (T \ S)" if "closedin X (S \ topspace X)" "closedin X T" for T proof - have "T \ topspace X" using that by (simp add: closedin_subset) with that show "closedin X (T \ S)" by (metis inf.absorb1 inf.left_commute inf_commute closedin_Int) qed show ?thesis by (fastforce simp add: closed_map_def Int_commute closedin_subtopology_alt intro: *) qed lemma closed_map_inclusion: "closedin X S \ closed_map (subtopology X S) X id" by (simp add: closed_map_inclusion_eq closedin_Int) lemma open_map_into_subtopology: "\open_map X X' f; f ` topspace X \ S\ \ open_map X (subtopology X' S) f" unfolding open_map_def openin_subtopology using openin_subset by fastforce lemma closed_map_into_subtopology: "\closed_map X X' f; f ` topspace X \ S\ \ closed_map X (subtopology X' S) f" unfolding closed_map_def closedin_subtopology using closedin_subset by fastforce lemma open_map_into_discrete_topology: "open_map X (discrete_topology U) f \ f ` (topspace X) \ U" unfolding open_map_def openin_discrete_topology using openin_subset by blast lemma closed_map_into_discrete_topology: "closed_map X (discrete_topology U) f \ f ` (topspace X) \ U" unfolding closed_map_def closedin_discrete_topology using closedin_subset by blast lemma bijective_open_imp_closed_map: "\open_map X X' f; f ` (topspace X) = topspace X'; inj_on f (topspace X)\ \ closed_map X X' f" unfolding open_map_def closed_map_def closedin_def by auto (metis Diff_subset inj_on_image_set_diff) lemma bijective_closed_imp_open_map: "\closed_map X X' f; f ` (topspace X) = topspace X'; inj_on f (topspace X)\ \ open_map X X' f" unfolding closed_map_def open_map_def openin_closedin_eq by auto (metis Diff_subset inj_on_image_set_diff) lemma open_map_from_subtopology: "\open_map X X' f; openin X U\ \ open_map (subtopology X U) X' f" unfolding open_map_def openin_subtopology_alt by blast lemma closed_map_from_subtopology: "\closed_map X X' f; closedin X U\ \ closed_map (subtopology X U) X' f" unfolding closed_map_def closedin_subtopology_alt by blast lemma open_map_restriction: "\open_map X X' f; {x. x \ topspace X \ f x \ V} = U\ \ open_map (subtopology X U) (subtopology X' V) f" unfolding open_map_def openin_subtopology_alt apply clarify apply (rename_tac T) apply (rule_tac x="f ` T" in image_eqI) using openin_closedin_eq by fastforce+ lemma closed_map_restriction: "\closed_map X X' f; {x. x \ topspace X \ f x \ V} = U\ \ closed_map (subtopology X U) (subtopology X' V) f" unfolding closed_map_def closedin_subtopology_alt apply clarify apply (rename_tac T) apply (rule_tac x="f ` T" in image_eqI) using closedin_def by fastforce+ subsection\Quotient maps\ definition quotient_map where "quotient_map X X' f \ f ` (topspace X) = topspace X' \ (\U. U \ topspace X' \ (openin X {x. x \ topspace X \ f x \ U} \ openin X' U))" lemma quotient_map_eq: assumes "quotient_map X X' f" "\x. x \ topspace X \ f x = g x" shows "quotient_map X X' g" proof - have eq: "{x \ topspace X. f x \ U} = {x \ topspace X. g x \ U}" for U using assms by auto show ?thesis using assms unfolding quotient_map_def by (metis (mono_tags, lifting) eq image_cong) qed lemma quotient_map_compose: assumes f: "quotient_map X X' f" and g: "quotient_map X' X'' g" shows "quotient_map X X'' (g \ f)" unfolding quotient_map_def proof (intro conjI allI impI) show "(g \ f) ` topspace X = topspace X''" using assms by (simp only: image_comp [symmetric]) (simp add: quotient_map_def) next fix U'' assume "U'' \ topspace X''" define U' where "U' \ {y \ topspace X'. g y \ U''}" have "U' \ topspace X'" by (auto simp add: U'_def) then have U': "openin X {x \ topspace X. f x \ U'} = openin X' U'" using assms unfolding quotient_map_def by simp have eq: "{x \ topspace X. f x \ topspace X' \ g (f x) \ U''} = {x \ topspace X. (g \ f) x \ U''}" using f quotient_map_def by fastforce have "openin X {x \ topspace X. (g \ f) x \ U''} = openin X {x \ topspace X. f x \ U'}" using assms by (simp add: quotient_map_def U'_def eq) also have "\ = openin X'' U''" using U'_def \U'' \ topspace X''\ U' g quotient_map_def by fastforce finally show "openin X {x \ topspace X. (g \ f) x \ U''} = openin X'' U''" . qed lemma quotient_map_from_composition: assumes f: "continuous_map X X' f" and g: "continuous_map X' X'' g" and gf: "quotient_map X X'' (g \ f)" shows "quotient_map X' X'' g" unfolding quotient_map_def proof (intro conjI allI impI) show "g ` topspace X' = topspace X''" using assms unfolding continuous_map_def quotient_map_def by fastforce next fix U'' :: "'c set" assume U'': "U'' \ topspace X''" have eq: "{x \ topspace X. g (f x) \ U''} = {x \ topspace X. f x \ {y. y \ topspace X' \ g y \ U''}}" using continuous_map_def f by fastforce show "openin X' {x \ topspace X'. g x \ U''} = openin X'' U''" using assms unfolding continuous_map_def quotient_map_def by (metis (mono_tags, lifting) Collect_cong U'' comp_apply eq) qed lemma quotient_imp_continuous_map: "quotient_map X X' f \ continuous_map X X' f" by (simp add: continuous_map openin_subset quotient_map_def) lemma quotient_imp_surjective_map: "quotient_map X X' f \ f ` (topspace X) = topspace X'" by (simp add: quotient_map_def) lemma quotient_map_closedin: "quotient_map X X' f \ f ` (topspace X) = topspace X' \ (\U. U \ topspace X' \ (closedin X {x. x \ topspace X \ f x \ U} \ closedin X' U))" proof - have eq: "(topspace X - {x \ topspace X. f x \ U'}) = {x \ topspace X. f x \ topspace X' \ f x \ U'}" if "f ` topspace X = topspace X'" "U' \ topspace X'" for U' using that by auto have "(\U\topspace X'. openin X {x \ topspace X. f x \ U} = openin X' U) = (\U\topspace X'. closedin X {x \ topspace X. f x \ U} = closedin X' U)" if "f ` topspace X = topspace X'" proof (rule iffI; intro allI impI subsetI) fix U' assume *[rule_format]: "\U\topspace X'. openin X {x \ topspace X. f x \ U} = openin X' U" and U': "U' \ topspace X'" show "closedin X {x \ topspace X. f x \ U'} = closedin X' U'" using U' by (auto simp add: closedin_def simp flip: * [of "topspace X' - U'"] eq [OF that]) next fix U' :: "'b set" assume *[rule_format]: "\U\topspace X'. closedin X {x \ topspace X. f x \ U} = closedin X' U" and U': "U' \ topspace X'" show "openin X {x \ topspace X. f x \ U'} = openin X' U'" using U' by (auto simp add: openin_closedin_eq simp flip: * [of "topspace X' - U'"] eq [OF that]) qed then show ?thesis unfolding quotient_map_def by force qed lemma continuous_open_imp_quotient_map: assumes "continuous_map X X' f" and om: "open_map X X' f" and feq: "f ` (topspace X) = topspace X'" shows "quotient_map X X' f" proof - { fix U assume U: "U \ topspace X'" and "openin X {x \ topspace X. f x \ U}" then have ope: "openin X' (f ` {x \ topspace X. f x \ U})" using om unfolding open_map_def by blast then have "openin X' U" using U feq by (subst openin_subopen) force } moreover have "openin X {x \ topspace X. f x \ U}" if "U \ topspace X'" and "openin X' U" for U using that assms unfolding continuous_map_def by blast ultimately show ?thesis unfolding quotient_map_def using assms by blast qed lemma continuous_closed_imp_quotient_map: assumes "continuous_map X X' f" and om: "closed_map X X' f" and feq: "f ` (topspace X) = topspace X'" shows "quotient_map X X' f" proof - have "f ` {x \ topspace X. f x \ U} = U" if "U \ topspace X'" for U using that feq by auto with assms show ?thesis unfolding quotient_map_closedin closed_map_def continuous_map_closedin by auto qed lemma continuous_open_quotient_map: "\continuous_map X X' f; open_map X X' f\ \ quotient_map X X' f \ f ` (topspace X) = topspace X'" by (meson continuous_open_imp_quotient_map quotient_map_def) lemma continuous_closed_quotient_map: "\continuous_map X X' f; closed_map X X' f\ \ quotient_map X X' f \ f ` (topspace X) = topspace X'" by (meson continuous_closed_imp_quotient_map quotient_map_def) lemma injective_quotient_map: assumes "inj_on f (topspace X)" shows "quotient_map X X' f \ continuous_map X X' f \ open_map X X' f \ closed_map X X' f \ f ` (topspace X) = topspace X'" (is "?lhs = ?rhs") proof assume L: ?lhs have "open_map X X' f" proof (clarsimp simp add: open_map_def) fix U assume "openin X U" then have "U \ topspace X" by (simp add: openin_subset) moreover have "{x \ topspace X. f x \ f ` U} = U" using \U \ topspace X\ assms inj_onD by fastforce ultimately show "openin X' (f ` U)" using L unfolding quotient_map_def by (metis (no_types, lifting) Collect_cong \openin X U\ image_mono) qed moreover have "closed_map X X' f" proof (clarsimp simp add: closed_map_def) fix U assume "closedin X U" then have "U \ topspace X" by (simp add: closedin_subset) moreover have "{x \ topspace X. f x \ f ` U} = U" using \U \ topspace X\ assms inj_onD by fastforce ultimately show "closedin X' (f ` U)" using L unfolding quotient_map_closedin by (metis (no_types, lifting) Collect_cong \closedin X U\ image_mono) qed ultimately show ?rhs using L by (simp add: quotient_imp_continuous_map quotient_imp_surjective_map) next assume ?rhs then show ?lhs by (simp add: continuous_closed_imp_quotient_map) qed lemma continuous_compose_quotient_map: assumes f: "quotient_map X X' f" and g: "continuous_map X X'' (g \ f)" shows "continuous_map X' X'' g" unfolding quotient_map_def continuous_map_def proof (intro conjI ballI allI impI) show "\x'. x' \ topspace X' \ g x' \ topspace X''" using assms unfolding quotient_map_def by (metis (no_types, hide_lams) continuous_map_image_subset_topspace image_comp image_subset_iff) next fix U'' :: "'c set" assume U'': "openin X'' U''" have "f ` topspace X = topspace X'" by (simp add: f quotient_imp_surjective_map) then have eq: "{x \ topspace X. f x \ topspace X' \ g (f x) \ U} = {x \ topspace X. g (f x) \ U}" for U by auto have "openin X {x \ topspace X. f x \ topspace X' \ g (f x) \ U''}" unfolding eq using U'' g openin_continuous_map_preimage by fastforce then have *: "openin X {x \ topspace X. f x \ {x \ topspace X'. g x \ U''}}" by auto show "openin X' {x \ topspace X'. g x \ U''}" using f unfolding quotient_map_def by (metis (no_types) Collect_subset *) qed lemma continuous_compose_quotient_map_eq: "quotient_map X X' f \ continuous_map X X'' (g \ f) \ continuous_map X' X'' g" using continuous_compose_quotient_map continuous_map_compose quotient_imp_continuous_map by blast lemma quotient_map_compose_eq: "quotient_map X X' f \ quotient_map X X'' (g \ f) \ quotient_map X' X'' g" apply safe apply (meson continuous_compose_quotient_map_eq quotient_imp_continuous_map quotient_map_from_composition) by (simp add: quotient_map_compose) lemma quotient_map_restriction: assumes quo: "quotient_map X Y f" and U: "{x \ topspace X. f x \ V} = U" and disj: "openin Y V \ closedin Y V" shows "quotient_map (subtopology X U) (subtopology Y V) f" using disj proof assume V: "openin Y V" with U have sub: "U \ topspace X" "V \ topspace Y" by (auto simp: openin_subset) have fim: "f ` topspace X = topspace Y" and Y: "\U. U \ topspace Y \ openin X {x \ topspace X. f x \ U} = openin Y U" using quo unfolding quotient_map_def by auto have "openin X U" using U V Y sub(2) by blast show ?thesis unfolding quotient_map_def proof (intro conjI allI impI) show "f ` topspace (subtopology X U) = topspace (subtopology Y V)" using sub U fim by (auto) next fix Y' :: "'b set" assume "Y' \ topspace (subtopology Y V)" then have "Y' \ topspace Y" "Y' \ V" by (simp_all) then have eq: "{x \ topspace X. x \ U \ f x \ Y'} = {x \ topspace X. f x \ Y'}" using U by blast then show "openin (subtopology X U) {x \ topspace (subtopology X U). f x \ Y'} = openin (subtopology Y V) Y'" using U V Y \openin X U\ \Y' \ topspace Y\ \Y' \ V\ by (simp add: openin_open_subtopology eq) (auto simp: openin_closedin_eq) qed next assume V: "closedin Y V" with U have sub: "U \ topspace X" "V \ topspace Y" by (auto simp: closedin_subset) have fim: "f ` topspace X = topspace Y" and Y: "\U. U \ topspace Y \ closedin X {x \ topspace X. f x \ U} = closedin Y U" using quo unfolding quotient_map_closedin by auto have "closedin X U" using U V Y sub(2) by blast show ?thesis unfolding quotient_map_closedin proof (intro conjI allI impI) show "f ` topspace (subtopology X U) = topspace (subtopology Y V)" using sub U fim by (auto) next fix Y' :: "'b set" assume "Y' \ topspace (subtopology Y V)" then have "Y' \ topspace Y" "Y' \ V" by (simp_all) then have eq: "{x \ topspace X. x \ U \ f x \ Y'} = {x \ topspace X. f x \ Y'}" using U by blast then show "closedin (subtopology X U) {x \ topspace (subtopology X U). f x \ Y'} = closedin (subtopology Y V) Y'" using U V Y \closedin X U\ \Y' \ topspace Y\ \Y' \ V\ by (simp add: closedin_closed_subtopology eq) (auto simp: closedin_def) qed qed lemma quotient_map_saturated_open: "quotient_map X Y f \ continuous_map X Y f \ f ` (topspace X) = topspace Y \ (\U. openin X U \ {x \ topspace X. f x \ f ` U} \ U \ openin Y (f ` U))" (is "?lhs = ?rhs") proof assume L: ?lhs then have fim: "f ` topspace X = topspace Y" and Y: "\U. U \ topspace Y \ openin Y U = openin X {x \ topspace X. f x \ U}" unfolding quotient_map_def by auto show ?rhs proof (intro conjI allI impI) show "continuous_map X Y f" by (simp add: L quotient_imp_continuous_map) show "f ` topspace X = topspace Y" by (simp add: fim) next fix U :: "'a set" assume U: "openin X U \ {x \ topspace X. f x \ f ` U} \ U" then have sub: "f ` U \ topspace Y" and eq: "{x \ topspace X. f x \ f ` U} = U" using fim openin_subset by fastforce+ show "openin Y (f ` U)" by (simp add: sub Y eq U) qed next assume ?rhs then have YX: "\U. openin Y U \ openin X {x \ topspace X. f x \ U}" and fim: "f ` topspace X = topspace Y" and XY: "\U. \openin X U; {x \ topspace X. f x \ f ` U} \ U\ \ openin Y (f ` U)" by (auto simp: quotient_map_def continuous_map_def) show ?lhs proof (simp add: quotient_map_def fim, intro allI impI iffI) fix U :: "'b set" assume "U \ topspace Y" and X: "openin X {x \ topspace X. f x \ U}" have feq: "f ` {x \ topspace X. f x \ U} = U" using \U \ topspace Y\ fim by auto show "openin Y U" using XY [OF X] by (simp add: feq) next fix U :: "'b set" assume "U \ topspace Y" and Y: "openin Y U" show "openin X {x \ topspace X. f x \ U}" by (metis YX [OF Y]) qed qed subsection\ Separated Sets\ definition separatedin :: "'a topology \ 'a set \ 'a set \ bool" where "separatedin X S T \ S \ topspace X \ T \ topspace X \ S \ X closure_of T = {} \ T \ X closure_of S = {}" lemma separatedin_empty [simp]: "separatedin X S {} \ S \ topspace X" "separatedin X {} S \ S \ topspace X" by (simp_all add: separatedin_def) lemma separatedin_refl [simp]: "separatedin X S S \ S = {}" proof - have "\x. \separatedin X S S; x \ S\ \ False" by (metis all_not_in_conv closure_of_subset inf.orderE separatedin_def) then show ?thesis by auto qed lemma separatedin_sym: "separatedin X S T \ separatedin X T S" by (auto simp: separatedin_def) lemma separatedin_imp_disjoint: "separatedin X S T \ disjnt S T" by (meson closure_of_subset disjnt_def disjnt_subset2 separatedin_def) lemma separatedin_mono: "\separatedin X S T; S' \ S; T' \ T\ \ separatedin X S' T'" unfolding separatedin_def using closure_of_mono by blast lemma separatedin_open_sets: "\openin X S; openin X T\ \ separatedin X S T \ disjnt S T" unfolding disjnt_def separatedin_def by (auto simp: openin_Int_closure_of_eq_empty openin_subset) lemma separatedin_closed_sets: "\closedin X S; closedin X T\ \ separatedin X S T \ disjnt S T" - by (metis closedin_def closure_of_eq disjnt_def inf_commute separatedin_def) + unfolding closure_of_eq disjnt_def separatedin_def + by (metis closedin_def closure_of_eq inf_commute) lemma separatedin_subtopology: "separatedin (subtopology X U) S T \ S \ U \ T \ U \ separatedin X S T" apply (simp add: separatedin_def closure_of_subtopology) apply (safe; metis Int_absorb1 inf.assoc inf.orderE insert_disjoint(2) mk_disjoint_insert) done lemma separatedin_discrete_topology: "separatedin (discrete_topology U) S T \ S \ U \ T \ U \ disjnt S T" by (metis openin_discrete_topology separatedin_def separatedin_open_sets topspace_discrete_topology) lemma separated_eq_distinguishable: "separatedin X {x} {y} \ x \ topspace X \ y \ topspace X \ (\U. openin X U \ x \ U \ (y \ U)) \ (\v. openin X v \ y \ v \ (x \ v))" by (force simp: separatedin_def closure_of_def) lemma separatedin_Un [simp]: "separatedin X S (T \ U) \ separatedin X S T \ separatedin X S U" "separatedin X (S \ T) U \ separatedin X S U \ separatedin X T U" by (auto simp: separatedin_def) lemma separatedin_Union: "finite \ \ separatedin X S (\\) \ S \ topspace X \ (\T \ \. separatedin X S T)" "finite \ \ separatedin X (\\) S \ (\T \ \. separatedin X S T) \ S \ topspace X" by (auto simp: separatedin_def closure_of_Union) lemma separatedin_openin_diff: "\openin X S; openin X T\ \ separatedin X (S - T) (T - S)" unfolding separatedin_def apply (intro conjI) apply (meson Diff_subset openin_subset subset_trans)+ using openin_Int_closure_of_eq_empty by fastforce+ lemma separatedin_closedin_diff: "\closedin X S; closedin X T\ \ separatedin X (S - T) (T - S)" apply (simp add: separatedin_def Diff_Int_distrib2 closure_of_minimal inf_absorb2) apply (meson Diff_subset closedin_subset subset_trans) done lemma separation_closedin_Un_gen: "separatedin X S T \ S \ topspace X \ T \ topspace X \ disjnt S T \ closedin (subtopology X (S \ T)) S \ closedin (subtopology X (S \ T)) T" apply (simp add: separatedin_def closedin_Int_closure_of disjnt_iff) using closure_of_subset apply blast done lemma separation_openin_Un_gen: "separatedin X S T \ S \ topspace X \ T \ topspace X \ disjnt S T \ openin (subtopology X (S \ T)) S \ openin (subtopology X (S \ T)) T" unfolding openin_closedin_eq topspace_subtopology separation_closedin_Un_gen disjnt_def by (auto simp: Diff_triv Int_commute Un_Diff inf_absorb1 topspace_def) subsection\Homeomorphisms\ text\(1-way and 2-way versions may be useful in places)\ definition homeomorphic_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "homeomorphic_map X Y f \ quotient_map X Y f \ inj_on f (topspace X)" definition homeomorphic_maps :: "'a topology \ 'b topology \ ('a \ 'b) \ ('b \ 'a) \ bool" where "homeomorphic_maps X Y f g \ continuous_map X Y f \ continuous_map Y X g \ (\x \ topspace X. g(f x) = x) \ (\y \ topspace Y. f(g y) = y)" lemma homeomorphic_map_eq: "\homeomorphic_map X Y f; \x. x \ topspace X \ f x = g x\ \ homeomorphic_map X Y g" by (meson homeomorphic_map_def inj_on_cong quotient_map_eq) lemma homeomorphic_maps_eq: "\homeomorphic_maps X Y f g; \x. x \ topspace X \ f x = f' x; \y. y \ topspace Y \ g y = g' y\ \ homeomorphic_maps X Y f' g'" apply (simp add: homeomorphic_maps_def) by (metis continuous_map_eq continuous_map_eq_image_closure_subset_gen image_subset_iff) lemma homeomorphic_maps_sym: "homeomorphic_maps X Y f g \ homeomorphic_maps Y X g f" by (auto simp: homeomorphic_maps_def) lemma homeomorphic_maps_id: "homeomorphic_maps X Y id id \ Y = X" (is "?lhs = ?rhs") proof assume L: ?lhs then have "topspace X = topspace Y" by (auto simp: homeomorphic_maps_def continuous_map_def) with L show ?rhs unfolding homeomorphic_maps_def by (metis topology_finer_continuous_id topology_eq) next assume ?rhs then show ?lhs unfolding homeomorphic_maps_def by auto qed lemma homeomorphic_map_id [simp]: "homeomorphic_map X Y id \ Y = X" (is "?lhs = ?rhs") proof assume L: ?lhs then have eq: "topspace X = topspace Y" by (auto simp: homeomorphic_map_def continuous_map_def quotient_map_def) then have "\S. openin X S \ openin Y S" by (meson L homeomorphic_map_def injective_quotient_map topology_finer_open_id) then show ?rhs using L unfolding homeomorphic_map_def by (metis eq quotient_imp_continuous_map topology_eq topology_finer_continuous_id) next assume ?rhs then show ?lhs unfolding homeomorphic_map_def by (simp add: closed_map_id continuous_closed_imp_quotient_map) qed lemma homeomorphic_map_compose: assumes "homeomorphic_map X Y f" "homeomorphic_map Y X'' g" shows "homeomorphic_map X X'' (g \ f)" proof - have "inj_on g (f ` topspace X)" by (metis (no_types) assms homeomorphic_map_def quotient_imp_surjective_map) then show ?thesis using assms by (meson comp_inj_on homeomorphic_map_def quotient_map_compose_eq) qed lemma homeomorphic_maps_compose: "homeomorphic_maps X Y f h \ homeomorphic_maps Y X'' g k \ homeomorphic_maps X X'' (g \ f) (h \ k)" unfolding homeomorphic_maps_def by (auto simp: continuous_map_compose; simp add: continuous_map_def) lemma homeomorphic_eq_everything_map: "homeomorphic_map X Y f \ continuous_map X Y f \ open_map X Y f \ closed_map X Y f \ f ` (topspace X) = topspace Y \ inj_on f (topspace X)" unfolding homeomorphic_map_def by (force simp: injective_quotient_map intro: injective_quotient_map) lemma homeomorphic_imp_continuous_map: "homeomorphic_map X Y f \ continuous_map X Y f" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_open_map: "homeomorphic_map X Y f \ open_map X Y f" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_closed_map: "homeomorphic_map X Y f \ closed_map X Y f" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_surjective_map: "homeomorphic_map X Y f \ f ` (topspace X) = topspace Y" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_injective_map: "homeomorphic_map X Y f \ inj_on f (topspace X)" by (simp add: homeomorphic_eq_everything_map) lemma bijective_open_imp_homeomorphic_map: "\continuous_map X Y f; open_map X Y f; f ` (topspace X) = topspace Y; inj_on f (topspace X)\ \ homeomorphic_map X Y f" by (simp add: homeomorphic_map_def continuous_open_imp_quotient_map) lemma bijective_closed_imp_homeomorphic_map: "\continuous_map X Y f; closed_map X Y f; f ` (topspace X) = topspace Y; inj_on f (topspace X)\ \ homeomorphic_map X Y f" by (simp add: continuous_closed_quotient_map homeomorphic_map_def) lemma open_eq_continuous_inverse_map: assumes X: "\x. x \ topspace X \ f x \ topspace Y \ g(f x) = x" and Y: "\y. y \ topspace Y \ g y \ topspace X \ f(g y) = y" shows "open_map X Y f \ continuous_map Y X g" proof - have eq: "{x \ topspace Y. g x \ U} = f ` U" if "openin X U" for U using openin_subset [OF that] by (force simp: X Y image_iff) show ?thesis by (auto simp: Y open_map_def continuous_map_def eq) qed lemma closed_eq_continuous_inverse_map: assumes X: "\x. x \ topspace X \ f x \ topspace Y \ g(f x) = x" and Y: "\y. y \ topspace Y \ g y \ topspace X \ f(g y) = y" shows "closed_map X Y f \ continuous_map Y X g" proof - have eq: "{x \ topspace Y. g x \ U} = f ` U" if "closedin X U" for U using closedin_subset [OF that] by (force simp: X Y image_iff) show ?thesis by (auto simp: Y closed_map_def continuous_map_closedin eq) qed lemma homeomorphic_maps_map: "homeomorphic_maps X Y f g \ homeomorphic_map X Y f \ homeomorphic_map Y X g \ (\x \ topspace X. g(f x) = x) \ (\y \ topspace Y. f(g y) = y)" (is "?lhs = ?rhs") proof assume ?lhs then have L: "continuous_map X Y f" "continuous_map Y X g" "\x\topspace X. g (f x) = x" "\x'\topspace Y. f (g x') = x'" by (auto simp: homeomorphic_maps_def) show ?rhs proof (intro conjI bijective_open_imp_homeomorphic_map L) show "open_map X Y f" using L using open_eq_continuous_inverse_map [of concl: X Y f g] by (simp add: continuous_map_def) show "open_map Y X g" using L using open_eq_continuous_inverse_map [of concl: Y X g f] by (simp add: continuous_map_def) show "f ` topspace X = topspace Y" "g ` topspace Y = topspace X" using L by (force simp: continuous_map_closedin)+ show "inj_on f (topspace X)" "inj_on g (topspace Y)" using L unfolding inj_on_def by metis+ qed next assume ?rhs then show ?lhs by (auto simp: homeomorphic_maps_def homeomorphic_imp_continuous_map) qed lemma homeomorphic_maps_imp_map: "homeomorphic_maps X Y f g \ homeomorphic_map X Y f" using homeomorphic_maps_map by blast lemma homeomorphic_map_maps: "homeomorphic_map X Y f \ (\g. homeomorphic_maps X Y f g)" (is "?lhs = ?rhs") proof assume ?lhs then have L: "continuous_map X Y f" "open_map X Y f" "closed_map X Y f" "f ` (topspace X) = topspace Y" "inj_on f (topspace X)" by (auto simp: homeomorphic_eq_everything_map) have X: "\x. x \ topspace X \ f x \ topspace Y \ inv_into (topspace X) f (f x) = x" using L by auto have Y: "\y. y \ topspace Y \ inv_into (topspace X) f y \ topspace X \ f (inv_into (topspace X) f y) = y" by (simp add: L f_inv_into_f inv_into_into) have "homeomorphic_maps X Y f (inv_into (topspace X) f)" unfolding homeomorphic_maps_def proof (intro conjI L) show "continuous_map Y X (inv_into (topspace X) f)" by (simp add: L X Y flip: open_eq_continuous_inverse_map [where f=f]) next show "\x\topspace X. inv_into (topspace X) f (f x) = x" "\y\topspace Y. f (inv_into (topspace X) f y) = y" using X Y by auto qed then show ?rhs by metis next assume ?rhs then show ?lhs using homeomorphic_maps_map by blast qed lemma homeomorphic_maps_involution: "\continuous_map X X f; \x. x \ topspace X \ f(f x) = x\ \ homeomorphic_maps X X f f" by (auto simp: homeomorphic_maps_def) lemma homeomorphic_map_involution: "\continuous_map X X f; \x. x \ topspace X \ f(f x) = x\ \ homeomorphic_map X X f" using homeomorphic_maps_involution homeomorphic_maps_map by blast lemma homeomorphic_map_openness: assumes hom: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "openin Y (f ` U) \ openin X U" proof - obtain g where "homeomorphic_maps X Y f g" using assms by (auto simp: homeomorphic_map_maps) then have g: "homeomorphic_map Y X g" and gf: "\x. x \ topspace X \ g(f x) = x" by (auto simp: homeomorphic_maps_map) then have "openin X U \ openin Y (f ` U)" using hom homeomorphic_imp_open_map open_map_def by blast show "openin Y (f ` U) = openin X U" proof assume L: "openin Y (f ` U)" have "U = g ` (f ` U)" using U gf by force then show "openin X U" by (metis L homeomorphic_imp_open_map open_map_def g) next assume "openin X U" then show "openin Y (f ` U)" using hom homeomorphic_imp_open_map open_map_def by blast qed qed lemma homeomorphic_map_closedness: assumes hom: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "closedin Y (f ` U) \ closedin X U" proof - obtain g where "homeomorphic_maps X Y f g" using assms by (auto simp: homeomorphic_map_maps) then have g: "homeomorphic_map Y X g" and gf: "\x. x \ topspace X \ g(f x) = x" by (auto simp: homeomorphic_maps_map) then have "closedin X U \ closedin Y (f ` U)" using hom homeomorphic_imp_closed_map closed_map_def by blast show "closedin Y (f ` U) = closedin X U" proof assume L: "closedin Y (f ` U)" have "U = g ` (f ` U)" using U gf by force then show "closedin X U" by (metis L homeomorphic_imp_closed_map closed_map_def g) next assume "closedin X U" then show "closedin Y (f ` U)" using hom homeomorphic_imp_closed_map closed_map_def by blast qed qed lemma homeomorphic_map_openness_eq: "homeomorphic_map X Y f \ openin X U \ U \ topspace X \ openin Y (f ` U)" by (meson homeomorphic_map_openness openin_closedin_eq) lemma homeomorphic_map_closedness_eq: "homeomorphic_map X Y f \ closedin X U \ U \ topspace X \ closedin Y (f ` U)" by (meson closedin_subset homeomorphic_map_closedness) lemma all_openin_homeomorphic_image: assumes "homeomorphic_map X Y f" shows "(\V. openin Y V \ P V) \ (\U. openin X U \ P(f ` U))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (meson assms homeomorphic_map_openness_eq) next assume ?rhs then show ?lhs by (metis (no_types, lifting) assms homeomorphic_imp_surjective_map homeomorphic_map_openness openin_subset subset_image_iff) qed lemma all_closedin_homeomorphic_image: assumes "homeomorphic_map X Y f" shows "(\V. closedin Y V \ P V) \ (\U. closedin X U \ P(f ` U))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (meson assms homeomorphic_map_closedness_eq) next assume ?rhs then show ?lhs by (metis (no_types, lifting) assms homeomorphic_imp_surjective_map homeomorphic_map_closedness closedin_subset subset_image_iff) qed lemma homeomorphic_map_derived_set_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y derived_set_of (f ` S) = f ` (X derived_set_of S)" proof - have fim: "f ` (topspace X) = topspace Y" and inj: "inj_on f (topspace X)" using hom by (auto simp: homeomorphic_eq_everything_map) have iff: "(\T. x \ T \ openin X T \ (\y. y \ x \ y \ S \ y \ T)) = (\T. T \ topspace Y \ f x \ T \ openin Y T \ (\y. y \ f x \ y \ f ` S \ y \ T))" if "x \ topspace X" for x proof - have 1: "(x \ T \ openin X T) = (T \ topspace X \ f x \ f ` T \ openin Y (f ` T))" for T by (meson hom homeomorphic_map_openness_eq inj inj_on_image_mem_iff that) have 2: "(\y. y \ x \ y \ S \ y \ T) = (\y. y \ f x \ y \ f ` S \ y \ f ` T)" (is "?lhs = ?rhs") if "T \ topspace X \ f x \ f ` T \ openin Y (f ` T)" for T proof show "?lhs \ ?rhs" by (meson "1" imageI inj inj_on_eq_iff inj_on_subset that) show "?rhs \ ?lhs" using S inj inj_onD that by fastforce qed show ?thesis apply (simp flip: fim add: all_subset_image) apply (simp flip: imp_conjL) by (intro all_cong1 imp_cong 1 2) qed have *: "\T = f ` S; \x. x \ S \ P x \ Q(f x)\ \ {y. y \ T \ Q y} = f ` {x \ S. P x}" for T S P Q by auto show ?thesis unfolding derived_set_of_def apply (rule *) using fim apply blast using iff openin_subset by force qed lemma homeomorphic_map_closure_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y closure_of (f ` S) = f ` (X closure_of S)" unfolding closure_of using homeomorphic_imp_surjective_map [OF hom] S by (auto simp: in_derived_set_of homeomorphic_map_derived_set_of [OF assms]) lemma homeomorphic_map_interior_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y interior_of (f ` S) = f ` (X interior_of S)" proof - { fix y assume "y \ topspace Y" and "y \ Y closure_of (topspace Y - f ` S)" then have "y \ f ` (topspace X - X closure_of (topspace X - S))" using homeomorphic_eq_everything_map [THEN iffD1, OF hom] homeomorphic_map_closure_of [OF hom] by (metis DiffI Diff_subset S closure_of_subset_topspace inj_on_image_set_diff) } moreover { fix x assume "x \ topspace X" then have "f x \ topspace Y" using hom homeomorphic_imp_surjective_map by blast } moreover { fix x assume "x \ topspace X" and "x \ X closure_of (topspace X - S)" and "f x \ Y closure_of (topspace Y - f ` S)" then have "False" using homeomorphic_map_closure_of [OF hom] hom unfolding homeomorphic_eq_everything_map by (metis (no_types, lifting) Diff_subset S closure_of_subset_topspace inj_on_image_mem_iff_alt inj_on_image_set_diff) } ultimately show ?thesis by (auto simp: interior_of_closure_of) qed lemma homeomorphic_map_frontier_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y frontier_of (f ` S) = f ` (X frontier_of S)" unfolding frontier_of_def proof (intro equalityI subsetI DiffI) fix y assume "y \ Y closure_of f ` S - Y interior_of f ` S" then show "y \ f ` (X closure_of S - X interior_of S)" using S hom homeomorphic_map_closure_of homeomorphic_map_interior_of by fastforce next fix y assume "y \ f ` (X closure_of S - X interior_of S)" then show "y \ Y closure_of f ` S" using S hom homeomorphic_map_closure_of by fastforce next fix x assume "x \ f ` (X closure_of S - X interior_of S)" then obtain y where y: "x = f y" "y \ X closure_of S" "y \ X interior_of S" by blast then have "y \ topspace X" by (simp add: in_closure_of) then have "f y \ f ` (X interior_of S)" by (meson hom homeomorphic_eq_everything_map inj_on_image_mem_iff_alt interior_of_subset_topspace y(3)) then show "x \ Y interior_of f ` S" using S hom homeomorphic_map_interior_of y(1) by blast qed lemma homeomorphic_maps_subtopologies: "\homeomorphic_maps X Y f g; f ` (topspace X \ S) = topspace Y \ T\ \ homeomorphic_maps (subtopology X S) (subtopology Y T) f g" unfolding homeomorphic_maps_def by (force simp: continuous_map_from_subtopology continuous_map_in_subtopology) lemma homeomorphic_maps_subtopologies_alt: "\homeomorphic_maps X Y f g; f ` (topspace X \ S) \ T; g ` (topspace Y \ T) \ S\ \ homeomorphic_maps (subtopology X S) (subtopology Y T) f g" unfolding homeomorphic_maps_def by (force simp: continuous_map_from_subtopology continuous_map_in_subtopology) lemma homeomorphic_map_subtopologies: "\homeomorphic_map X Y f; f ` (topspace X \ S) = topspace Y \ T\ \ homeomorphic_map (subtopology X S) (subtopology Y T) f" by (meson homeomorphic_map_maps homeomorphic_maps_subtopologies) lemma homeomorphic_map_subtopologies_alt: "\homeomorphic_map X Y f; \x. \x \ topspace X; f x \ topspace Y\ \ f x \ T \ x \ S\ \ homeomorphic_map (subtopology X S) (subtopology Y T) f" unfolding homeomorphic_map_maps apply (erule ex_forward) apply (rule homeomorphic_maps_subtopologies) apply (auto simp: homeomorphic_maps_def continuous_map_def) by (metis IntI image_iff) subsection\Relation of homeomorphism between topological spaces\ definition homeomorphic_space (infixr "homeomorphic'_space" 50) where "X homeomorphic_space Y \ \f g. homeomorphic_maps X Y f g" lemma homeomorphic_space_refl: "X homeomorphic_space X" by (meson homeomorphic_maps_id homeomorphic_space_def) lemma homeomorphic_space_sym: "X homeomorphic_space Y \ Y homeomorphic_space X" unfolding homeomorphic_space_def by (metis homeomorphic_maps_sym) lemma homeomorphic_space_trans [trans]: "\X1 homeomorphic_space X2; X2 homeomorphic_space X3\ \ X1 homeomorphic_space X3" unfolding homeomorphic_space_def by (metis homeomorphic_maps_compose) lemma homeomorphic_space: "X homeomorphic_space Y \ (\f. homeomorphic_map X Y f)" by (simp add: homeomorphic_map_maps homeomorphic_space_def) lemma homeomorphic_maps_imp_homeomorphic_space: "homeomorphic_maps X Y f g \ X homeomorphic_space Y" unfolding homeomorphic_space_def by metis lemma homeomorphic_map_imp_homeomorphic_space: "homeomorphic_map X Y f \ X homeomorphic_space Y" unfolding homeomorphic_map_maps using homeomorphic_space_def by blast lemma homeomorphic_empty_space: "X homeomorphic_space Y \ topspace X = {} \ topspace Y = {}" by (metis homeomorphic_imp_surjective_map homeomorphic_space image_is_empty) lemma homeomorphic_empty_space_eq: assumes "topspace X = {}" shows "X homeomorphic_space Y \ topspace Y = {}" proof - have "\f t. continuous_map X (t::'b topology) f" using assms continuous_map_on_empty by blast then show ?thesis by (metis (no_types) assms continuous_map_on_empty empty_iff homeomorphic_empty_space homeomorphic_maps_def homeomorphic_space_def) qed subsection\Connected topological spaces\ definition connected_space :: "'a topology \ bool" where "connected_space X \ \(\E1 E2. openin X E1 \ openin X E2 \ topspace X \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" definition connectedin :: "'a topology \ 'a set \ bool" where "connectedin X S \ S \ topspace X \ connected_space (subtopology X S)" lemma connected_spaceD: "\connected_space X; openin X U; openin X V; topspace X \ U \ V; U \ V = {}; U \ {}; V \ {}\ \ False" by (auto simp: connected_space_def) lemma connectedin_subset_topspace: "connectedin X S \ S \ topspace X" by (simp add: connectedin_def) lemma connectedin_topspace: "connectedin X (topspace X) \ connected_space X" by (simp add: connectedin_def) lemma connected_space_subtopology: "connectedin X S \ connected_space (subtopology X S)" by (simp add: connectedin_def) lemma connectedin_subtopology: "connectedin (subtopology X S) T \ connectedin X T \ T \ S" by (force simp: connectedin_def subtopology_subtopology inf_absorb2) lemma connected_space_eq: "connected_space X \ (\E1 E2. openin X E1 \ openin X E2 \ E1 \ E2 = topspace X \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" unfolding connected_space_def by (metis openin_Un openin_subset subset_antisym) lemma connected_space_closedin: "connected_space X \ (\E1 E2. closedin X E1 \ closedin X E2 \ topspace X \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" (is "?lhs = ?rhs") proof assume ?lhs then have L: "\E1 E2. \openin X E1; E1 \ E2 = {}; topspace X \ E1 \ E2; openin X E2\ \ E1 = {} \ E2 = {}" by (simp add: connected_space_def) show ?rhs unfolding connected_space_def proof clarify fix E1 E2 assume "closedin X E1" and "closedin X E2" and "topspace X \ E1 \ E2" and "E1 \ E2 = {}" and "E1 \ {}" and "E2 \ {}" have "E1 \ E2 = topspace X" by (meson Un_subset_iff \closedin X E1\ \closedin X E2\ \topspace X \ E1 \ E2\ closedin_def subset_antisym) then have "topspace X - E2 = E1" using \E1 \ E2 = {}\ by fastforce then have "topspace X = E1" using \E1 \ {}\ L \closedin X E1\ \closedin X E2\ by blast then show "False" using \E1 \ E2 = {}\ \E1 \ E2 = topspace X\ \E2 \ {}\ by blast qed next assume R: ?rhs show ?lhs unfolding connected_space_def proof clarify fix E1 E2 assume "openin X E1" and "openin X E2" and "topspace X \ E1 \ E2" and "E1 \ E2 = {}" and "E1 \ {}" and "E2 \ {}" have "E1 \ E2 = topspace X" by (meson Un_subset_iff \openin X E1\ \openin X E2\ \topspace X \ E1 \ E2\ openin_closedin_eq subset_antisym) then have "topspace X - E2 = E1" using \E1 \ E2 = {}\ by fastforce then have "topspace X = E1" using \E1 \ {}\ R \openin X E1\ \openin X E2\ by blast then show "False" using \E1 \ E2 = {}\ \E1 \ E2 = topspace X\ \E2 \ {}\ by blast qed qed lemma connected_space_closedin_eq: "connected_space X \ (\E1 E2. closedin X E1 \ closedin X E2 \ E1 \ E2 = topspace X \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" apply (simp add: connected_space_closedin) apply (intro all_cong) using closedin_subset apply blast done lemma connected_space_clopen_in: "connected_space X \ (\T. openin X T \ closedin X T \ T = {} \ T = topspace X)" proof - have eq: "openin X E1 \ openin X E2 \ E1 \ E2 = topspace X \ E1 \ E2 = {} \ P \ E2 = topspace X - E1 \ openin X E1 \ openin X E2 \ P" for E1 E2 P using openin_subset by blast show ?thesis unfolding connected_space_eq eq closedin_def by (auto simp: openin_closedin_eq) qed lemma connectedin: "connectedin X S \ S \ topspace X \ (\E1 E2. openin X E1 \ openin X E2 \ S \ E1 \ E2 \ E1 \ E2 \ S = {} \ E1 \ S \ {} \ E2 \ S \ {})" proof - have *: "(\E1:: 'a set. \E2:: 'a set. (\T1:: 'a set. P1 T1 \ E1 = f1 T1) \ (\T2:: 'a set. P2 T2 \ E2 = f2 T2) \ R E1 E2) \ (\T1 T2. P1 T1 \ P2 T2 \ R(f1 T1) (f2 T2))" for P1 f1 P2 f2 R by auto show ?thesis unfolding connectedin_def connected_space_def openin_subtopology topspace_subtopology Not_eq_iff * apply (intro conj_cong arg_cong [where f=Not] ex_cong1 refl) apply (blast elim: dest!: openin_subset)+ done qed lemma connectedin_iff_connected [simp]: "connectedin euclidean S \ connected S" by (simp add: connected_def connectedin) lemma connectedin_closedin: "connectedin X S \ S \ topspace X \ \(\E1 E2. closedin X E1 \ closedin X E2 \ S \ (E1 \ E2) \ (E1 \ E2 \ S = {}) \ \(E1 \ S = {}) \ \(E2 \ S = {}))" proof - have *: "(\E1:: 'a set. \E2:: 'a set. (\T1:: 'a set. P1 T1 \ E1 = f1 T1) \ (\T2:: 'a set. P2 T2 \ E2 = f2 T2) \ R E1 E2) \ (\T1 T2. P1 T1 \ P2 T2 \ R(f1 T1) (f2 T2))" for P1 f1 P2 f2 R by auto show ?thesis unfolding connectedin_def connected_space_closedin closedin_subtopology topspace_subtopology Not_eq_iff * apply (intro conj_cong arg_cong [where f=Not] ex_cong1 refl) apply (blast elim: dest!: openin_subset)+ done qed lemma connectedin_empty [simp]: "connectedin X {}" by (simp add: connectedin) lemma connected_space_topspace_empty: "topspace X = {} \ connected_space X" using connectedin_topspace by fastforce lemma connectedin_sing [simp]: "connectedin X {a} \ a \ topspace X" by (simp add: connectedin) lemma connectedin_absolute [simp]: "connectedin (subtopology X S) S \ connectedin X S" apply (simp only: connectedin_def topspace_subtopology subtopology_subtopology) apply (intro conj_cong imp_cong arg_cong [where f=Not] all_cong1 ex_cong1 refl) by auto lemma connectedin_Union: assumes \: "\S. S \ \ \ connectedin X S" and ne: "\\ \ {}" shows "connectedin X (\\)" proof - have "\\ \ topspace X" using \ by (simp add: Union_least connectedin_def) moreover have False if "openin X E1" "openin X E2" and cover: "\\ \ E1 \ E2" and disj: "E1 \ E2 \ \\ = {}" and overlap1: "E1 \ \\ \ {}" and overlap2: "E2 \ \\ \ {}" for E1 E2 proof - have disjS: "E1 \ E2 \ S = {}" if "S \ \" for S using Diff_triv that disj by auto have coverS: "S \ E1 \ E2" if "S \ \" for S using that cover by blast have "\ \ {}" using overlap1 by blast obtain a where a: "\U. U \ \ \ a \ U" using ne by force with \\ \ {}\ have "a \ \\" by blast then consider "a \ E1" | "a \ E2" using \\\ \ E1 \ E2\ by auto then show False proof cases case 1 then obtain b S where "b \ E2" "b \ S" "S \ \" using overlap2 by blast then show ?thesis using "1" \openin X E1\ \openin X E2\ disjS coverS a [OF \S \ \\] \[OF \S \ \\] unfolding connectedin by (meson disjoint_iff_not_equal) next case 2 then obtain b S where "b \ E1" "b \ S" "S \ \" using overlap1 by blast then show ?thesis using "2" \openin X E1\ \openin X E2\ disjS coverS a [OF \S \ \\] \[OF \S \ \\] unfolding connectedin by (meson disjoint_iff_not_equal) qed qed ultimately show ?thesis unfolding connectedin by blast qed lemma connectedin_Un: "\connectedin X S; connectedin X T; S \ T \ {}\ \ connectedin X (S \ T)" using connectedin_Union [of "{S,T}"] by auto lemma connected_space_subconnected: "connected_space X \ (\x \ topspace X. \y \ topspace X. \S. connectedin X S \ x \ S \ y \ S)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using connectedin_topspace by blast next assume R [rule_format]: ?rhs have False if "openin X U" "openin X V" and disj: "U \ V = {}" and cover: "topspace X \ U \ V" and "U \ {}" "V \ {}" for U V proof - obtain u v where "u \ U" "v \ V" using \U \ {}\ \V \ {}\ by auto then obtain T where "u \ T" "v \ T" and T: "connectedin X T" using R [of u v] that by (meson \openin X U\ \openin X V\ subsetD openin_subset) then show False using that unfolding connectedin by (metis IntI \u \ U\ \v \ V\ empty_iff inf_bot_left subset_trans) qed then show ?lhs by (auto simp: connected_space_def) qed lemma connectedin_intermediate_closure_of: assumes "connectedin X S" "S \ T" "T \ X closure_of S" shows "connectedin X T" proof - have S: "S \ topspace X"and T: "T \ topspace X" using assms by (meson closure_of_subset_topspace dual_order.trans)+ show ?thesis using assms apply (simp add: connectedin closure_of_subset_topspace S T) apply (elim all_forward imp_forward2 asm_rl) apply (blast dest: openin_Int_closure_of_eq_empty [of X _ S])+ done qed lemma connectedin_closure_of: "connectedin X S \ connectedin X (X closure_of S)" by (meson closure_of_subset connectedin_def connectedin_intermediate_closure_of subset_refl) lemma connectedin_separation: "connectedin X S \ S \ topspace X \ (\C1 C2. C1 \ C2 = S \ C1 \ {} \ C2 \ {} \ C1 \ X closure_of C2 = {} \ C2 \ X closure_of C1 = {})" (is "?lhs = ?rhs") unfolding connectedin_def connected_space_closedin_eq closedin_Int_closure_of topspace_subtopology apply (intro conj_cong refl arg_cong [where f=Not]) apply (intro ex_cong1 iffI, blast) using closure_of_subset_Int by force lemma connectedin_eq_not_separated: "connectedin X S \ S \ topspace X \ (\C1 C2. C1 \ C2 = S \ C1 \ {} \ C2 \ {} \ separatedin X C1 C2)" apply (simp add: separatedin_def connectedin_separation) apply (intro conj_cong all_cong1 refl, blast) done lemma connectedin_eq_not_separated_subset: "connectedin X S \ S \ topspace X \ (\C1 C2. S \ C1 \ C2 \ S \ C1 \ {} \ S \ C2 \ {} \ separatedin X C1 C2)" proof - have *: "\C1 C2. S \ C1 \ C2 \ S \ C1 = {} \ S \ C2 = {} \ \ separatedin X C1 C2" if "\C1 C2. C1 \ C2 = S \ C1 = {} \ C2 = {} \ \ separatedin X C1 C2" proof (intro allI) fix C1 C2 show "S \ C1 \ C2 \ S \ C1 = {} \ S \ C2 = {} \ \ separatedin X C1 C2" using that [of "S \ C1" "S \ C2"] by (auto simp: separatedin_mono) qed show ?thesis apply (simp add: connectedin_eq_not_separated) apply (intro conj_cong refl iffI *) apply (blast elim!: all_forward)+ done qed lemma connected_space_eq_not_separated: "connected_space X \ (\C1 C2. C1 \ C2 = topspace X \ C1 \ {} \ C2 \ {} \ separatedin X C1 C2)" by (simp add: connectedin_eq_not_separated flip: connectedin_topspace) lemma connected_space_eq_not_separated_subset: "connected_space X \ (\C1 C2. topspace X \ C1 \ C2 \ C1 \ {} \ C2 \ {} \ separatedin X C1 C2)" apply (simp add: connected_space_eq_not_separated) apply (intro all_cong1) by (metis Un_absorb dual_order.antisym separatedin_def subset_refl sup_mono) lemma connectedin_subset_separated_union: "\connectedin X C; separatedin X S T; C \ S \ T\ \ C \ S \ C \ T" unfolding connectedin_eq_not_separated_subset by blast lemma connectedin_nonseparated_union: "\connectedin X S; connectedin X T; \separatedin X S T\ \ connectedin X (S \ T)" apply (simp add: connectedin_eq_not_separated_subset, auto) apply (metis (no_types, hide_lams) Diff_subset_conv Diff_triv disjoint_iff_not_equal separatedin_mono sup_commute) apply (metis (no_types, hide_lams) Diff_subset_conv Diff_triv disjoint_iff_not_equal separatedin_mono separatedin_sym sup_commute) by (meson disjoint_iff_not_equal) lemma connected_space_closures: "connected_space X \ (\e1 e2. e1 \ e2 = topspace X \ X closure_of e1 \ X closure_of e2 = {} \ e1 \ {} \ e2 \ {})" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding connected_space_closedin_eq by (metis Un_upper1 Un_upper2 closedin_closure_of closure_of_Un closure_of_eq_empty closure_of_topspace) next assume ?rhs then show ?lhs unfolding connected_space_closedin_eq by (metis closure_of_eq) qed lemma connectedin_inter_frontier_of: assumes "connectedin X S" "S \ T \ {}" "S - T \ {}" shows "S \ X frontier_of T \ {}" proof - have "S \ topspace X" and *: "\E1 E2. openin X E1 \ openin X E2 \ E1 \ E2 \ S = {} \ S \ E1 \ E2 \ E1 \ S = {} \ E2 \ S = {}" using \connectedin X S\ by (auto simp: connectedin) have "S - (topspace X \ T) \ {}" using assms(3) by blast moreover have "S \ topspace X \ T \ {}" using assms(1) assms(2) connectedin by fastforce moreover have False if "S \ T \ {}" "S - T \ {}" "T \ topspace X" "S \ X frontier_of T = {}" for T proof - have null: "S \ (X closure_of T - X interior_of T) = {}" using that unfolding frontier_of_def by blast have 1: "X interior_of T \ (topspace X - X closure_of T) \ S = {}" by (metis Diff_disjoint inf_bot_left interior_of_Int interior_of_complement interior_of_empty) have 2: "S \ X interior_of T \ (topspace X - X closure_of T)" using that \S \ topspace X\ null by auto have 3: "S \ X interior_of T \ {}" using closure_of_subset that(1) that(3) null by fastforce show ?thesis using null \S \ topspace X\ that * [of "X interior_of T" "topspace X - X closure_of T"] apply (clarsimp simp add: openin_diff 1 2) apply (simp add: Int_commute Diff_Int_distrib 3) by (metis Int_absorb2 contra_subsetD interior_of_subset) qed ultimately show ?thesis by (metis Int_lower1 frontier_of_restrict inf_assoc) qed lemma connectedin_continuous_map_image: assumes f: "continuous_map X Y f" and "connectedin X S" shows "connectedin Y (f ` S)" proof - have "S \ topspace X" and *: "\E1 E2. openin X E1 \ openin X E2 \ E1 \ E2 \ S = {} \ S \ E1 \ E2 \ E1 \ S = {} \ E2 \ S = {}" using \connectedin X S\ by (auto simp: connectedin) show ?thesis unfolding connectedin connected_space_def proof (intro conjI notI; clarify) show "f x \ topspace Y" if "x \ S" for x using \S \ topspace X\ continuous_map_image_subset_topspace f that by blast next fix U V let ?U = "{x \ topspace X. f x \ U}" let ?V = "{x \ topspace X. f x \ V}" assume UV: "openin Y U" "openin Y V" "f ` S \ U \ V" "U \ V \ f ` S = {}" "U \ f ` S \ {}" "V \ f ` S \ {}" then have 1: "?U \ ?V \ S = {}" by auto have 2: "openin X ?U" "openin X ?V" using \openin Y U\ \openin Y V\ continuous_map f by fastforce+ show "False" using * [of ?U ?V] UV \S \ topspace X\ by (auto simp: 1 2) qed qed lemma homeomorphic_connected_space: "X homeomorphic_space Y \ connected_space X \ connected_space Y" unfolding homeomorphic_space_def homeomorphic_maps_def apply safe apply (metis connectedin_continuous_map_image connected_space_subconnected continuous_map_image_subset_topspace image_eqI image_subset_iff) by (metis (no_types, hide_lams) connectedin_continuous_map_image connectedin_topspace continuous_map_def continuous_map_image_subset_topspace imageI set_eq_subset subsetI) lemma homeomorphic_map_connectedness: assumes f: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "connectedin Y (f ` U) \ connectedin X U" proof - have 1: "f ` U \ topspace Y \ U \ topspace X" using U f homeomorphic_imp_surjective_map by blast moreover have "connected_space (subtopology Y (f ` U)) \ connected_space (subtopology X U)" proof (rule homeomorphic_connected_space) have "f ` U \ topspace Y" by (simp add: U 1) then have "topspace Y \ f ` U = f ` U" by (simp add: subset_antisym) then show "subtopology Y (f ` U) homeomorphic_space subtopology X U" by (metis (no_types) Int_subset_iff U f homeomorphic_map_imp_homeomorphic_space homeomorphic_map_subtopologies homeomorphic_space_sym subset_antisym subset_refl) qed ultimately show ?thesis by (auto simp: connectedin_def) qed lemma homeomorphic_map_connectedness_eq: "homeomorphic_map X Y f \ connectedin X U \ U \ topspace X \ connectedin Y (f ` U)" using homeomorphic_map_connectedness connectedin_subset_topspace by metis lemma connectedin_discrete_topology: "connectedin (discrete_topology U) S \ S \ U \ (\a. S \ {a})" proof (cases "S \ U") case True show ?thesis proof (cases "S = {}") case False moreover have "connectedin (discrete_topology U) S \ (\a. S = {a})" apply safe using False connectedin_inter_frontier_of insert_Diff apply fastforce using True by auto ultimately show ?thesis by auto qed simp next case False then show ?thesis by (simp add: connectedin_def) qed lemma connected_space_discrete_topology: "connected_space (discrete_topology U) \ (\a. U \ {a})" by (metis connectedin_discrete_topology connectedin_topspace order_refl topspace_discrete_topology) subsection\Compact sets\ definition compactin where "compactin X S \ S \ topspace X \ (\\. (\U \ \. openin X U) \ S \ \\ \ (\\. finite \ \ \ \ \ \ S \ \\))" definition compact_space where "compact_space X \ compactin X (topspace X)" lemma compact_space_alt: "compact_space X \ (\\. (\U \ \. openin X U) \ topspace X \ \\ \ (\\. finite \ \ \ \ \ \ topspace X \ \\))" by (simp add: compact_space_def compactin_def) lemma compact_space: "compact_space X \ (\\. (\U \ \. openin X U) \ \\ = topspace X \ (\\. finite \ \ \ \ \ \ \\ = topspace X))" unfolding compact_space_alt using openin_subset by fastforce lemma compactinD: "\compactin X S; \U. U \ \ \ openin X U; S \ \\\ \ \\. finite \ \ \ \ \ \ S \ \\" by (auto simp: compactin_def) lemma compactin_euclidean_iff [simp]: "compactin euclidean S \ compact S" by (simp add: compact_eq_Heine_Borel compactin_def) meson lemma compactin_absolute [simp]: "compactin (subtopology X S) S \ compactin X S" proof - have eq: "(\U \ \. \Y. openin X Y \ U = Y \ S) \ \ \ (\Y. Y \ S) ` {y. openin X y}" for \ by auto show ?thesis by (auto simp: compactin_def openin_subtopology eq imp_conjL all_subset_image ex_finite_subset_image) qed lemma compactin_subspace: "compactin X S \ S \ topspace X \ compact_space (subtopology X S)" unfolding compact_space_def topspace_subtopology by (metis compactin_absolute compactin_def inf.absorb2) lemma compact_space_subtopology: "compactin X S \ compact_space (subtopology X S)" by (simp add: compactin_subspace) lemma compactin_subtopology: "compactin (subtopology X S) T \ compactin X T \ T \ S" apply (simp add: compactin_subspace) by (metis inf.orderE inf_commute subtopology_subtopology) lemma compactin_subset_topspace: "compactin X S \ S \ topspace X" by (simp add: compactin_subspace) lemma compactin_contractive: "\compactin X' S; topspace X' = topspace X; \U. openin X U \ openin X' U\ \ compactin X S" by (simp add: compactin_def) lemma finite_imp_compactin: "\S \ topspace X; finite S\ \ compactin X S" by (metis compactin_subspace compact_space finite_UnionD inf.absorb_iff2 order_refl topspace_subtopology) lemma compactin_empty [iff]: "compactin X {}" by (simp add: finite_imp_compactin) lemma compact_space_topspace_empty: "topspace X = {} \ compact_space X" by (simp add: compact_space_def) lemma finite_imp_compactin_eq: "finite S \ (compactin X S \ S \ topspace X)" using compactin_subset_topspace finite_imp_compactin by blast lemma compactin_sing [simp]: "compactin X {a} \ a \ topspace X" by (simp add: finite_imp_compactin_eq) lemma closed_compactin: assumes XK: "compactin X K" and "C \ K" and XC: "closedin X C" shows "compactin X C" unfolding compactin_def proof (intro conjI allI impI) show "C \ topspace X" by (simp add: XC closedin_subset) next fix \ :: "'a set set" assume \: "Ball \ (openin X) \ C \ \\" have "(\U\insert (topspace X - C) \. openin X U)" using XC \ by blast moreover have "K \ \(insert (topspace X - C) \)" using \ XK compactin_subset_topspace by fastforce ultimately obtain \ where "finite \" "\ \ insert (topspace X - C) \" "K \ \\" using assms unfolding compactin_def by metis moreover have "openin X (topspace X - C)" using XC by auto ultimately show "\\. finite \ \ \ \ \ \ C \ \\" using \C \ K\ by (rule_tac x="\ - {topspace X - C}" in exI) auto qed lemma closedin_compact_space: "\compact_space X; closedin X S\ \ compactin X S" by (simp add: closed_compactin closedin_subset compact_space_def) lemma compact_Int_closedin: assumes "compactin X S" "closedin X T" shows "compactin X (S \ T)" proof - have "compactin (subtopology X S) (S \ T)" by (metis assms closedin_compact_space closedin_subtopology compactin_subspace inf_commute) then show ?thesis by (simp add: compactin_subtopology) qed lemma closed_Int_compactin: "\closedin X S; compactin X T\ \ compactin X (S \ T)" by (metis compact_Int_closedin inf_commute) lemma compactin_Un: assumes S: "compactin X S" and T: "compactin X T" shows "compactin X (S \ T)" unfolding compactin_def proof (intro conjI allI impI) show "S \ T \ topspace X" using assms by (auto simp: compactin_def) next fix \ :: "'a set set" assume \: "Ball \ (openin X) \ S \ T \ \\" with S obtain \ where \: "finite \" "\ \ \" "S \ \\" unfolding compactin_def by (meson sup.bounded_iff) obtain \ where "finite \" "\ \ \" "T \ \\" using \ T unfolding compactin_def by (meson sup.bounded_iff) with \ show "\\. finite \ \ \ \ \ \ S \ T \ \\" by (rule_tac x="\ \ \" in exI) auto qed lemma compactin_Union: "\finite \; \S. S \ \ \ compactin X S\ \ compactin X (\\)" by (induction rule: finite_induct) (simp_all add: compactin_Un) lemma compactin_subtopology_imp_compact: assumes "compactin (subtopology X S) K" shows "compactin X K" using assms proof (clarsimp simp add: compactin_def) fix \ define \ where "\ \ (\U. U \ S) ` \" assume "K \ topspace X" and "K \ S" and "\x\\. openin X x" and "K \ \\" then have "\V \ \. openin (subtopology X S) V" "K \ \\" unfolding \_def by (auto simp: openin_subtopology) moreover assume "\\. (\x\\. openin (subtopology X S) x) \ K \ \\ \ (\\. finite \ \ \ \ \ \ K \ \\)" ultimately obtain \ where "finite \" "\ \ \" "K \ \\" by meson then have \: "\U. U \ \ \ V = U \ S" if "V \ \" for V unfolding \_def using that by blast let ?\ = "(\F. @U. U \ \ \ F = U \ S) ` \" show "\\. finite \ \ \ \ \ \ K \ \\" proof (intro exI conjI) show "finite ?\" using \finite \\ by blast show "?\ \ \" using someI_ex [OF \] by blast show "K \ \?\" proof clarsimp fix x assume "x \ K" then show "\V \ \. x \ (SOME U. U \ \ \ V = U \ S)" using \K \ \\\ someI_ex [OF \] by (metis (no_types, lifting) IntD1 Union_iff subsetCE) qed qed qed lemma compact_imp_compactin_subtopology: assumes "compactin X K" "K \ S" shows "compactin (subtopology X S) K" using assms proof (clarsimp simp add: compactin_def) fix \ :: "'a set set" define \ where "\ \ {V. openin X V \ (\U \ \. U = V \ S)}" assume "K \ S" and "K \ topspace X" and "\U\\. openin (subtopology X S) U" and "K \ \\" then have "\V \ \. openin X V" "K \ \\" unfolding \_def by (fastforce simp: subset_eq openin_subtopology)+ moreover assume "\\. (\U\\. openin X U) \ K \ \\ \ (\\. finite \ \ \ \ \ \ K \ \\)" ultimately obtain \ where "finite \" "\ \ \" "K \ \\" by meson let ?\ = "(\F. F \ S) ` \" show "\\. finite \ \ \ \ \ \ K \ \\" proof (intro exI conjI) show "finite ?\" using \finite \\ by blast show "?\ \ \" using \_def \\ \ \\ by blast show "K \ \?\" using \K \ \\\ assms(2) by auto qed qed proposition compact_space_fip: "compact_space X \ (\\. (\C\\. closedin X C) \ (\\. finite \ \ \ \ \ \ \\ \ {}) \ \\ \ {})" (is "_ = ?rhs") proof (cases "topspace X = {}") case True then show ?thesis apply (clarsimp simp add: compact_space_def closedin_topspace_empty) by (metis finite.emptyI finite_insert infinite_super insertI1 subsetI) next case False show ?thesis proof safe fix \ :: "'a set set" assume * [rule_format]: "\\. finite \ \ \ \ \ \ \\ \ {}" define \ where "\ \ (\S. topspace X - S) ` \" assume clo: "\C\\. closedin X C" and [simp]: "\\ = {}" then have "\V \ \. openin X V" "topspace X \ \\" by (auto simp: \_def) moreover assume [unfolded compact_space_alt, rule_format, of \]: "compact_space X" ultimately obtain \ where \: "finite \" "\ \ \" "topspace X \ topspace X - \\" by (auto simp: ex_finite_subset_image \_def) moreover have "\ \ {}" using \ \topspace X \ {}\ by blast ultimately show "False" using * [of \] by auto (metis Diff_iff Inter_iff clo closedin_def subsetD) next assume R [rule_format]: ?rhs show "compact_space X" unfolding compact_space_alt proof clarify fix \ :: "'a set set" define \ where "\ \ (\S. topspace X - S) ` \" assume "\C\\. openin X C" and "topspace X \ \\" with \topspace X \ {}\ have *: "\V \ \. closedin X V" "\ \ {}" by (auto simp: \_def) show "\\. finite \ \ \ \ \ \ topspace X \ \\" proof (rule ccontr; simp) assume "\\\\. finite \ \ \ topspace X \ \\" then have "\\. finite \ \ \ \ \ \ \\ \ {}" by (simp add: \_def all_finite_subset_image) with \topspace X \ \\\ show False using R [of \] * by (simp add: \_def) qed qed qed qed corollary compactin_fip: "compactin X S \ S \ topspace X \ (\\. (\C\\. closedin X C) \ (\\. finite \ \ \ \ \ \ S \ \\ \ {}) \ S \ \\ \ {})" proof (cases "S = {}") case False show ?thesis proof (cases "S \ topspace X") case True then have "compactin X S \ (\\. \ \ (\T. S \ T) ` {T. closedin X T} \ (\\. finite \ \ \ \ \ \ \\ \ {}) \ \\ \ {})" by (simp add: compact_space_fip compactin_subspace closedin_subtopology image_def subset_eq Int_commute imp_conjL) also have "\ = (\\\Collect (closedin X). (\\. finite \ \ \ \ (\) S ` \ \ \\ \ {}) \ \ ((\) S ` \) \ {})" by (simp add: all_subset_image) also have "\ = (\\. (\C\\. closedin X C) \ (\\. finite \ \ \ \ \ \ S \ \\ \ {}) \ S \ \\ \ {})" proof - have eq: "((\\. finite \ \ \ \ \ \ \ ((\) S ` \) \ {}) \ \ ((\) S ` \) \ {}) \ ((\\. finite \ \ \ \ \ \ S \ \\ \ {}) \ S \ \\ \ {})" for \ by simp (use \S \ {}\ in blast) show ?thesis apply (simp only: imp_conjL [symmetric] all_finite_subset_image eq) apply (simp add: subset_eq) done qed finally show ?thesis using True by simp qed (simp add: compactin_subspace) qed force corollary compact_space_imp_nest: fixes C :: "nat \ 'a set" assumes "compact_space X" and clo: "\n. closedin X (C n)" and ne: "\n. C n \ {}" and inc: "\m n. m \ n \ C n \ C m" shows "(\n. C n) \ {}" proof - let ?\ = "range (\n. \m \ n. C m)" have "closedin X A" if "A \ ?\" for A using that clo by auto moreover have "(\n\K. \m \ n. C m) \ {}" if "finite K" for K proof - obtain n where "\k. k \ K \ k \ n" using Max.coboundedI \finite K\ by blast with inc have "C n \ (\n\K. \m \ n. C m)" by blast with ne [of n] show ?thesis by blast qed ultimately show ?thesis using \compact_space X\ [unfolded compact_space_fip, rule_format, of ?\] by (simp add: all_finite_subset_image INT_extend_simps UN_atMost_UNIV del: INT_simps) qed lemma compactin_discrete_topology: "compactin (discrete_topology X) S \ S \ X \ finite S" (is "?lhs = ?rhs") proof (intro iffI conjI) assume L: ?lhs then show "S \ X" by (auto simp: compactin_def) have *: "\\. Ball \ (openin (discrete_topology X)) \ S \ \\ \ (\\. finite \ \ \ \ \ \ S \ \\)" using L by (auto simp: compactin_def) show "finite S" using * [of "(\x. {x}) ` X"] \S \ X\ by clarsimp (metis UN_singleton finite_subset_image infinite_super) next assume ?rhs then show ?lhs by (simp add: finite_imp_compactin) qed lemma compact_space_discrete_topology: "compact_space(discrete_topology X) \ finite X" by (simp add: compactin_discrete_topology compact_space_def) lemma compact_space_imp_Bolzano_Weierstrass: assumes "compact_space X" "infinite S" "S \ topspace X" shows "X derived_set_of S \ {}" proof assume X: "X derived_set_of S = {}" then have "closedin X S" by (simp add: closedin_contains_derived_set assms) then have "compactin X S" by (rule closedin_compact_space [OF \compact_space X\]) with X show False by (metis \infinite S\ compactin_subspace compact_space_discrete_topology inf_bot_right subtopology_eq_discrete_topology_eq) qed lemma compactin_imp_Bolzano_Weierstrass: "\compactin X S; infinite T \ T \ S\ \ S \ X derived_set_of T \ {}" using compact_space_imp_Bolzano_Weierstrass [of "subtopology X S"] by (simp add: compactin_subspace derived_set_of_subtopology inf_absorb2) lemma compact_closure_of_imp_Bolzano_Weierstrass: "\compactin X (X closure_of S); infinite T; T \ S; T \ topspace X\ \ X derived_set_of T \ {}" using closure_of_mono closure_of_subset compactin_imp_Bolzano_Weierstrass by fastforce lemma discrete_compactin_eq_finite: "S \ X derived_set_of S = {} \ compactin X S \ S \ topspace X \ finite S" apply (rule iffI) using compactin_imp_Bolzano_Weierstrass compactin_subset_topspace apply blast by (simp add: finite_imp_compactin_eq) lemma discrete_compact_space_eq_finite: "X derived_set_of (topspace X) = {} \ (compact_space X \ finite(topspace X))" by (metis compact_space_discrete_topology discrete_topology_unique_derived_set) lemma image_compactin: assumes cpt: "compactin X S" and cont: "continuous_map X Y f" shows "compactin Y (f ` S)" unfolding compactin_def proof (intro conjI allI impI) show "f ` S \ topspace Y" using compactin_subset_topspace cont continuous_map_image_subset_topspace cpt by blast next fix \ :: "'b set set" assume \: "Ball \ (openin Y) \ f ` S \ \\" define \ where "\ \ (\U. {x \ topspace X. f x \ U}) ` \" have "S \ topspace X" and *: "\\. \\U\\. openin X U; S \ \\\ \ \\. finite \ \ \ \ \ \ S \ \\" using cpt by (auto simp: compactin_def) obtain \ where \: "finite \" "\ \ \" "S \ \\" proof - have 1: "\U\\. openin X U" unfolding \_def using \ cont[unfolded continuous_map] by blast have 2: "S \ \\" unfolding \_def using compactin_subset_topspace cpt \ by fastforce show thesis using * [OF 1 2] that by metis qed have "\v \ \. \U. U \ \ \ v = {x \ topspace X. f x \ U}" using \_def by blast then obtain U where U: "\v \ \. U v \ \ \ v = {x \ topspace X. f x \ U v}" by metis show "\\. finite \ \ \ \ \ \ f ` S \ \\" proof (intro conjI exI) show "finite (U ` \)" by (simp add: \finite \\) next show "U ` \ \ \" using \\ \ \\ U by auto next show "f ` S \ \ (U ` \)" using \(2-3) U UnionE subset_eq U by fastforce qed qed lemma homeomorphic_compact_space: assumes "X homeomorphic_space Y" shows "compact_space X \ compact_space Y" using homeomorphic_space_sym by (metis assms compact_space_def homeomorphic_eq_everything_map homeomorphic_space image_compactin) lemma homeomorphic_map_compactness: assumes hom: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "compactin Y (f ` U) \ compactin X U" proof - have "f ` U \ topspace Y" using hom U homeomorphic_imp_surjective_map by blast moreover have "homeomorphic_map (subtopology X U) (subtopology Y (f ` U)) f" using U hom homeomorphic_imp_surjective_map by (blast intro: homeomorphic_map_subtopologies) then have "compact_space (subtopology Y (f ` U)) = compact_space (subtopology X U)" using homeomorphic_compact_space homeomorphic_map_imp_homeomorphic_space by blast ultimately show ?thesis by (simp add: compactin_subspace U) qed lemma homeomorphic_map_compactness_eq: "homeomorphic_map X Y f \ compactin X U \ U \ topspace X \ compactin Y (f ` U)" by (meson compactin_subset_topspace homeomorphic_map_compactness) subsection\Embedding maps\ definition embedding_map where "embedding_map X Y f \ homeomorphic_map X (subtopology Y (f ` (topspace X))) f" lemma embedding_map_eq: "\embedding_map X Y f; \x. x \ topspace X \ f x = g x\ \ embedding_map X Y g" unfolding embedding_map_def by (metis homeomorphic_map_eq image_cong) lemma embedding_map_compose: assumes "embedding_map X X' f" "embedding_map X' X'' g" shows "embedding_map X X'' (g \ f)" proof - have hm: "homeomorphic_map X (subtopology X' (f ` topspace X)) f" "homeomorphic_map X' (subtopology X'' (g ` topspace X')) g" using assms by (auto simp: embedding_map_def) then obtain C where "g ` topspace X' \ C = (g \ f) ` topspace X" by (metis (no_types) Int_absorb1 continuous_map_image_subset_topspace continuous_map_in_subtopology homeomorphic_eq_everything_map image_comp image_mono) then have "homeomorphic_map (subtopology X' (f ` topspace X)) (subtopology X'' ((g \ f) ` topspace X)) g" by (metis hm homeomorphic_imp_surjective_map homeomorphic_map_subtopologies image_comp subtopology_subtopology topspace_subtopology) then show ?thesis unfolding embedding_map_def using hm(1) homeomorphic_map_compose by blast qed lemma surjective_embedding_map: "embedding_map X Y f \ f ` (topspace X) = topspace Y \ homeomorphic_map X Y f" by (force simp: embedding_map_def homeomorphic_eq_everything_map) lemma embedding_map_in_subtopology: "embedding_map X (subtopology Y S) f \ embedding_map X Y f \ f ` (topspace X) \ S" apply (auto simp: embedding_map_def subtopology_subtopology Int_absorb1) apply (metis (no_types) homeomorphic_imp_surjective_map subtopology_subtopology subtopology_topspace topspace_subtopology) apply (simp add: continuous_map_def homeomorphic_eq_everything_map) done lemma injective_open_imp_embedding_map: "\continuous_map X Y f; open_map X Y f; inj_on f (topspace X)\ \ embedding_map X Y f" unfolding embedding_map_def apply (rule bijective_open_imp_homeomorphic_map) using continuous_map_in_subtopology apply blast apply (auto simp: continuous_map_in_subtopology open_map_into_subtopology continuous_map) done lemma injective_closed_imp_embedding_map: "\continuous_map X Y f; closed_map X Y f; inj_on f (topspace X)\ \ embedding_map X Y f" unfolding embedding_map_def apply (rule bijective_closed_imp_homeomorphic_map) apply (simp_all add: continuous_map_into_subtopology closed_map_into_subtopology) apply (simp add: continuous_map inf.absorb_iff2) done lemma embedding_map_imp_homeomorphic_space: "embedding_map X Y f \ X homeomorphic_space (subtopology Y (f ` (topspace X)))" unfolding embedding_map_def using homeomorphic_space by blast lemma embedding_imp_closed_map: "\embedding_map X Y f; closedin Y (f ` topspace X)\ \ closed_map X Y f" unfolding closed_map_def by (auto simp: closedin_closed_subtopology embedding_map_def homeomorphic_map_closedness_eq) subsection\Retraction and section maps\ definition retraction_maps :: "'a topology \ 'b topology \ ('a \ 'b) \ ('b \ 'a) \ bool" where "retraction_maps X Y f g \ continuous_map X Y f \ continuous_map Y X g \ (\x \ topspace Y. f(g x) = x)" definition section_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "section_map X Y f \ \g. retraction_maps Y X g f" definition retraction_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "retraction_map X Y f \ \g. retraction_maps X Y f g" lemma retraction_maps_eq: "\retraction_maps X Y f g; \x. x \ topspace X \ f x = f' x; \x. x \ topspace Y \ g x = g' x\ \ retraction_maps X Y f' g'" unfolding retraction_maps_def by (metis (no_types, lifting) continuous_map_def continuous_map_eq) lemma section_map_eq: "\section_map X Y f; \x. x \ topspace X \ f x = g x\ \ section_map X Y g" unfolding section_map_def using retraction_maps_eq by blast lemma retraction_map_eq: "\retraction_map X Y f; \x. x \ topspace X \ f x = g x\ \ retraction_map X Y g" unfolding retraction_map_def using retraction_maps_eq by blast lemma homeomorphic_imp_retraction_maps: "homeomorphic_maps X Y f g \ retraction_maps X Y f g" by (simp add: homeomorphic_maps_def retraction_maps_def) lemma section_and_retraction_eq_homeomorphic_map: "section_map X Y f \ retraction_map X Y f \ homeomorphic_map X Y f" apply (auto simp: section_map_def retraction_map_def homeomorphic_map_maps retraction_maps_def homeomorphic_maps_def) by (metis (full_types) continuous_map_image_subset_topspace image_subset_iff) lemma section_imp_embedding_map: "section_map X Y f \ embedding_map X Y f" unfolding section_map_def embedding_map_def homeomorphic_map_maps retraction_maps_def homeomorphic_maps_def by (force simp: continuous_map_in_subtopology continuous_map_from_subtopology) lemma retraction_imp_quotient_map: assumes "retraction_map X Y f" shows "quotient_map X Y f" unfolding quotient_map_def proof (intro conjI subsetI allI impI) show "f ` topspace X = topspace Y" using assms by (force simp: retraction_map_def retraction_maps_def continuous_map_def) next fix U assume U: "U \ topspace Y" have "openin Y U" if "\x\topspace Y. g x \ topspace X" "\x\topspace Y. f (g x) = x" "openin Y {x \ topspace Y. g x \ {x \ topspace X. f x \ U}}" for g using openin_subopen U that by fastforce then show "openin X {x \ topspace X. f x \ U} = openin Y U" using assms by (auto simp: retraction_map_def retraction_maps_def continuous_map_def) qed lemma retraction_maps_compose: "\retraction_maps X Y f f'; retraction_maps Y Z g g'\ \ retraction_maps X Z (g \ f) (f' \ g')" by (clarsimp simp: retraction_maps_def continuous_map_compose) (simp add: continuous_map_def) lemma retraction_map_compose: "\retraction_map X Y f; retraction_map Y Z g\ \ retraction_map X Z (g \ f)" by (meson retraction_map_def retraction_maps_compose) lemma section_map_compose: "\section_map X Y f; section_map Y Z g\ \ section_map X Z (g \ f)" by (meson retraction_maps_compose section_map_def) lemma surjective_section_eq_homeomorphic_map: "section_map X Y f \ f ` (topspace X) = topspace Y \ homeomorphic_map X Y f" by (meson section_and_retraction_eq_homeomorphic_map section_imp_embedding_map surjective_embedding_map) lemma surjective_retraction_or_section_map: "f ` (topspace X) = topspace Y \ retraction_map X Y f \ section_map X Y f \ retraction_map X Y f" using section_and_retraction_eq_homeomorphic_map surjective_section_eq_homeomorphic_map by fastforce lemma retraction_imp_surjective_map: "retraction_map X Y f \ f ` (topspace X) = topspace Y" by (simp add: retraction_imp_quotient_map quotient_imp_surjective_map) lemma section_imp_injective_map: "\section_map X Y f; x \ topspace X; y \ topspace X\ \ f x = f y \ x = y" by (metis (mono_tags, hide_lams) retraction_maps_def section_map_def) lemma retraction_maps_to_retract_maps: "retraction_maps X Y r s \ retraction_maps X (subtopology X (s ` (topspace Y))) (s \ r) id" unfolding retraction_maps_def by (auto simp: continuous_map_compose continuous_map_into_subtopology continuous_map_from_subtopology) subsection \Continuity\ lemma continuous_on_open: "continuous_on S f \ (\T. openin (top_of_set (f ` S)) T \ openin (top_of_set S) (S \ f -` T))" unfolding continuous_on_open_invariant openin_open Int_def vimage_def Int_commute by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong) lemma continuous_on_closed: "continuous_on S f \ (\T. closedin (top_of_set (f ` S)) T \ closedin (top_of_set S) (S \ f -` T))" unfolding continuous_on_closed_invariant closedin_closed Int_def vimage_def Int_commute by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong) lemma continuous_on_imp_closedin: assumes "continuous_on S f" "closedin (top_of_set (f ` S)) T" shows "closedin (top_of_set S) (S \ f -` T)" using assms continuous_on_closed by blast lemma continuous_map_subtopology_eu [simp]: "continuous_map (top_of_set S) (subtopology euclidean T) h \ continuous_on S h \ h ` S \ T" apply safe apply (metis continuous_map_closedin_preimage_eq continuous_map_in_subtopology continuous_on_closed order_refl topspace_euclidean_subtopology) apply (simp add: continuous_map_closedin_preimage_eq image_subset_iff) by (metis (no_types, hide_lams) continuous_map_closedin_preimage_eq continuous_map_in_subtopology continuous_on_closed order_refl topspace_euclidean_subtopology) lemma continuous_map_euclidean_top_of_set: assumes eq: "f -` S = UNIV" and cont: "continuous_on UNIV f" shows "continuous_map euclidean (top_of_set S) f" by (simp add: cont continuous_map_into_subtopology eq image_subset_iff_subset_vimage) subsection\<^marker>\tag unimportant\ \Half-global and completely global cases\ lemma continuous_openin_preimage_gen: assumes "continuous_on S f" "open T" shows "openin (top_of_set S) (S \ f -` T)" proof - have *: "(S \ f -` T) = (S \ f -` (T \ f ` S))" by auto have "openin (top_of_set (f ` S)) (T \ f ` S)" using openin_open_Int[of T "f ` S", OF assms(2)] unfolding openin_open by auto then show ?thesis using assms(1)[unfolded continuous_on_open, THEN spec[where x="T \ f ` S"]] using * by auto qed lemma continuous_closedin_preimage: assumes "continuous_on S f" and "closed T" shows "closedin (top_of_set S) (S \ f -` T)" proof - have *: "(S \ f -` T) = (S \ f -` (T \ f ` S))" by auto have "closedin (top_of_set (f ` S)) (T \ f ` S)" using closedin_closed_Int[of T "f ` S", OF assms(2)] by (simp add: Int_commute) then show ?thesis using assms(1)[unfolded continuous_on_closed, THEN spec[where x="T \ f ` S"]] using * by auto qed lemma continuous_openin_preimage_eq: "continuous_on S f \ (\T. open T \ openin (top_of_set S) (S \ f -` T))" apply safe apply (simp add: continuous_openin_preimage_gen) apply (fastforce simp add: continuous_on_open openin_open) done lemma continuous_closedin_preimage_eq: "continuous_on S f \ (\T. closed T \ closedin (top_of_set S) (S \ f -` T))" apply safe apply (simp add: continuous_closedin_preimage) apply (fastforce simp add: continuous_on_closed closedin_closed) done lemma continuous_open_preimage: assumes contf: "continuous_on S f" and "open S" "open T" shows "open (S \ f -` T)" proof- obtain U where "open U" "(S \ f -` T) = S \ U" using continuous_openin_preimage_gen[OF contf \open T\] unfolding openin_open by auto then show ?thesis using open_Int[of S U, OF \open S\] by auto qed lemma continuous_closed_preimage: assumes contf: "continuous_on S f" and "closed S" "closed T" shows "closed (S \ f -` T)" proof- obtain U where "closed U" "(S \ f -` T) = S \ U" using continuous_closedin_preimage[OF contf \closed T\] unfolding closedin_closed by auto then show ?thesis using closed_Int[of S U, OF \closed S\] by auto qed lemma continuous_open_vimage: "open S \ (\x. continuous (at x) f) \ open (f -` S)" by (metis continuous_on_eq_continuous_within open_vimage) lemma continuous_closed_vimage: "closed S \ (\x. continuous (at x) f) \ closed (f -` S)" by (simp add: closed_vimage continuous_on_eq_continuous_within) lemma Times_in_interior_subtopology: assumes "(x, y) \ U" "openin (top_of_set (S \ T)) U" obtains V W where "openin (top_of_set S) V" "x \ V" "openin (top_of_set T) W" "y \ W" "(V \ W) \ U" proof - from assms obtain E where "open E" "U = S \ T \ E" "(x, y) \ E" "x \ S" "y \ T" by (auto simp: openin_open) from open_prod_elim[OF \open E\ \(x, y) \ E\] obtain E1 E2 where "open E1" "open E2" "(x, y) \ E1 \ E2" "E1 \ E2 \ E" by blast show ?thesis proof show "openin (top_of_set S) (E1 \ S)" "openin (top_of_set T) (E2 \ T)" using \open E1\ \open E2\ by (auto simp: openin_open) show "x \ E1 \ S" "y \ E2 \ T" using \(x, y) \ E1 \ E2\ \x \ S\ \y \ T\ by auto show "(E1 \ S) \ (E2 \ T) \ U" using \E1 \ E2 \ E\ \U = _\ by (auto simp: ) qed qed lemma closedin_Times: "closedin (top_of_set S) S' \ closedin (top_of_set T) T' \ closedin (top_of_set (S \ T)) (S' \ T')" unfolding closedin_closed using closed_Times by blast lemma openin_Times: "openin (top_of_set S) S' \ openin (top_of_set T) T' \ openin (top_of_set (S \ T)) (S' \ T')" unfolding openin_open using open_Times by blast lemma openin_Times_eq: fixes S :: "'a::topological_space set" and T :: "'b::topological_space set" shows "openin (top_of_set (S \ T)) (S' \ T') \ S' = {} \ T' = {} \ openin (top_of_set S) S' \ openin (top_of_set T) T'" (is "?lhs = ?rhs") proof (cases "S' = {} \ T' = {}") case True then show ?thesis by auto next case False then obtain x y where "x \ S'" "y \ T'" by blast show ?thesis proof assume ?lhs have "openin (top_of_set S) S'" apply (subst openin_subopen, clarify) apply (rule Times_in_interior_subtopology [OF _ \?lhs\]) using \y \ T'\ apply auto done moreover have "openin (top_of_set T) T'" apply (subst openin_subopen, clarify) apply (rule Times_in_interior_subtopology [OF _ \?lhs\]) using \x \ S'\ apply auto done ultimately show ?rhs by simp next assume ?rhs with False show ?lhs by (simp add: openin_Times) qed qed lemma Lim_transform_within_openin: assumes f: "(f \ l) (at a within T)" and "openin (top_of_set T) S" "a \ S" and eq: "\x. \x \ S; x \ a\ \ f x = g x" shows "(g \ l) (at a within T)" proof - have "\\<^sub>F x in at a within T. x \ T \ x \ a" by (simp add: eventually_at_filter) moreover from \openin _ _\ obtain U where "open U" "S = T \ U" by (auto simp: openin_open) then have "a \ U" using \a \ S\ by auto from topological_tendstoD[OF tendsto_ident_at \open U\ \a \ U\] have "\\<^sub>F x in at a within T. x \ U" by auto ultimately have "\\<^sub>F x in at a within T. f x = g x" by eventually_elim (auto simp: \S = _\ eq) with f show ?thesis by (rule Lim_transform_eventually) qed lemma continuous_on_open_gen: assumes "f ` S \ T" shows "continuous_on S f \ (\U. openin (top_of_set T) U \ openin (top_of_set S) (S \ f -` U))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (clarsimp simp add: continuous_openin_preimage_eq openin_open) (metis Int_assoc assms image_subset_iff_subset_vimage inf.absorb_iff1) next assume R [rule_format]: ?rhs show ?lhs proof (clarsimp simp add: continuous_openin_preimage_eq) fix U::"'a set" assume "open U" then have "openin (top_of_set S) (S \ f -` (U \ T))" by (metis R inf_commute openin_open) then show "openin (top_of_set S) (S \ f -` U)" by (metis Int_assoc Int_commute assms image_subset_iff_subset_vimage inf.absorb_iff2 vimage_Int) qed qed lemma continuous_openin_preimage: "\continuous_on S f; f ` S \ T; openin (top_of_set T) U\ \ openin (top_of_set S) (S \ f -` U)" by (simp add: continuous_on_open_gen) lemma continuous_on_closed_gen: assumes "f ` S \ T" shows "continuous_on S f \ (\U. closedin (top_of_set T) U \ closedin (top_of_set S) (S \ f -` U))" (is "?lhs = ?rhs") proof - have *: "U \ T \ S \ f -` (T - U) = S - (S \ f -` U)" for U using assms by blast show ?thesis proof assume L: ?lhs show ?rhs proof clarify fix U assume "closedin (top_of_set T) U" then show "closedin (top_of_set S) (S \ f -` U)" using L unfolding continuous_on_open_gen [OF assms] by (metis * closedin_def inf_le1 topspace_euclidean_subtopology) qed next assume R [rule_format]: ?rhs show ?lhs unfolding continuous_on_open_gen [OF assms] by (metis * R inf_le1 openin_closedin_eq topspace_euclidean_subtopology) qed qed lemma continuous_closedin_preimage_gen: assumes "continuous_on S f" "f ` S \ T" "closedin (top_of_set T) U" shows "closedin (top_of_set S) (S \ f -` U)" using assms continuous_on_closed_gen by blast lemma continuous_transform_within_openin: assumes "continuous (at a within T) f" and "openin (top_of_set T) S" "a \ S" and eq: "\x. x \ S \ f x = g x" shows "continuous (at a within T) g" using assms by (simp add: Lim_transform_within_openin continuous_within) subsection\<^marker>\tag important\ \The topology generated by some (open) subsets\ text \In the definition below of a generated topology, the \Empty\ case is not necessary, as it follows from \UN\ taking for \K\ the empty set. However, it is convenient to have, and is never a problem in proofs, so I prefer to write it down explicitly. We do not require \UNIV\ to be an open set, as this will not be the case in applications. (We are thinking of a topology on a subset of \UNIV\, the remaining part of \UNIV\ being irrelevant.)\ inductive generate_topology_on for S where Empty: "generate_topology_on S {}" | Int: "generate_topology_on S a \ generate_topology_on S b \ generate_topology_on S (a \ b)" | UN: "(\k. k \ K \ generate_topology_on S k) \ generate_topology_on S (\K)" | Basis: "s \ S \ generate_topology_on S s" lemma istopology_generate_topology_on: "istopology (generate_topology_on S)" unfolding istopology_def by (auto intro: generate_topology_on.intros) text \The basic property of the topology generated by a set \S\ is that it is the smallest topology containing all the elements of \S\:\ lemma generate_topology_on_coarsest: assumes "istopology T" "\s. s \ S \ T s" "generate_topology_on S s0" shows "T s0" using assms(3) apply (induct rule: generate_topology_on.induct) using assms(1) assms(2) unfolding istopology_def by auto abbreviation\<^marker>\tag unimportant\ topology_generated_by::"('a set set) \ ('a topology)" where "topology_generated_by S \ topology (generate_topology_on S)" lemma openin_topology_generated_by_iff: "openin (topology_generated_by S) s \ generate_topology_on S s" using topology_inverse'[OF istopology_generate_topology_on[of S]] by simp lemma openin_topology_generated_by: "openin (topology_generated_by S) s \ generate_topology_on S s" using openin_topology_generated_by_iff by auto lemma topology_generated_by_topspace [simp]: "topspace (topology_generated_by S) = (\S)" proof { fix s assume "openin (topology_generated_by S) s" then have "generate_topology_on S s" by (rule openin_topology_generated_by) then have "s \ (\S)" by (induct, auto) } then show "topspace (topology_generated_by S) \ (\S)" unfolding topspace_def by auto next have "generate_topology_on S (\S)" using generate_topology_on.UN[OF generate_topology_on.Basis, of S S] by simp then show "(\S) \ topspace (topology_generated_by S)" unfolding topspace_def using openin_topology_generated_by_iff by auto qed lemma topology_generated_by_Basis: "s \ S \ openin (topology_generated_by S) s" by (simp only: openin_topology_generated_by_iff, auto simp: generate_topology_on.Basis) lemma generate_topology_on_Inter: "\finite \; \K. K \ \ \ generate_topology_on \ K; \ \ {}\ \ generate_topology_on \ (\\)" by (induction \ rule: finite_induct; force intro: generate_topology_on.intros) subsection\Topology bases and sub-bases\ lemma istopology_base_alt: "istopology (arbitrary union_of P) \ (\S T. (arbitrary union_of P) S \ (arbitrary union_of P) T \ (arbitrary union_of P) (S \ T))" by (simp add: istopology_def) (blast intro: arbitrary_union_of_Union) lemma istopology_base_eq: "istopology (arbitrary union_of P) \ (\S T. P S \ P T \ (arbitrary union_of P) (S \ T))" by (simp add: istopology_base_alt arbitrary_union_of_Int_eq) lemma istopology_base: "(\S T. \P S; P T\ \ P(S \ T)) \ istopology (arbitrary union_of P)" by (simp add: arbitrary_def istopology_base_eq union_of_inc) lemma openin_topology_base_unique: "openin X = arbitrary union_of P \ (\V. P V \ openin X V) \ (\U x. openin X U \ x \ U \ (\V. P V \ x \ V \ V \ U))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (auto simp: union_of_def arbitrary_def) next assume R: ?rhs then have *: "\\\Collect P. \\ = S" if "openin X S" for S using that by (rule_tac x="{V. P V \ V \ S}" in exI) fastforce from R show ?lhs by (fastforce simp add: union_of_def arbitrary_def intro: *) qed lemma topology_base_unique: - "\\S. P S \ openin X S; - \U x. \openin X U; x \ U\ \ \B. P B \ x \ B \ B \ U\ - \ topology(arbitrary union_of P) = X" - by (metis openin_topology_base_unique openin_inverse [of X]) + assumes "\S. P S \ openin X S" + "\U x. \openin X U; x \ U\ \ \B. P B \ x \ B \ B \ U" + shows "topology (arbitrary union_of P) = X" +proof - + have "X = topology (openin X)" + by (simp add: openin_inverse) + also from assms have "openin X = arbitrary union_of P" + by (subst openin_topology_base_unique) auto + finally show ?thesis .. +qed lemma topology_bases_eq_aux: "\(arbitrary union_of P) S; \U x. \P U; x \ U\ \ \V. Q V \ x \ V \ V \ U\ \ (arbitrary union_of Q) S" by (metis arbitrary_union_of_alt arbitrary_union_of_idempot) lemma topology_bases_eq: "\\U x. \P U; x \ U\ \ \V. Q V \ x \ V \ V \ U; \V x. \Q V; x \ V\ \ \U. P U \ x \ U \ U \ V\ \ topology (arbitrary union_of P) = topology (arbitrary union_of Q)" by (fastforce intro: arg_cong [where f=topology] elim: topology_bases_eq_aux) lemma istopology_subbase: "istopology (arbitrary union_of (finite intersection_of P relative_to S))" by (simp add: finite_intersection_of_Int istopology_base relative_to_Int) lemma openin_subbase: "openin (topology (arbitrary union_of (finite intersection_of B relative_to U))) S \ (arbitrary union_of (finite intersection_of B relative_to U)) S" by (simp add: istopology_subbase topology_inverse') lemma topspace_subbase [simp]: "topspace(topology (arbitrary union_of (finite intersection_of B relative_to U))) = U" (is "?lhs = _") proof show "?lhs \ U" by (metis arbitrary_union_of_relative_to openin_subbase openin_topspace relative_to_imp_subset) show "U \ ?lhs" by (metis arbitrary_union_of_inc finite_intersection_of_empty inf.orderE istopology_subbase openin_subset relative_to_inc subset_UNIV topology_inverse') qed lemma minimal_topology_subbase: "\\S. P S \ openin X S; openin X U; openin(topology(arbitrary union_of (finite intersection_of P relative_to U))) S\ \ openin X S" apply (simp add: istopology_subbase topology_inverse) apply (simp add: union_of_def intersection_of_def relative_to_def) apply (blast intro: openin_Int_Inter) done lemma istopology_subbase_UNIV: "istopology (arbitrary union_of (finite intersection_of P))" by (simp add: istopology_base finite_intersection_of_Int) lemma generate_topology_on_eq: "generate_topology_on S = arbitrary union_of finite' intersection_of (\x. x \ S)" (is "?lhs = ?rhs") proof (intro ext iffI) fix A assume "?lhs A" then show "?rhs A" proof induction case (Int a b) then show ?case by (metis (mono_tags, lifting) istopology_base_alt finite'_intersection_of_Int istopology_base) next case (UN K) then show ?case by (simp add: arbitrary_union_of_Union) next case (Basis s) then show ?case by (simp add: Sup_upper arbitrary_union_of_inc finite'_intersection_of_inc relative_to_subset) qed auto next fix A assume "?rhs A" then obtain \ where \: "\T. T \ \ \ \\. finite' \ \ \ \ S \ \\ = T" and eq: "A = \\" unfolding union_of_def intersection_of_def by auto show "?lhs A" unfolding eq proof (rule generate_topology_on.UN) fix T assume "T \ \" with \ obtain \ where "finite' \" "\ \ S" "\\ = T" by blast have "generate_topology_on S (\\)" proof (rule generate_topology_on_Inter) show "finite \" "\ \ {}" by (auto simp: \finite' \\) show "\K. K \ \ \ generate_topology_on S K" by (metis \\ \ S\ generate_topology_on.simps subset_iff) qed then show "generate_topology_on S T" using \\\ = T\ by blast qed qed lemma continuous_on_generated_topo_iff: "continuous_map T1 (topology_generated_by S) f \ ((\U. U \ S \ openin T1 (f-`U \ topspace(T1))) \ (f`(topspace T1) \ (\ S)))" unfolding continuous_map_alt topology_generated_by_topspace proof (auto simp add: topology_generated_by_Basis) assume H: "\U. U \ S \ openin T1 (f -` U \ topspace T1)" fix U assume "openin (topology_generated_by S) U" then have "generate_topology_on S U" by (rule openin_topology_generated_by) then show "openin T1 (f -` U \ topspace T1)" proof (induct) fix a b assume H: "openin T1 (f -` a \ topspace T1)" "openin T1 (f -` b \ topspace T1)" have "f -` (a \ b) \ topspace T1 = (f-`a \ topspace T1) \ (f-`b \ topspace T1)" by auto then show "openin T1 (f -` (a \ b) \ topspace T1)" using H by auto next fix K assume H: "openin T1 (f -` k \ topspace T1)" if "k\ K" for k define L where "L = {f -` k \ topspace T1|k. k \ K}" have *: "openin T1 l" if "l \L" for l using that H unfolding L_def by auto have "openin T1 (\L)" using openin_Union[OF *] by simp moreover have "(\L) = (f -` \K \ topspace T1)" unfolding L_def by auto ultimately show "openin T1 (f -` \K \ topspace T1)" by simp qed (auto simp add: H) qed lemma continuous_on_generated_topo: assumes "\U. U \S \ openin T1 (f-`U \ topspace(T1))" "f`(topspace T1) \ (\ S)" shows "continuous_map T1 (topology_generated_by S) f" using assms continuous_on_generated_topo_iff by blast subsection\<^marker>\tag important\ \Pullback topology\ text \Pulling back a topology by map gives again a topology. \subtopology\ is a special case of this notion, pulling back by the identity. We introduce the general notion as we will need it to define the strong operator topology on the space of continuous linear operators, by pulling back the product topology on the space of all functions.\ text \\pullback_topology A f T\ is the pullback of the topology \T\ by the map \f\ on the set \A\.\ definition\<^marker>\tag important\ pullback_topology::"('a set) \ ('a \ 'b) \ ('b topology) \ ('a topology)" where "pullback_topology A f T = topology (\S. \U. openin T U \ S = f-`U \ A)" lemma istopology_pullback_topology: "istopology (\S. \U. openin T U \ S = f-`U \ A)" unfolding istopology_def proof (auto) fix K assume "\S\K. \U. openin T U \ S = f -` U \ A" then have "\U. \S\K. openin T (U S) \ S = f-`(U S) \ A" by (rule bchoice) then obtain U where U: "\S\K. openin T (U S) \ S = f-`(U S) \ A" by blast define V where "V = (\S\K. U S)" have "openin T V" "\K = f -` V \ A" unfolding V_def using U by auto then show "\V. openin T V \ \K = f -` V \ A" by auto qed lemma openin_pullback_topology: "openin (pullback_topology A f T) S \ (\U. openin T U \ S = f-`U \ A)" unfolding pullback_topology_def topology_inverse'[OF istopology_pullback_topology] by auto lemma topspace_pullback_topology: "topspace (pullback_topology A f T) = f-`(topspace T) \ A" by (auto simp add: topspace_def openin_pullback_topology) proposition continuous_map_pullback [intro]: assumes "continuous_map T1 T2 g" shows "continuous_map (pullback_topology A f T1) T2 (g o f)" unfolding continuous_map_alt proof (auto) fix U::"'b set" assume "openin T2 U" then have "openin T1 (g-`U \ topspace T1)" using assms unfolding continuous_map_alt by auto have "(g o f)-`U \ topspace (pullback_topology A f T1) = (g o f)-`U \ A \ f-`(topspace T1)" unfolding topspace_pullback_topology by auto also have "... = f-`(g-`U \ topspace T1) \ A " by auto also have "openin (pullback_topology A f T1) (...)" unfolding openin_pullback_topology using \openin T1 (g-`U \ topspace T1)\ by auto finally show "openin (pullback_topology A f T1) ((g \ f) -` U \ topspace (pullback_topology A f T1))" by auto next fix x assume "x \ topspace (pullback_topology A f T1)" then have "f x \ topspace T1" unfolding topspace_pullback_topology by auto then show "g (f x) \ topspace T2" using assms unfolding continuous_map_def by auto qed proposition continuous_map_pullback' [intro]: assumes "continuous_map T1 T2 (f o g)" "topspace T1 \ g-`A" shows "continuous_map T1 (pullback_topology A f T2) g" unfolding continuous_map_alt proof (auto) fix U assume "openin (pullback_topology A f T2) U" then have "\V. openin T2 V \ U = f-`V \ A" unfolding openin_pullback_topology by auto then obtain V where "openin T2 V" "U = f-`V \ A" by blast then have "g -` U \ topspace T1 = g-`(f-`V \ A) \ topspace T1" by blast also have "... = (f o g)-`V \ (g-`A \ topspace T1)" by auto also have "... = (f o g)-`V \ topspace T1" using assms(2) by auto also have "openin T1 (...)" using assms(1) \openin T2 V\ by auto finally show "openin T1 (g -` U \ topspace T1)" by simp next fix x assume "x \ topspace T1" have "(f o g) x \ topspace T2" using assms(1) \x \ topspace T1\ unfolding continuous_map_def by auto then have "g x \ f-`(topspace T2)" unfolding comp_def by blast moreover have "g x \ A" using assms(2) \x \ topspace T1\ by blast ultimately show "g x \ topspace (pullback_topology A f T2)" unfolding topspace_pullback_topology by blast qed subsection\Proper maps (not a priori assumed continuous) \ definition proper_map where "proper_map X Y f \ closed_map X Y f \ (\y \ topspace Y. compactin X {x \ topspace X. f x = y})" lemma proper_imp_closed_map: "proper_map X Y f \ closed_map X Y f" by (simp add: proper_map_def) lemma proper_map_imp_subset_topspace: "proper_map X Y f \ f ` (topspace X) \ topspace Y" by (simp add: closed_map_imp_subset_topspace proper_map_def) lemma closed_injective_imp_proper_map: assumes f: "closed_map X Y f" and inj: "inj_on f (topspace X)" shows "proper_map X Y f" unfolding proper_map_def proof (clarsimp simp: f) show "compactin X {x \ topspace X. f x = y}" if "y \ topspace Y" for y proof - have "{x \ topspace X. f x = y} = {} \ (\a \ topspace X. {x \ topspace X. f x = y} = {a})" using inj_on_eq_iff [OF inj] by auto then show ?thesis using that by (metis (no_types, lifting) compactin_empty compactin_sing) qed qed lemma injective_imp_proper_eq_closed_map: "inj_on f (topspace X) \ (proper_map X Y f \ closed_map X Y f)" using closed_injective_imp_proper_map proper_imp_closed_map by blast lemma homeomorphic_imp_proper_map: "homeomorphic_map X Y f \ proper_map X Y f" by (simp add: closed_injective_imp_proper_map homeomorphic_eq_everything_map) lemma compactin_proper_map_preimage: assumes f: "proper_map X Y f" and "compactin Y K" shows "compactin X {x. x \ topspace X \ f x \ K}" proof - have "f ` (topspace X) \ topspace Y" by (simp add: f proper_map_imp_subset_topspace) have *: "\y. y \ topspace Y \ compactin X {x \ topspace X. f x = y}" using f by (auto simp: proper_map_def) show ?thesis unfolding compactin_def proof clarsimp show "\\. finite \ \ \ \ \ \ {x \ topspace X. f x \ K} \ \\" if \: "\U\\. openin X U" and sub: "{x \ topspace X. f x \ K} \ \\" for \ proof - have "\y \ K. \\. finite \ \ \ \ \ \ {x \ topspace X. f x = y} \ \\" proof fix y assume "y \ K" then have "compactin X {x \ topspace X. f x = y}" by (metis "*" \compactin Y K\ compactin_subspace subsetD) with \y \ K\ show "\\. finite \ \ \ \ \ \ {x \ topspace X. f x = y} \ \\" unfolding compactin_def using \ sub by fastforce qed then obtain \ where \: "\y. y \ K \ finite (\ y) \ \ y \ \ \ {x \ topspace X. f x = y} \ \(\ y)" by (metis (full_types)) define F where "F \ \y. topspace Y - f ` (topspace X - \(\ y))" have "\\. finite \ \ \ \ F ` K \ K \ \\" proof (rule compactinD [OF \compactin Y K\]) have "\x. x \ K \ closedin Y (f ` (topspace X - \(\ x)))" using f unfolding proper_map_def closed_map_def by (meson \ \ openin_Union openin_closedin_eq subsetD) then show "openin Y U" if "U \ F ` K" for U using that by (auto simp: F_def) show "K \ \(F ` K)" using \ \compactin Y K\ unfolding F_def compactin_def by fastforce qed then obtain J where "finite J" "J \ K" and J: "K \ \(F ` J)" by (auto simp: ex_finite_subset_image) show ?thesis unfolding F_def proof (intro exI conjI) show "finite (\(\ ` J))" using \ \J \ K\ \finite J\ by blast show "\(\ ` J) \ \" using \ \J \ K\ by blast show "{x \ topspace X. f x \ K} \ \(\(\ ` J))" using J \J \ K\ unfolding F_def by auto qed qed qed qed lemma compact_space_proper_map_preimage: assumes f: "proper_map X Y f" and fim: "f ` (topspace X) = topspace Y" and "compact_space Y" shows "compact_space X" proof - have eq: "topspace X = {x \ topspace X. f x \ topspace Y}" using fim by blast moreover have "compactin Y (topspace Y)" using \compact_space Y\ compact_space_def by auto ultimately show ?thesis unfolding compact_space_def using eq f compactin_proper_map_preimage by fastforce qed lemma proper_map_alt: "proper_map X Y f \ closed_map X Y f \ (\K. compactin Y K \ compactin X {x. x \ topspace X \ f x \ K})" proof (intro iffI conjI allI impI) show "compactin X {x \ topspace X. f x \ K}" if "proper_map X Y f" and "compactin Y K" for K using that by (simp add: compactin_proper_map_preimage) show "proper_map X Y f" if f: "closed_map X Y f \ (\K. compactin Y K \ compactin X {x \ topspace X. f x \ K})" proof - have "compactin X {x \ topspace X. f x = y}" if "y \ topspace Y" for y proof - have "compactin X {x \ topspace X. f x \ {y}}" using f compactin_sing that by fastforce then show ?thesis by auto qed with f show ?thesis by (auto simp: proper_map_def) qed qed (simp add: proper_imp_closed_map) lemma proper_map_on_empty: "topspace X = {} \ proper_map X Y f" by (auto simp: proper_map_def closed_map_on_empty) lemma proper_map_id [simp]: "proper_map X X id" proof (clarsimp simp: proper_map_alt closed_map_id) fix K assume K: "compactin X K" then have "{a \ topspace X. a \ K} = K" by (simp add: compactin_subspace subset_antisym subset_iff) then show "compactin X {a \ topspace X. a \ K}" using K by auto qed lemma proper_map_compose: assumes "proper_map X Y f" "proper_map Y Z g" shows "proper_map X Z (g \ f)" proof - have "closed_map X Y f" and f: "\K. compactin Y K \ compactin X {x \ topspace X. f x \ K}" and "closed_map Y Z g" and g: "\K. compactin Z K \ compactin Y {x \ topspace Y. g x \ K}" using assms by (auto simp: proper_map_alt) show ?thesis unfolding proper_map_alt proof (intro conjI allI impI) show "closed_map X Z (g \ f)" using \closed_map X Y f\ \closed_map Y Z g\ closed_map_compose by blast have "{x \ topspace X. g (f x) \ K} = {x \ topspace X. f x \ {b \ topspace Y. g b \ K}}" for K using \closed_map X Y f\ closed_map_imp_subset_topspace by blast then show "compactin X {x \ topspace X. (g \ f) x \ K}" if "compactin Z K" for K using f [OF g [OF that]] by auto qed qed lemma proper_map_const: "proper_map X Y (\x. c) \ compact_space X \ (topspace X = {} \ closedin Y {c})" proof (cases "topspace X = {}") case True then show ?thesis by (simp add: compact_space_topspace_empty proper_map_on_empty) next case False have *: "compactin X {x \ topspace X. c = y}" if "compact_space X" for y proof (cases "c = y") case True then show ?thesis using compact_space_def \compact_space X\ by auto qed auto then show ?thesis using closed_compactin closedin_subset by (force simp: False proper_map_def closed_map_const compact_space_def) qed lemma proper_map_inclusion: "s \ topspace X \ proper_map (subtopology X s) X id \ closedin X s \ (\k. compactin X k \ compactin X (s \ k))" by (auto simp: proper_map_alt closed_map_inclusion_eq inf.absorb_iff2 Collect_conj_eq compactin_subtopology intro: closed_Int_compactin) subsection\Perfect maps (proper, continuous and surjective)\ definition perfect_map where "perfect_map X Y f \ continuous_map X Y f \ proper_map X Y f \ f ` (topspace X) = topspace Y" lemma homeomorphic_imp_perfect_map: "homeomorphic_map X Y f \ perfect_map X Y f" by (simp add: homeomorphic_eq_everything_map homeomorphic_imp_proper_map perfect_map_def) lemma perfect_imp_quotient_map: "perfect_map X Y f \ quotient_map X Y f" by (simp add: continuous_closed_imp_quotient_map perfect_map_def proper_map_def) lemma homeomorphic_eq_injective_perfect_map: "homeomorphic_map X Y f \ perfect_map X Y f \ inj_on f (topspace X)" using homeomorphic_imp_perfect_map homeomorphic_map_def perfect_imp_quotient_map by blast lemma perfect_injective_eq_homeomorphic_map: "perfect_map X Y f \ inj_on f (topspace X) \ homeomorphic_map X Y f" by (simp add: homeomorphic_eq_injective_perfect_map) lemma perfect_map_id [simp]: "perfect_map X X id" by (simp add: homeomorphic_imp_perfect_map) lemma perfect_map_compose: "\perfect_map X Y f; perfect_map Y Z g\ \ perfect_map X Z (g \ f)" by (meson continuous_map_compose perfect_imp_quotient_map perfect_map_def proper_map_compose quotient_map_compose_eq quotient_map_def) lemma perfect_imp_continuous_map: "perfect_map X Y f \ continuous_map X Y f" using perfect_map_def by blast lemma perfect_imp_closed_map: "perfect_map X Y f \ closed_map X Y f" by (simp add: perfect_map_def proper_map_def) lemma perfect_imp_proper_map: "perfect_map X Y f \ proper_map X Y f" by (simp add: perfect_map_def) lemma perfect_imp_surjective_map: "perfect_map X Y f \ f ` (topspace X) = topspace Y" by (simp add: perfect_map_def) end diff --git a/src/HOL/Analysis/Affine.thy b/src/HOL/Analysis/Affine.thy --- a/src/HOL/Analysis/Affine.thy +++ b/src/HOL/Analysis/Affine.thy @@ -1,1655 +1,1655 @@ section "Affine Sets" theory Affine imports Linear_Algebra begin lemma if_smult: "(if P then x else (y::real)) *\<^sub>R v = (if P then x *\<^sub>R v else y *\<^sub>R v)" by (fact if_distrib) lemma sum_delta_notmem: assumes "x \ s" shows "sum (\y. if (y = x) then P x else Q y) s = sum Q s" and "sum (\y. if (x = y) then P x else Q y) s = sum Q s" and "sum (\y. if (y = x) then P y else Q y) s = sum Q s" and "sum (\y. if (x = y) then P y else Q y) s = sum Q s" apply (rule_tac [!] sum.cong) using assms apply auto done lemmas independent_finite = independent_imp_finite lemma span_substd_basis: assumes d: "d \ Basis" shows "span d = {x. \i\Basis. i \ d \ x\i = 0}" (is "_ = ?B") proof - have "d \ ?B" using d by (auto simp: inner_Basis) moreover have s: "subspace ?B" using subspace_substandard[of "\i. i \ d"] . ultimately have "span d \ ?B" using span_mono[of d "?B"] span_eq_iff[of "?B"] by blast moreover have *: "card d \ dim (span d)" using independent_card_le_dim[of d "span d"] independent_substdbasis[OF assms] span_superset[of d] by auto moreover from * have "dim ?B \ dim (span d)" using dim_substandard[OF assms] by auto ultimately show ?thesis using s subspace_dim_equal[of "span d" "?B"] subspace_span[of d] by auto qed lemma basis_to_substdbasis_subspace_isomorphism: fixes B :: "'a::euclidean_space set" assumes "independent B" shows "\f d::'a set. card d = card B \ linear f \ f ` B = d \ f ` span B = {x. \i\Basis. i \ d \ x \ i = 0} \ inj_on f (span B) \ d \ Basis" proof - have B: "card B = dim B" using dim_unique[of B B "card B"] assms span_superset[of B] by auto have "dim B \ card (Basis :: 'a set)" using dim_subset_UNIV[of B] by simp from ex_card[OF this] obtain d :: "'a set" where d: "d \ Basis" and t: "card d = dim B" by auto let ?t = "{x::'a::euclidean_space. \i\Basis. i \ d \ x\i = 0}" have "\f. linear f \ f ` B = d \ f ` span B = ?t \ inj_on f (span B)" proof (intro basis_to_basis_subspace_isomorphism subspace_span subspace_substandard span_superset) show "d \ {x. \i\Basis. i \ d \ x \ i = 0}" using d inner_not_same_Basis by blast qed (auto simp: span_substd_basis independent_substdbasis dim_substandard d t B assms) with t \card B = dim B\ d show ?thesis by auto qed subsection \Affine set and affine hull\ definition\<^marker>\tag important\ affine :: "'a::real_vector set \ bool" where "affine s \ (\x\s. \y\s. \u v. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ s)" lemma affine_alt: "affine s \ (\x\s. \y\s. \u::real. (1 - u) *\<^sub>R x + u *\<^sub>R y \ s)" unfolding affine_def by (metis eq_diff_eq') lemma affine_empty [iff]: "affine {}" unfolding affine_def by auto lemma affine_sing [iff]: "affine {x}" unfolding affine_alt by (auto simp: scaleR_left_distrib [symmetric]) lemma affine_UNIV [iff]: "affine UNIV" unfolding affine_def by auto lemma affine_Inter [intro]: "(\s. s\f \ affine s) \ affine (\f)" unfolding affine_def by auto lemma affine_Int[intro]: "affine s \ affine t \ affine (s \ t)" unfolding affine_def by auto lemma affine_scaling: "affine s \ affine (image (\x. c *\<^sub>R x) s)" apply (clarsimp simp add: affine_def) apply (rule_tac x="u *\<^sub>R x + v *\<^sub>R y" in image_eqI) apply (auto simp: algebra_simps) done lemma affine_affine_hull [simp]: "affine(affine hull s)" unfolding hull_def using affine_Inter[of "{t. affine t \ s \ t}"] by auto lemma affine_hull_eq[simp]: "(affine hull s = s) \ affine s" by (metis affine_affine_hull hull_same) lemma affine_hyperplane: "affine {x. a \ x = b}" by (simp add: affine_def algebra_simps) (metis distrib_right mult.left_neutral) subsubsection\<^marker>\tag unimportant\ \Some explicit formulations\ text "Formalized by Lars Schewe." lemma affine: fixes V::"'a::real_vector set" shows "affine V \ (\S u. finite S \ S \ {} \ S \ V \ sum u S = 1 \ (\x\S. u x *\<^sub>R x) \ V)" proof - have "u *\<^sub>R x + v *\<^sub>R y \ V" if "x \ V" "y \ V" "u + v = (1::real)" and *: "\S u. \finite S; S \ {}; S \ V; sum u S = 1\ \ (\x\S. u x *\<^sub>R x) \ V" for x y u v proof (cases "x = y") case True then show ?thesis using that by (metis scaleR_add_left scaleR_one) next case False then show ?thesis using that *[of "{x,y}" "\w. if w = x then u else v"] by auto qed moreover have "(\x\S. u x *\<^sub>R x) \ V" if *: "\x y u v. \x\V; y\V; u + v = 1\ \ u *\<^sub>R x + v *\<^sub>R y \ V" and "finite S" "S \ {}" "S \ V" "sum u S = 1" for S u proof - define n where "n = card S" consider "card S = 0" | "card S = 1" | "card S = 2" | "card S > 2" by linarith then show "(\x\S. u x *\<^sub>R x) \ V" proof cases assume "card S = 1" then obtain a where "S={a}" by (auto simp: card_Suc_eq) then show ?thesis using that by simp next assume "card S = 2" then obtain a b where "S = {a, b}" by (metis Suc_1 card_1_singletonE card_Suc_eq) then show ?thesis using *[of a b] that by (auto simp: sum_clauses(2)) next assume "card S > 2" then show ?thesis using that n_def proof (induct n arbitrary: u S) case 0 then show ?case by auto next case (Suc n u S) have "sum u S = card S" if "\ (\x\S. u x \ 1)" using that unfolding card_eq_sum by auto with Suc.prems obtain x where "x \ S" and x: "u x \ 1" by force have c: "card (S - {x}) = card S - 1" by (simp add: Suc.prems(3) \x \ S\) have "sum u (S - {x}) = 1 - u x" by (simp add: Suc.prems sum_diff1 \x \ S\) with x have eq1: "inverse (1 - u x) * sum u (S - {x}) = 1" by auto have inV: "(\y\S - {x}. (inverse (1 - u x) * u y) *\<^sub>R y) \ V" proof (cases "card (S - {x}) > 2") case True then have S: "S - {x} \ {}" "card (S - {x}) = n" using Suc.prems c by force+ show ?thesis proof (rule Suc.hyps) show "(\a\S - {x}. inverse (1 - u x) * u a) = 1" by (auto simp: eq1 sum_distrib_left[symmetric]) qed (use S Suc.prems True in auto) next case False then have "card (S - {x}) = Suc (Suc 0)" using Suc.prems c by auto then obtain a b where ab: "(S - {x}) = {a, b}" "a\b" unfolding card_Suc_eq by auto then show ?thesis using eq1 \S \ V\ by (auto simp: sum_distrib_left distrib_left intro!: Suc.prems(2)[of a b]) qed have "u x + (1 - u x) = 1 \ u x *\<^sub>R x + (1 - u x) *\<^sub>R ((\y\S - {x}. u y *\<^sub>R y) /\<^sub>R (1 - u x)) \ V" by (rule Suc.prems) (use \x \ S\ Suc.prems inV in \auto simp: scaleR_right.sum\) moreover have "(\a\S. u a *\<^sub>R a) = u x *\<^sub>R x + (\a\S - {x}. u a *\<^sub>R a)" by (meson Suc.prems(3) sum.remove \x \ S\) ultimately show "(\x\S. u x *\<^sub>R x) \ V" by (simp add: x) qed qed (use \S\{}\ \finite S\ in auto) qed ultimately show ?thesis unfolding affine_def by meson qed lemma affine_hull_explicit: "affine hull p = {y. \S u. finite S \ S \ {} \ S \ p \ sum u S = 1 \ sum (\v. u v *\<^sub>R v) S = y}" (is "_ = ?rhs") proof (rule hull_unique) show "p \ ?rhs" proof (intro subsetI CollectI exI conjI) show "\x. sum (\z. 1) {x} = 1" by auto qed auto show "?rhs \ T" if "p \ T" "affine T" for T using that unfolding affine by blast show "affine ?rhs" unfolding affine_def proof clarify fix u v :: real and sx ux sy uy assume uv: "u + v = 1" and x: "finite sx" "sx \ {}" "sx \ p" "sum ux sx = (1::real)" and y: "finite sy" "sy \ {}" "sy \ p" "sum uy sy = (1::real)" have **: "(sx \ sy) \ sx = sx" "(sx \ sy) \ sy = sy" by auto show "\S w. finite S \ S \ {} \ S \ p \ sum w S = 1 \ (\v\S. w v *\<^sub>R v) = u *\<^sub>R (\v\sx. ux v *\<^sub>R v) + v *\<^sub>R (\v\sy. uy v *\<^sub>R v)" proof (intro exI conjI) show "finite (sx \ sy)" using x y by auto show "sum (\i. (if i\sx then u * ux i else 0) + (if i\sy then v * uy i else 0)) (sx \ sy) = 1" using x y uv by (simp add: sum_Un sum.distrib sum.inter_restrict[symmetric] sum_distrib_left [symmetric] **) have "(\i\sx \ sy. ((if i \ sx then u * ux i else 0) + (if i \ sy then v * uy i else 0)) *\<^sub>R i) = (\i\sx. (u * ux i) *\<^sub>R i) + (\i\sy. (v * uy i) *\<^sub>R i)" using x y unfolding scaleR_left_distrib scaleR_zero_left if_smult by (simp add: sum_Un sum.distrib sum.inter_restrict[symmetric] **) also have "\ = u *\<^sub>R (\v\sx. ux v *\<^sub>R v) + v *\<^sub>R (\v\sy. uy v *\<^sub>R v)" unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] by blast finally show "(\i\sx \ sy. ((if i \ sx then u * ux i else 0) + (if i \ sy then v * uy i else 0)) *\<^sub>R i) = u *\<^sub>R (\v\sx. ux v *\<^sub>R v) + v *\<^sub>R (\v\sy. uy v *\<^sub>R v)" . qed (use x y in auto) qed qed lemma affine_hull_finite: assumes "finite S" shows "affine hull S = {y. \u. sum u S = 1 \ sum (\v. u v *\<^sub>R v) S = y}" proof - have *: "\h. sum h S = 1 \ (\v\S. h v *\<^sub>R v) = x" if "F \ S" "finite F" "F \ {}" and sum: "sum u F = 1" and x: "(\v\F. u v *\<^sub>R v) = x" for x F u proof - have "S \ F = F" using that by auto show ?thesis proof (intro exI conjI) show "(\x\S. if x \ F then u x else 0) = 1" by (metis (mono_tags, lifting) \S \ F = F\ assms sum.inter_restrict sum) show "(\v\S. (if v \ F then u v else 0) *\<^sub>R v) = x" by (simp add: if_smult cong: if_cong) (metis (no_types) \S \ F = F\ assms sum.inter_restrict x) qed qed show ?thesis unfolding affine_hull_explicit using assms by (fastforce dest: *) qed subsubsection\<^marker>\tag unimportant\ \Stepping theorems and hence small special cases\ lemma affine_hull_empty[simp]: "affine hull {} = {}" by simp lemma affine_hull_finite_step: fixes y :: "'a::real_vector" shows "finite S \ (\u. sum u (insert a S) = w \ sum (\x. u x *\<^sub>R x) (insert a S) = y) \ (\v u. sum u S = w - v \ sum (\x. u x *\<^sub>R x) S = y - v *\<^sub>R a)" (is "_ \ ?lhs = ?rhs") proof - assume fin: "finite S" show "?lhs = ?rhs" proof assume ?lhs then obtain u where u: "sum u (insert a S) = w \ (\x\insert a S. u x *\<^sub>R x) = y" by auto show ?rhs proof (cases "a \ S") case True then show ?thesis using u by (simp add: insert_absorb) (metis diff_zero real_vector.scale_zero_left) next case False show ?thesis by (rule exI [where x="u a"]) (use u fin False in auto) qed next assume ?rhs then obtain v u where vu: "sum u S = w - v" "(\x\S. u x *\<^sub>R x) = y - v *\<^sub>R a" by auto have *: "\x M. (if x = a then v else M) *\<^sub>R x = (if x = a then v *\<^sub>R x else M *\<^sub>R x)" by auto show ?lhs proof (cases "a \ S") case True show ?thesis by (rule exI [where x="\x. (if x=a then v else 0) + u x"]) (simp add: True scaleR_left_distrib sum.distrib sum_clauses fin vu * cong: if_cong) next case False then show ?thesis apply (rule_tac x="\x. if x=a then v else u x" in exI) apply (simp add: vu sum_clauses(2)[OF fin] *) by (simp add: sum_delta_notmem(3) vu) qed qed qed lemma affine_hull_2: fixes a b :: "'a::real_vector" shows "affine hull {a,b} = {u *\<^sub>R a + v *\<^sub>R b| u v. (u + v = 1)}" (is "?lhs = ?rhs") proof - have *: "\x y z. z = x - y \ y + z = (x::real)" "\x y z. z = x - y \ y + z = (x::'a)" by auto have "?lhs = {y. \u. sum u {a, b} = 1 \ (\v\{a, b}. u v *\<^sub>R v) = y}" using affine_hull_finite[of "{a,b}"] by auto also have "\ = {y. \v u. u b = 1 - v \ u b *\<^sub>R b = y - v *\<^sub>R a}" by (simp add: affine_hull_finite_step[of "{b}" a]) also have "\ = ?rhs" unfolding * by auto finally show ?thesis by auto qed lemma affine_hull_3: fixes a b c :: "'a::real_vector" shows "affine hull {a,b,c} = { u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c| u v w. u + v + w = 1}" proof - have *: "\x y z. z = x - y \ y + z = (x::real)" "\x y z. z = x - y \ y + z = (x::'a)" by auto show ?thesis apply (simp add: affine_hull_finite affine_hull_finite_step) unfolding * apply safe apply (metis add.assoc) apply (rule_tac x=u in exI, force) done qed lemma mem_affine: assumes "affine S" "x \ S" "y \ S" "u + v = 1" shows "u *\<^sub>R x + v *\<^sub>R y \ S" using assms affine_def[of S] by auto lemma mem_affine_3: assumes "affine S" "x \ S" "y \ S" "z \ S" "u + v + w = 1" shows "u *\<^sub>R x + v *\<^sub>R y + w *\<^sub>R z \ S" proof - have "u *\<^sub>R x + v *\<^sub>R y + w *\<^sub>R z \ affine hull {x, y, z}" using affine_hull_3[of x y z] assms by auto moreover have "affine hull {x, y, z} \ affine hull S" using hull_mono[of "{x, y, z}" "S"] assms by auto moreover have "affine hull S = S" using assms affine_hull_eq[of S] by auto ultimately show ?thesis by auto qed lemma mem_affine_3_minus: assumes "affine S" "x \ S" "y \ S" "z \ S" shows "x + v *\<^sub>R (y-z) \ S" using mem_affine_3[of S x y z 1 v "-v"] assms by (simp add: algebra_simps) corollary%unimportant mem_affine_3_minus2: "\affine S; x \ S; y \ S; z \ S\ \ x - v *\<^sub>R (y-z) \ S" by (metis add_uminus_conv_diff mem_affine_3_minus real_vector.scale_minus_left) subsubsection\<^marker>\tag unimportant\ \Some relations between affine hull and subspaces\ lemma affine_hull_insert_subset_span: "affine hull (insert a S) \ {a + v| v . v \ span {x - a | x . x \ S}}" proof - have "\v T u. x = a + v \ (finite T \ T \ {x - a |x. x \ S} \ (\v\T. u v *\<^sub>R v) = v)" if "finite F" "F \ {}" "F \ insert a S" "sum u F = 1" "(\v\F. u v *\<^sub>R v) = x" for x F u proof - have *: "(\x. x - a) ` (F - {a}) \ {x - a |x. x \ S}" using that by auto show ?thesis proof (intro exI conjI) show "finite ((\x. x - a) ` (F - {a}))" by (simp add: that(1)) show "(\v\(\x. x - a) ` (F - {a}). u(v+a) *\<^sub>R v) = x-a" by (simp add: sum.reindex[unfolded inj_on_def] algebra_simps sum_subtractf scaleR_left.sum[symmetric] sum_diff1 that) qed (use \F \ insert a S\ in auto) qed then show ?thesis - unfolding affine_hull_explicit span_explicit by blast + unfolding affine_hull_explicit span_explicit by fast qed lemma affine_hull_insert_span: assumes "a \ S" shows "affine hull (insert a S) = {a + v | v . v \ span {x - a | x. x \ S}}" proof - have *: "\G u. finite G \ G \ {} \ G \ insert a S \ sum u G = 1 \ (\v\G. u v *\<^sub>R v) = y" if "v \ span {x - a |x. x \ S}" "y = a + v" for y v proof - from that obtain T u where u: "finite T" "T \ {x - a |x. x \ S}" "a + (\v\T. u v *\<^sub>R v) = y" unfolding span_explicit by auto define F where "F = (\x. x + a) ` T" have F: "finite F" "F \ S" "(\v\F. u (v - a) *\<^sub>R (v - a)) = y - a" unfolding F_def using u by (auto simp: sum.reindex[unfolded inj_on_def]) have *: "F \ {a} = {}" "F \ - {a} = F" using F assms by auto show "\G u. finite G \ G \ {} \ G \ insert a S \ sum u G = 1 \ (\v\G. u v *\<^sub>R v) = y" apply (rule_tac x = "insert a F" in exI) apply (rule_tac x = "\x. if x=a then 1 - sum (\x. u (x - a)) F else u (x - a)" in exI) using assms F apply (auto simp: sum_clauses sum.If_cases if_smult sum_subtractf scaleR_left.sum algebra_simps *) done qed show ?thesis by (intro subset_antisym affine_hull_insert_subset_span) (auto simp: affine_hull_explicit dest!: *) qed lemma affine_hull_span: assumes "a \ S" shows "affine hull S = {a + v | v. v \ span {x - a | x. x \ S - {a}}}" using affine_hull_insert_span[of a "S - {a}", unfolded insert_Diff[OF assms]] by auto subsubsection\<^marker>\tag unimportant\ \Parallel affine sets\ definition affine_parallel :: "'a::real_vector set \ 'a::real_vector set \ bool" where "affine_parallel S T \ (\a. T = (\x. a + x) ` S)" lemma affine_parallel_expl_aux: fixes S T :: "'a::real_vector set" assumes "\x. x \ S \ a + x \ T" shows "T = (\x. a + x) ` S" proof - have "x \ ((\x. a + x) ` S)" if "x \ T" for x using that by (simp add: image_iff) (metis add.commute diff_add_cancel assms) moreover have "T \ (\x. a + x) ` S" using assms by auto ultimately show ?thesis by auto qed lemma affine_parallel_expl: "affine_parallel S T \ (\a. \x. x \ S \ a + x \ T)" by (auto simp add: affine_parallel_def) (use affine_parallel_expl_aux [of S _ T] in blast) lemma affine_parallel_reflex: "affine_parallel S S" unfolding affine_parallel_def using image_add_0 by blast lemma affine_parallel_commut: assumes "affine_parallel A B" shows "affine_parallel B A" proof - from assms obtain a where B: "B = (\x. a + x) ` A" unfolding affine_parallel_def by auto have [simp]: "(\x. x - a) = plus (- a)" by (simp add: fun_eq_iff) from B show ?thesis using translation_galois [of B a A] unfolding affine_parallel_def by blast qed lemma affine_parallel_assoc: assumes "affine_parallel A B" and "affine_parallel B C" shows "affine_parallel A C" proof - from assms obtain ab where "B = (\x. ab + x) ` A" unfolding affine_parallel_def by auto moreover from assms obtain bc where "C = (\x. bc + x) ` B" unfolding affine_parallel_def by auto ultimately show ?thesis using translation_assoc[of bc ab A] unfolding affine_parallel_def by auto qed lemma affine_translation_aux: fixes a :: "'a::real_vector" assumes "affine ((\x. a + x) ` S)" shows "affine S" proof - { fix x y u v assume xy: "x \ S" "y \ S" "(u :: real) + v = 1" then have "(a + x) \ ((\x. a + x) ` S)" "(a + y) \ ((\x. a + x) ` S)" by auto then have h1: "u *\<^sub>R (a + x) + v *\<^sub>R (a + y) \ (\x. a + x) ` S" using xy assms unfolding affine_def by auto have "u *\<^sub>R (a + x) + v *\<^sub>R (a + y) = (u + v) *\<^sub>R a + (u *\<^sub>R x + v *\<^sub>R y)" by (simp add: algebra_simps) also have "\ = a + (u *\<^sub>R x + v *\<^sub>R y)" using \u + v = 1\ by auto ultimately have "a + (u *\<^sub>R x + v *\<^sub>R y) \ (\x. a + x) ` S" using h1 by auto then have "u *\<^sub>R x + v *\<^sub>R y \ S" by auto } then show ?thesis unfolding affine_def by auto qed lemma affine_translation: "affine S \ affine ((+) a ` S)" for a :: "'a::real_vector" proof show "affine ((+) a ` S)" if "affine S" using that translation_assoc [of "- a" a S] by (auto intro: affine_translation_aux [of "- a" "((+) a ` S)"]) show "affine S" if "affine ((+) a ` S)" using that by (rule affine_translation_aux) qed lemma parallel_is_affine: fixes S T :: "'a::real_vector set" assumes "affine S" "affine_parallel S T" shows "affine T" proof - from assms obtain a where "T = (\x. a + x) ` S" unfolding affine_parallel_def by auto then show ?thesis using affine_translation assms by auto qed lemma subspace_imp_affine: "subspace s \ affine s" unfolding subspace_def affine_def by auto lemma affine_hull_subset_span: "(affine hull s) \ (span s)" by (metis hull_minimal span_superset subspace_imp_affine subspace_span) subsubsection\<^marker>\tag unimportant\ \Subspace parallel to an affine set\ lemma subspace_affine: "subspace S \ affine S \ 0 \ S" proof - have h0: "subspace S \ affine S \ 0 \ S" using subspace_imp_affine[of S] subspace_0 by auto { assume assm: "affine S \ 0 \ S" { fix c :: real fix x assume x: "x \ S" have "c *\<^sub>R x = (1-c) *\<^sub>R 0 + c *\<^sub>R x" by auto moreover have "(1 - c) *\<^sub>R 0 + c *\<^sub>R x \ S" using affine_alt[of S] assm x by auto ultimately have "c *\<^sub>R x \ S" by auto } then have h1: "\c. \x \ S. c *\<^sub>R x \ S" by auto { fix x y assume xy: "x \ S" "y \ S" define u where "u = (1 :: real)/2" have "(1/2) *\<^sub>R (x+y) = (1/2) *\<^sub>R (x+y)" by auto moreover have "(1/2) *\<^sub>R (x+y)=(1/2) *\<^sub>R x + (1-(1/2)) *\<^sub>R y" by (simp add: algebra_simps) moreover have "(1 - u) *\<^sub>R x + u *\<^sub>R y \ S" using affine_alt[of S] assm xy by auto ultimately have "(1/2) *\<^sub>R (x+y) \ S" using u_def by auto moreover have "x + y = 2 *\<^sub>R ((1/2) *\<^sub>R (x+y))" by auto ultimately have "x + y \ S" using h1[rule_format, of "(1/2) *\<^sub>R (x+y)" "2"] by auto } then have "\x \ S. \y \ S. x + y \ S" by auto then have "subspace S" using h1 assm unfolding subspace_def by auto } then show ?thesis using h0 by metis qed lemma affine_diffs_subspace: assumes "affine S" "a \ S" shows "subspace ((\x. (-a)+x) ` S)" proof - have [simp]: "(\x. x - a) = plus (- a)" by (simp add: fun_eq_iff) have "affine ((\x. (-a)+x) ` S)" using affine_translation assms by blast moreover have "0 \ ((\x. (-a)+x) ` S)" using assms exI[of "(\x. x\S \ -a+x = 0)" a] by auto ultimately show ?thesis using subspace_affine by auto qed lemma affine_diffs_subspace_subtract: "subspace ((\x. x - a) ` S)" if "affine S" "a \ S" using that affine_diffs_subspace [of _ a] by simp lemma parallel_subspace_explicit: assumes "affine S" and "a \ S" assumes "L \ {y. \x \ S. (-a) + x = y}" shows "subspace L \ affine_parallel S L" proof - from assms have "L = plus (- a) ` S" by auto then have par: "affine_parallel S L" unfolding affine_parallel_def .. then have "affine L" using assms parallel_is_affine by auto moreover have "0 \ L" using assms by auto ultimately show ?thesis using subspace_affine par by auto qed lemma parallel_subspace_aux: assumes "subspace A" and "subspace B" and "affine_parallel A B" shows "A \ B" proof - from assms obtain a where a: "\x. x \ A \ a + x \ B" using affine_parallel_expl[of A B] by auto then have "-a \ A" using assms subspace_0[of B] by auto then have "a \ A" using assms subspace_neg[of A "-a"] by auto then show ?thesis using assms a unfolding subspace_def by auto qed lemma parallel_subspace: assumes "subspace A" and "subspace B" and "affine_parallel A B" shows "A = B" proof show "A \ B" using assms parallel_subspace_aux by auto show "A \ B" using assms parallel_subspace_aux[of B A] affine_parallel_commut by auto qed lemma affine_parallel_subspace: assumes "affine S" "S \ {}" shows "\!L. subspace L \ affine_parallel S L" proof - have ex: "\L. subspace L \ affine_parallel S L" using assms parallel_subspace_explicit by auto { fix L1 L2 assume ass: "subspace L1 \ affine_parallel S L1" "subspace L2 \ affine_parallel S L2" then have "affine_parallel L1 L2" using affine_parallel_commut[of S L1] affine_parallel_assoc[of L1 S L2] by auto then have "L1 = L2" using ass parallel_subspace by auto } then show ?thesis using ex by auto qed subsection \Affine Dependence\ text "Formalized by Lars Schewe." definition\<^marker>\tag important\ affine_dependent :: "'a::real_vector set \ bool" where "affine_dependent s \ (\x\s. x \ affine hull (s - {x}))" lemma affine_dependent_imp_dependent: "affine_dependent s \ dependent s" unfolding affine_dependent_def dependent_def using affine_hull_subset_span by auto lemma affine_dependent_subset: "\affine_dependent s; s \ t\ \ affine_dependent t" apply (simp add: affine_dependent_def Bex_def) apply (blast dest: hull_mono [OF Diff_mono [OF _ subset_refl]]) done lemma affine_independent_subset: shows "\\ affine_dependent t; s \ t\ \ \ affine_dependent s" by (metis affine_dependent_subset) lemma affine_independent_Diff: "\ affine_dependent s \ \ affine_dependent(s - t)" by (meson Diff_subset affine_dependent_subset) proposition affine_dependent_explicit: "affine_dependent p \ (\S u. finite S \ S \ p \ sum u S = 0 \ (\v\S. u v \ 0) \ sum (\v. u v *\<^sub>R v) S = 0)" proof - have "\S u. finite S \ S \ p \ sum u S = 0 \ (\v\S. u v \ 0) \ (\w\S. u w *\<^sub>R w) = 0" if "(\w\S. u w *\<^sub>R w) = x" "x \ p" "finite S" "S \ {}" "S \ p - {x}" "sum u S = 1" for x S u proof (intro exI conjI) have "x \ S" using that by auto then show "(\v \ insert x S. if v = x then - 1 else u v) = 0" using that by (simp add: sum_delta_notmem) show "(\w \ insert x S. (if w = x then - 1 else u w) *\<^sub>R w) = 0" using that \x \ S\ by (simp add: if_smult sum_delta_notmem cong: if_cong) qed (use that in auto) moreover have "\x\p. \S u. finite S \ S \ {} \ S \ p - {x} \ sum u S = 1 \ (\v\S. u v *\<^sub>R v) = x" if "(\v\S. u v *\<^sub>R v) = 0" "finite S" "S \ p" "sum u S = 0" "v \ S" "u v \ 0" for S u v proof (intro bexI exI conjI) have "S \ {v}" using that by auto then show "S - {v} \ {}" using that by auto show "(\x \ S - {v}. - (1 / u v) * u x) = 1" unfolding sum_distrib_left[symmetric] sum_diff1[OF \finite S\] by (simp add: that) show "(\x\S - {v}. (- (1 / u v) * u x) *\<^sub>R x) = v" unfolding sum_distrib_left [symmetric] scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] sum_diff1[OF \finite S\] using that by auto show "S - {v} \ p - {v}" using that by auto qed (use that in auto) ultimately show ?thesis unfolding affine_dependent_def affine_hull_explicit by auto qed lemma affine_dependent_explicit_finite: fixes S :: "'a::real_vector set" assumes "finite S" shows "affine_dependent S \ (\u. sum u S = 0 \ (\v\S. u v \ 0) \ sum (\v. u v *\<^sub>R v) S = 0)" (is "?lhs = ?rhs") proof have *: "\vt u v. (if vt then u v else 0) *\<^sub>R v = (if vt then (u v) *\<^sub>R v else 0::'a)" by auto assume ?lhs then obtain t u v where "finite t" "t \ S" "sum u t = 0" "v\t" "u v \ 0" "(\v\t. u v *\<^sub>R v) = 0" unfolding affine_dependent_explicit by auto then show ?rhs apply (rule_tac x="\x. if x\t then u x else 0" in exI) apply (auto simp: * sum.inter_restrict[OF assms, symmetric] Int_absorb1[OF \t\S\]) done next assume ?rhs then obtain u v where "sum u S = 0" "v\S" "u v \ 0" "(\v\S. u v *\<^sub>R v) = 0" by auto then show ?lhs unfolding affine_dependent_explicit using assms by auto qed lemma dependent_imp_affine_dependent: assumes "dependent {x - a| x . x \ s}" and "a \ s" shows "affine_dependent (insert a s)" proof - from assms(1)[unfolded dependent_explicit] obtain S u v where obt: "finite S" "S \ {x - a |x. x \ s}" "v\S" "u v \ 0" "(\v\S. u v *\<^sub>R v) = 0" by auto define t where "t = (\x. x + a) ` S" have inj: "inj_on (\x. x + a) S" unfolding inj_on_def by auto have "0 \ S" using obt(2) assms(2) unfolding subset_eq by auto have fin: "finite t" and "t \ s" unfolding t_def using obt(1,2) by auto then have "finite (insert a t)" and "insert a t \ insert a s" by auto moreover have *: "\P Q. (\x\t. (if x = a then P x else Q x)) = (\x\t. Q x)" apply (rule sum.cong) using \a\s\ \t\s\ apply auto done have "(\x\insert a t. if x = a then - (\x\t. u (x - a)) else u (x - a)) = 0" unfolding sum_clauses(2)[OF fin] * using \a\s\ \t\s\ by auto moreover have "\v\insert a t. (if v = a then - (\x\t. u (x - a)) else u (v - a)) \ 0" using obt(3,4) \0\S\ by (rule_tac x="v + a" in bexI) (auto simp: t_def) moreover have *: "\P Q. (\x\t. (if x = a then P x else Q x) *\<^sub>R x) = (\x\t. Q x *\<^sub>R x)" using \a\s\ \t\s\ by (auto intro!: sum.cong) have "(\x\t. u (x - a)) *\<^sub>R a = (\v\t. u (v - a) *\<^sub>R v)" unfolding scaleR_left.sum unfolding t_def and sum.reindex[OF inj] and o_def using obt(5) by (auto simp: sum.distrib scaleR_right_distrib) then have "(\v\insert a t. (if v = a then - (\x\t. u (x - a)) else u (v - a)) *\<^sub>R v) = 0" unfolding sum_clauses(2)[OF fin] using \a\s\ \t\s\ by (auto simp: *) ultimately show ?thesis unfolding affine_dependent_explicit apply (rule_tac x="insert a t" in exI, auto) done qed lemma affine_dependent_biggerset: fixes s :: "'a::euclidean_space set" assumes "finite s" "card s \ DIM('a) + 2" shows "affine_dependent s" proof - have "s \ {}" using assms by auto then obtain a where "a\s" by auto have *: "{x - a |x. x \ s - {a}} = (\x. x - a) ` (s - {a})" by auto have "card {x - a |x. x \ s - {a}} = card (s - {a})" unfolding * by (simp add: card_image inj_on_def) also have "\ > DIM('a)" using assms(2) unfolding card_Diff_singleton[OF assms(1) \a\s\] by auto finally show ?thesis apply (subst insert_Diff[OF \a\s\, symmetric]) apply (rule dependent_imp_affine_dependent) apply (rule dependent_biggerset, auto) done qed lemma affine_dependent_biggerset_general: assumes "finite (S :: 'a::euclidean_space set)" and "card S \ dim S + 2" shows "affine_dependent S" proof - from assms(2) have "S \ {}" by auto then obtain a where "a\S" by auto have *: "{x - a |x. x \ S - {a}} = (\x. x - a) ` (S - {a})" by auto have **: "card {x - a |x. x \ S - {a}} = card (S - {a})" by (metis (no_types, lifting) "*" card_image diff_add_cancel inj_on_def) have "dim {x - a |x. x \ S - {a}} \ dim S" using \a\S\ by (auto simp: span_base span_diff intro: subset_le_dim) also have "\ < dim S + 1" by auto also have "\ \ card (S - {a})" using assms using card_Diff_singleton[OF assms(1) \a\S\] by auto finally show ?thesis apply (subst insert_Diff[OF \a\S\, symmetric]) apply (rule dependent_imp_affine_dependent) apply (rule dependent_biggerset_general) unfolding ** apply auto done qed subsection\<^marker>\tag unimportant\ \Some Properties of Affine Dependent Sets\ lemma affine_independent_0 [simp]: "\ affine_dependent {}" by (simp add: affine_dependent_def) lemma affine_independent_1 [simp]: "\ affine_dependent {a}" by (simp add: affine_dependent_def) lemma affine_independent_2 [simp]: "\ affine_dependent {a,b}" by (simp add: affine_dependent_def insert_Diff_if hull_same) lemma affine_hull_translation: "affine hull ((\x. a + x) ` S) = (\x. a + x) ` (affine hull S)" proof - have "affine ((\x. a + x) ` (affine hull S))" using affine_translation affine_affine_hull by blast moreover have "(\x. a + x) ` S \ (\x. a + x) ` (affine hull S)" using hull_subset[of S] by auto ultimately have h1: "affine hull ((\x. a + x) ` S) \ (\x. a + x) ` (affine hull S)" by (metis hull_minimal) have "affine((\x. -a + x) ` (affine hull ((\x. a + x) ` S)))" using affine_translation affine_affine_hull by blast moreover have "(\x. -a + x) ` (\x. a + x) ` S \ (\x. -a + x) ` (affine hull ((\x. a + x) ` S))" using hull_subset[of "(\x. a + x) ` S"] by auto moreover have "S = (\x. -a + x) ` (\x. a + x) ` S" using translation_assoc[of "-a" a] by auto ultimately have "(\x. -a + x) ` (affine hull ((\x. a + x) ` S)) >= (affine hull S)" by (metis hull_minimal) then have "affine hull ((\x. a + x) ` S) >= (\x. a + x) ` (affine hull S)" by auto then show ?thesis using h1 by auto qed lemma affine_dependent_translation: assumes "affine_dependent S" shows "affine_dependent ((\x. a + x) ` S)" proof - obtain x where x: "x \ S \ x \ affine hull (S - {x})" using assms affine_dependent_def by auto have "(+) a ` (S - {x}) = (+) a ` S - {a + x}" by auto then have "a + x \ affine hull ((\x. a + x) ` S - {a + x})" using affine_hull_translation[of a "S - {x}"] x by auto moreover have "a + x \ (\x. a + x) ` S" using x by auto ultimately show ?thesis unfolding affine_dependent_def by auto qed lemma affine_dependent_translation_eq: "affine_dependent S \ affine_dependent ((\x. a + x) ` S)" proof - { assume "affine_dependent ((\x. a + x) ` S)" then have "affine_dependent S" using affine_dependent_translation[of "((\x. a + x) ` S)" "-a"] translation_assoc[of "-a" a] by auto } then show ?thesis using affine_dependent_translation by auto qed lemma affine_hull_0_dependent: assumes "0 \ affine hull S" shows "dependent S" proof - obtain s u where s_u: "finite s \ s \ {} \ s \ S \ sum u s = 1 \ (\v\s. u v *\<^sub>R v) = 0" using assms affine_hull_explicit[of S] by auto then have "\v\s. u v \ 0" by auto then have "finite s \ s \ S \ (\v\s. u v \ 0 \ (\v\s. u v *\<^sub>R v) = 0)" using s_u by auto then show ?thesis unfolding dependent_explicit[of S] by auto qed lemma affine_dependent_imp_dependent2: assumes "affine_dependent (insert 0 S)" shows "dependent S" proof - obtain x where x: "x \ insert 0 S \ x \ affine hull (insert 0 S - {x})" using affine_dependent_def[of "(insert 0 S)"] assms by blast then have "x \ span (insert 0 S - {x})" using affine_hull_subset_span by auto moreover have "span (insert 0 S - {x}) = span (S - {x})" using insert_Diff_if[of "0" S "{x}"] span_insert_0[of "S-{x}"] by auto ultimately have "x \ span (S - {x})" by auto then have "x \ 0 \ dependent S" using x dependent_def by auto moreover { assume "x = 0" then have "0 \ affine hull S" using x hull_mono[of "S - {0}" S] by auto then have "dependent S" using affine_hull_0_dependent by auto } ultimately show ?thesis by auto qed lemma affine_dependent_iff_dependent: assumes "a \ S" shows "affine_dependent (insert a S) \ dependent ((\x. -a + x) ` S)" proof - have "((+) (- a) ` S) = {x - a| x . x \ S}" by auto then show ?thesis using affine_dependent_translation_eq[of "(insert a S)" "-a"] affine_dependent_imp_dependent2 assms dependent_imp_affine_dependent[of a S] by (auto simp del: uminus_add_conv_diff) qed lemma affine_dependent_iff_dependent2: assumes "a \ S" shows "affine_dependent S \ dependent ((\x. -a + x) ` (S-{a}))" proof - have "insert a (S - {a}) = S" using assms by auto then show ?thesis using assms affine_dependent_iff_dependent[of a "S-{a}"] by auto qed lemma affine_hull_insert_span_gen: "affine hull (insert a s) = (\x. a + x) ` span ((\x. - a + x) ` s)" proof - have h1: "{x - a |x. x \ s} = ((\x. -a+x) ` s)" by auto { assume "a \ s" then have ?thesis using affine_hull_insert_span[of a s] h1 by auto } moreover { assume a1: "a \ s" have "\x. x \ s \ -a+x=0" apply (rule exI[of _ a]) using a1 apply auto done then have "insert 0 ((\x. -a+x) ` (s - {a})) = (\x. -a+x) ` s" by auto then have "span ((\x. -a+x) ` (s - {a}))=span ((\x. -a+x) ` s)" using span_insert_0[of "(+) (- a) ` (s - {a})"] by (auto simp del: uminus_add_conv_diff) moreover have "{x - a |x. x \ (s - {a})} = ((\x. -a+x) ` (s - {a}))" by auto moreover have "insert a (s - {a}) = insert a s" by auto ultimately have ?thesis using affine_hull_insert_span[of "a" "s-{a}"] by auto } ultimately show ?thesis by auto qed lemma affine_hull_span2: assumes "a \ s" shows "affine hull s = (\x. a+x) ` span ((\x. -a+x) ` (s-{a}))" using affine_hull_insert_span_gen[of a "s - {a}", unfolded insert_Diff[OF assms]] by auto lemma affine_hull_span_gen: assumes "a \ affine hull s" shows "affine hull s = (\x. a+x) ` span ((\x. -a+x) ` s)" proof - have "affine hull (insert a s) = affine hull s" using hull_redundant[of a affine s] assms by auto then show ?thesis using affine_hull_insert_span_gen[of a "s"] by auto qed lemma affine_hull_span_0: assumes "0 \ affine hull S" shows "affine hull S = span S" using affine_hull_span_gen[of "0" S] assms by auto lemma extend_to_affine_basis_nonempty: fixes S V :: "'n::euclidean_space set" assumes "\ affine_dependent S" "S \ V" "S \ {}" shows "\T. \ affine_dependent T \ S \ T \ T \ V \ affine hull T = affine hull V" proof - obtain a where a: "a \ S" using assms by auto then have h0: "independent ((\x. -a + x) ` (S-{a}))" using affine_dependent_iff_dependent2 assms by auto obtain B where B: "(\x. -a+x) ` (S - {a}) \ B \ B \ (\x. -a+x) ` V \ independent B \ (\x. -a+x) ` V \ span B" using assms by (blast intro: maximal_independent_subset_extend[OF _ h0, of "(\x. -a + x) ` V"]) define T where "T = (\x. a+x) ` insert 0 B" then have "T = insert a ((\x. a+x) ` B)" by auto then have "affine hull T = (\x. a+x) ` span B" using affine_hull_insert_span_gen[of a "((\x. a+x) ` B)"] translation_assoc[of "-a" a B] by auto then have "V \ affine hull T" using B assms translation_inverse_subset[of a V "span B"] by auto moreover have "T \ V" using T_def B a assms by auto ultimately have "affine hull T = affine hull V" by (metis Int_absorb1 Int_absorb2 hull_hull hull_mono) moreover have "S \ T" using T_def B translation_inverse_subset[of a "S-{a}" B] by auto moreover have "\ affine_dependent T" using T_def affine_dependent_translation_eq[of "insert 0 B"] affine_dependent_imp_dependent2 B by auto ultimately show ?thesis using \T \ V\ by auto qed lemma affine_basis_exists: fixes V :: "'n::euclidean_space set" shows "\B. B \ V \ \ affine_dependent B \ affine hull V = affine hull B" proof (cases "V = {}") case True then show ?thesis using affine_independent_0 by auto next case False then obtain x where "x \ V" by auto then show ?thesis using affine_dependent_def[of "{x}"] extend_to_affine_basis_nonempty[of "{x}" V] by auto qed proposition extend_to_affine_basis: fixes S V :: "'n::euclidean_space set" assumes "\ affine_dependent S" "S \ V" obtains T where "\ affine_dependent T" "S \ T" "T \ V" "affine hull T = affine hull V" proof (cases "S = {}") case True then show ?thesis using affine_basis_exists by (metis empty_subsetI that) next case False then show ?thesis by (metis assms extend_to_affine_basis_nonempty that) qed subsection \Affine Dimension of a Set\ definition\<^marker>\tag important\ aff_dim :: "('a::euclidean_space) set \ int" where "aff_dim V = (SOME d :: int. \B. affine hull B = affine hull V \ \ affine_dependent B \ of_nat (card B) = d + 1)" lemma aff_dim_basis_exists: fixes V :: "('n::euclidean_space) set" shows "\B. affine hull B = affine hull V \ \ affine_dependent B \ of_nat (card B) = aff_dim V + 1" proof - obtain B where "\ affine_dependent B \ affine hull B = affine hull V" using affine_basis_exists[of V] by auto then show ?thesis unfolding aff_dim_def some_eq_ex[of "\d. \B. affine hull B = affine hull V \ \ affine_dependent B \ of_nat (card B) = d + 1"] apply auto apply (rule exI[of _ "int (card B) - (1 :: int)"]) apply (rule exI[of _ "B"], auto) done qed lemma affine_hull_eq_empty [simp]: "affine hull S = {} \ S = {}" by (metis affine_empty subset_empty subset_hull) lemma empty_eq_affine_hull[simp]: "{} = affine hull S \ S = {}" by (metis affine_hull_eq_empty) lemma aff_dim_parallel_subspace_aux: fixes B :: "'n::euclidean_space set" assumes "\ affine_dependent B" "a \ B" shows "finite B \ ((card B) - 1 = dim (span ((\x. -a+x) ` (B-{a}))))" proof - have "independent ((\x. -a + x) ` (B-{a}))" using affine_dependent_iff_dependent2 assms by auto then have fin: "dim (span ((\x. -a+x) ` (B-{a}))) = card ((\x. -a + x) ` (B-{a}))" "finite ((\x. -a + x) ` (B - {a}))" using indep_card_eq_dim_span[of "(\x. -a+x) ` (B-{a})"] by auto show ?thesis proof (cases "(\x. -a + x) ` (B - {a}) = {}") case True have "B = insert a ((\x. a + x) ` (\x. -a + x) ` (B - {a}))" using translation_assoc[of "a" "-a" "(B - {a})"] assms by auto then have "B = {a}" using True by auto then show ?thesis using assms fin by auto next case False then have "card ((\x. -a + x) ` (B - {a})) > 0" using fin by auto moreover have h1: "card ((\x. -a + x) ` (B-{a})) = card (B-{a})" by (rule card_image) (use translate_inj_on in blast) ultimately have "card (B-{a}) > 0" by auto then have *: "finite (B - {a})" using card_gt_0_iff[of "(B - {a})"] by auto then have "card (B - {a}) = card B - 1" using card_Diff_singleton assms by auto with * show ?thesis using fin h1 by auto qed qed lemma aff_dim_parallel_subspace: fixes V L :: "'n::euclidean_space set" assumes "V \ {}" and "subspace L" and "affine_parallel (affine hull V) L" shows "aff_dim V = int (dim L)" proof - obtain B where B: "affine hull B = affine hull V \ \ affine_dependent B \ int (card B) = aff_dim V + 1" using aff_dim_basis_exists by auto then have "B \ {}" using assms B by auto then obtain a where a: "a \ B" by auto define Lb where "Lb = span ((\x. -a+x) ` (B-{a}))" moreover have "affine_parallel (affine hull B) Lb" using Lb_def B assms affine_hull_span2[of a B] a affine_parallel_commut[of "Lb" "(affine hull B)"] unfolding affine_parallel_def by auto moreover have "subspace Lb" using Lb_def subspace_span by auto moreover have "affine hull B \ {}" using assms B by auto ultimately have "L = Lb" using assms affine_parallel_subspace[of "affine hull B"] affine_affine_hull[of B] B by auto then have "dim L = dim Lb" by auto moreover have "card B - 1 = dim Lb" and "finite B" using Lb_def aff_dim_parallel_subspace_aux a B by auto ultimately show ?thesis using B \B \ {}\ card_gt_0_iff[of B] by auto qed lemma aff_independent_finite: fixes B :: "'n::euclidean_space set" assumes "\ affine_dependent B" shows "finite B" proof - { assume "B \ {}" then obtain a where "a \ B" by auto then have ?thesis using aff_dim_parallel_subspace_aux assms by auto } then show ?thesis by auto qed lemma aff_dim_empty: fixes S :: "'n::euclidean_space set" shows "S = {} \ aff_dim S = -1" proof - obtain B where *: "affine hull B = affine hull S" and "\ affine_dependent B" and "int (card B) = aff_dim S + 1" using aff_dim_basis_exists by auto moreover from * have "S = {} \ B = {}" by auto ultimately show ?thesis using aff_independent_finite[of B] card_gt_0_iff[of B] by auto qed lemma aff_dim_empty_eq [simp]: "aff_dim ({}::'a::euclidean_space set) = -1" by (simp add: aff_dim_empty [symmetric]) lemma aff_dim_affine_hull [simp]: "aff_dim (affine hull S) = aff_dim S" unfolding aff_dim_def using hull_hull[of _ S] by auto lemma aff_dim_affine_hull2: assumes "affine hull S = affine hull T" shows "aff_dim S = aff_dim T" unfolding aff_dim_def using assms by auto lemma aff_dim_unique: fixes B V :: "'n::euclidean_space set" assumes "affine hull B = affine hull V \ \ affine_dependent B" shows "of_nat (card B) = aff_dim V + 1" proof (cases "B = {}") case True then have "V = {}" using assms by auto then have "aff_dim V = (-1::int)" using aff_dim_empty by auto then show ?thesis using \B = {}\ by auto next case False then obtain a where a: "a \ B" by auto define Lb where "Lb = span ((\x. -a+x) ` (B-{a}))" have "affine_parallel (affine hull B) Lb" using Lb_def affine_hull_span2[of a B] a affine_parallel_commut[of "Lb" "(affine hull B)"] unfolding affine_parallel_def by auto moreover have "subspace Lb" using Lb_def subspace_span by auto ultimately have "aff_dim B = int(dim Lb)" using aff_dim_parallel_subspace[of B Lb] \B \ {}\ by auto moreover have "(card B) - 1 = dim Lb" "finite B" using Lb_def aff_dim_parallel_subspace_aux a assms by auto ultimately have "of_nat (card B) = aff_dim B + 1" using \B \ {}\ card_gt_0_iff[of B] by auto then show ?thesis using aff_dim_affine_hull2 assms by auto qed lemma aff_dim_affine_independent: fixes B :: "'n::euclidean_space set" assumes "\ affine_dependent B" shows "of_nat (card B) = aff_dim B + 1" using aff_dim_unique[of B B] assms by auto lemma affine_independent_iff_card: fixes s :: "'a::euclidean_space set" shows "\ affine_dependent s \ finite s \ aff_dim s = int(card s) - 1" apply (rule iffI) apply (simp add: aff_dim_affine_independent aff_independent_finite) by (metis affine_basis_exists [of s] aff_dim_unique card_subset_eq diff_add_cancel of_nat_eq_iff) lemma aff_dim_sing [simp]: fixes a :: "'n::euclidean_space" shows "aff_dim {a} = 0" using aff_dim_affine_independent[of "{a}"] affine_independent_1 by auto lemma aff_dim_2 [simp]: "aff_dim {a,b} = (if a = b then 0 else 1)" proof (clarsimp) assume "a \ b" then have "aff_dim{a,b} = card{a,b} - 1" using affine_independent_2 [of a b] aff_dim_affine_independent by fastforce also have "\ = 1" using \a \ b\ by simp finally show "aff_dim {a, b} = 1" . qed lemma aff_dim_inner_basis_exists: fixes V :: "('n::euclidean_space) set" shows "\B. B \ V \ affine hull B = affine hull V \ \ affine_dependent B \ of_nat (card B) = aff_dim V + 1" proof - obtain B where B: "\ affine_dependent B" "B \ V" "affine hull B = affine hull V" using affine_basis_exists[of V] by auto then have "of_nat(card B) = aff_dim V+1" using aff_dim_unique by auto with B show ?thesis by auto qed lemma aff_dim_le_card: fixes V :: "'n::euclidean_space set" assumes "finite V" shows "aff_dim V \ of_nat (card V) - 1" proof - obtain B where B: "B \ V" "of_nat (card B) = aff_dim V + 1" using aff_dim_inner_basis_exists[of V] by auto then have "card B \ card V" using assms card_mono by auto with B show ?thesis by auto qed lemma aff_dim_parallel_eq: fixes S T :: "'n::euclidean_space set" assumes "affine_parallel (affine hull S) (affine hull T)" shows "aff_dim S = aff_dim T" proof - { assume "T \ {}" "S \ {}" then obtain L where L: "subspace L \ affine_parallel (affine hull T) L" using affine_parallel_subspace[of "affine hull T"] affine_affine_hull[of T] by auto then have "aff_dim T = int (dim L)" using aff_dim_parallel_subspace \T \ {}\ by auto moreover have *: "subspace L \ affine_parallel (affine hull S) L" using L affine_parallel_assoc[of "affine hull S" "affine hull T" L] assms by auto moreover from * have "aff_dim S = int (dim L)" using aff_dim_parallel_subspace \S \ {}\ by auto ultimately have ?thesis by auto } moreover { assume "S = {}" then have "S = {}" and "T = {}" using assms unfolding affine_parallel_def by auto then have ?thesis using aff_dim_empty by auto } moreover { assume "T = {}" then have "S = {}" and "T = {}" using assms unfolding affine_parallel_def by auto then have ?thesis using aff_dim_empty by auto } ultimately show ?thesis by blast qed lemma aff_dim_translation_eq: "aff_dim ((+) a ` S) = aff_dim S" for a :: "'n::euclidean_space" proof - have "affine_parallel (affine hull S) (affine hull ((\x. a + x) ` S))" unfolding affine_parallel_def apply (rule exI[of _ "a"]) using affine_hull_translation[of a S] apply auto done then show ?thesis using aff_dim_parallel_eq[of S "(\x. a + x) ` S"] by auto qed lemma aff_dim_translation_eq_subtract: "aff_dim ((\x. x - a) ` S) = aff_dim S" for a :: "'n::euclidean_space" using aff_dim_translation_eq [of "- a"] by (simp cong: image_cong_simp) lemma aff_dim_affine: fixes S L :: "'n::euclidean_space set" assumes "S \ {}" and "affine S" and "subspace L" and "affine_parallel S L" shows "aff_dim S = int (dim L)" proof - have *: "affine hull S = S" using assms affine_hull_eq[of S] by auto then have "affine_parallel (affine hull S) L" using assms by (simp add: *) then show ?thesis using assms aff_dim_parallel_subspace[of S L] by blast qed lemma dim_affine_hull: fixes S :: "'n::euclidean_space set" shows "dim (affine hull S) = dim S" proof - have "dim (affine hull S) \ dim S" using dim_subset by auto moreover have "dim (span S) \ dim (affine hull S)" using dim_subset affine_hull_subset_span by blast moreover have "dim (span S) = dim S" using dim_span by auto ultimately show ?thesis by auto qed lemma aff_dim_subspace: fixes S :: "'n::euclidean_space set" assumes "subspace S" shows "aff_dim S = int (dim S)" proof (cases "S={}") case True with assms show ?thesis by (simp add: subspace_affine) next case False with aff_dim_affine[of S S] assms subspace_imp_affine[of S] affine_parallel_reflex[of S] subspace_affine show ?thesis by auto qed lemma aff_dim_zero: fixes S :: "'n::euclidean_space set" assumes "0 \ affine hull S" shows "aff_dim S = int (dim S)" proof - have "subspace (affine hull S)" using subspace_affine[of "affine hull S"] affine_affine_hull assms by auto then have "aff_dim (affine hull S) = int (dim (affine hull S))" using assms aff_dim_subspace[of "affine hull S"] by auto then show ?thesis using aff_dim_affine_hull[of S] dim_affine_hull[of S] by auto qed lemma aff_dim_eq_dim: "aff_dim S = int (dim ((+) (- a) ` S))" if "a \ affine hull S" for S :: "'n::euclidean_space set" proof - have "0 \ affine hull (+) (- a) ` S" unfolding affine_hull_translation using that by (simp add: ac_simps) with aff_dim_zero show ?thesis by (metis aff_dim_translation_eq) qed lemma aff_dim_eq_dim_subtract: "aff_dim S = int (dim ((\x. x - a) ` S))" if "a \ affine hull S" for S :: "'n::euclidean_space set" using aff_dim_eq_dim [of a] that by (simp cong: image_cong_simp) lemma aff_dim_UNIV [simp]: "aff_dim (UNIV :: 'n::euclidean_space set) = int(DIM('n))" using aff_dim_subspace[of "(UNIV :: 'n::euclidean_space set)"] dim_UNIV[where 'a="'n::euclidean_space"] by auto lemma aff_dim_geq: fixes V :: "'n::euclidean_space set" shows "aff_dim V \ -1" proof - obtain B where "affine hull B = affine hull V" and "\ affine_dependent B" and "int (card B) = aff_dim V + 1" using aff_dim_basis_exists by auto then show ?thesis by auto qed lemma aff_dim_negative_iff [simp]: fixes S :: "'n::euclidean_space set" shows "aff_dim S < 0 \S = {}" by (metis aff_dim_empty aff_dim_geq diff_0 eq_iff zle_diff1_eq) lemma aff_lowdim_subset_hyperplane: fixes S :: "'a::euclidean_space set" assumes "aff_dim S < DIM('a)" obtains a b where "a \ 0" "S \ {x. a \ x = b}" proof (cases "S={}") case True moreover have "(SOME b. b \ Basis) \ 0" by (metis norm_some_Basis norm_zero zero_neq_one) ultimately show ?thesis using that by blast next case False then obtain c S' where "c \ S'" "S = insert c S'" by (meson equals0I mk_disjoint_insert) have "dim ((+) (-c) ` S) < DIM('a)" by (metis \S = insert c S'\ aff_dim_eq_dim assms hull_inc insertI1 of_nat_less_imp_less) then obtain a where "a \ 0" "span ((+) (-c) ` S) \ {x. a \ x = 0}" using lowdim_subset_hyperplane by blast moreover have "a \ w = a \ c" if "span ((+) (- c) ` S) \ {x. a \ x = 0}" "w \ S" for w proof - have "w-c \ span ((+) (- c) ` S)" by (simp add: span_base \w \ S\) with that have "w-c \ {x. a \ x = 0}" by blast then show ?thesis by (auto simp: algebra_simps) qed ultimately have "S \ {x. a \ x = a \ c}" by blast then show ?thesis by (rule that[OF \a \ 0\]) qed lemma affine_independent_card_dim_diffs: fixes S :: "'a :: euclidean_space set" assumes "\ affine_dependent S" "a \ S" shows "card S = dim {x - a|x. x \ S} + 1" proof - have 1: "{b - a|b. b \ (S - {a})} \ {x - a|x. x \ S}" by auto have 2: "x - a \ span {b - a |b. b \ S - {a}}" if "x \ S" for x proof (cases "x = a") case True then show ?thesis by (simp add: span_clauses) next case False then show ?thesis using assms by (blast intro: span_base that) qed have "\ affine_dependent (insert a S)" by (simp add: assms insert_absorb) then have 3: "independent {b - a |b. b \ S - {a}}" using dependent_imp_affine_dependent by fastforce have "{b - a |b. b \ S - {a}} = (\b. b-a) ` (S - {a})" by blast then have "card {b - a |b. b \ S - {a}} = card ((\b. b-a) ` (S - {a}))" by simp also have "\ = card (S - {a})" by (metis (no_types, lifting) card_image diff_add_cancel inj_onI) also have "\ = card S - 1" by (simp add: aff_independent_finite assms) finally have 4: "card {b - a |b. b \ S - {a}} = card S - 1" . have "finite S" by (meson assms aff_independent_finite) with \a \ S\ have "card S \ 0" by auto moreover have "dim {x - a |x. x \ S} = card S - 1" using 2 by (blast intro: dim_unique [OF 1 _ 3 4]) ultimately show ?thesis by auto qed lemma independent_card_le_aff_dim: fixes B :: "'n::euclidean_space set" assumes "B \ V" assumes "\ affine_dependent B" shows "int (card B) \ aff_dim V + 1" proof - obtain T where T: "\ affine_dependent T \ B \ T \ T \ V \ affine hull T = affine hull V" by (metis assms extend_to_affine_basis[of B V]) then have "of_nat (card T) = aff_dim V + 1" using aff_dim_unique by auto then show ?thesis using T card_mono[of T B] aff_independent_finite[of T] by auto qed lemma aff_dim_subset: fixes S T :: "'n::euclidean_space set" assumes "S \ T" shows "aff_dim S \ aff_dim T" proof - obtain B where B: "\ affine_dependent B" "B \ S" "affine hull B = affine hull S" "of_nat (card B) = aff_dim S + 1" using aff_dim_inner_basis_exists[of S] by auto then have "int (card B) \ aff_dim T + 1" using assms independent_card_le_aff_dim[of B T] by auto with B show ?thesis by auto qed lemma aff_dim_le_DIM: fixes S :: "'n::euclidean_space set" shows "aff_dim S \ int (DIM('n))" proof - have "aff_dim (UNIV :: 'n::euclidean_space set) = int(DIM('n))" using aff_dim_UNIV by auto then show "aff_dim (S:: 'n::euclidean_space set) \ int(DIM('n))" using aff_dim_subset[of S "(UNIV :: ('n::euclidean_space) set)"] subset_UNIV by auto qed lemma affine_dim_equal: fixes S :: "'n::euclidean_space set" assumes "affine S" "affine T" "S \ {}" "S \ T" "aff_dim S = aff_dim T" shows "S = T" proof - obtain a where "a \ S" using assms by auto then have "a \ T" using assms by auto define LS where "LS = {y. \x \ S. (-a) + x = y}" then have ls: "subspace LS" "affine_parallel S LS" using assms parallel_subspace_explicit[of S a LS] \a \ S\ by auto then have h1: "int(dim LS) = aff_dim S" using assms aff_dim_affine[of S LS] by auto have "T \ {}" using assms by auto define LT where "LT = {y. \x \ T. (-a) + x = y}" then have lt: "subspace LT \ affine_parallel T LT" using assms parallel_subspace_explicit[of T a LT] \a \ T\ by auto then have "int(dim LT) = aff_dim T" using assms aff_dim_affine[of T LT] \T \ {}\ by auto then have "dim LS = dim LT" using h1 assms by auto moreover have "LS \ LT" using LS_def LT_def assms by auto ultimately have "LS = LT" using subspace_dim_equal[of LS LT] ls lt by auto moreover have "S = {x. \y \ LS. a+y=x}" using LS_def by auto moreover have "T = {x. \y \ LT. a+y=x}" using LT_def by auto ultimately show ?thesis by auto qed lemma aff_dim_eq_0: fixes S :: "'a::euclidean_space set" shows "aff_dim S = 0 \ (\a. S = {a})" proof (cases "S = {}") case True then show ?thesis by auto next case False then obtain a where "a \ S" by auto show ?thesis proof safe assume 0: "aff_dim S = 0" have "\ {a,b} \ S" if "b \ a" for b by (metis "0" aff_dim_2 aff_dim_subset not_one_le_zero that) then show "\a. S = {a}" using \a \ S\ by blast qed auto qed lemma affine_hull_UNIV: fixes S :: "'n::euclidean_space set" assumes "aff_dim S = int(DIM('n))" shows "affine hull S = (UNIV :: ('n::euclidean_space) set)" proof - have "S \ {}" using assms aff_dim_empty[of S] by auto have h0: "S \ affine hull S" using hull_subset[of S _] by auto have h1: "aff_dim (UNIV :: ('n::euclidean_space) set) = aff_dim S" using aff_dim_UNIV assms by auto then have h2: "aff_dim (affine hull S) \ aff_dim (UNIV :: ('n::euclidean_space) set)" using aff_dim_le_DIM[of "affine hull S"] assms h0 by auto have h3: "aff_dim S \ aff_dim (affine hull S)" using h0 aff_dim_subset[of S "affine hull S"] assms by auto then have h4: "aff_dim (affine hull S) = aff_dim (UNIV :: ('n::euclidean_space) set)" using h0 h1 h2 by auto then show ?thesis using affine_dim_equal[of "affine hull S" "(UNIV :: ('n::euclidean_space) set)"] affine_affine_hull[of S] affine_UNIV assms h4 h0 \S \ {}\ by auto qed lemma disjoint_affine_hull: fixes s :: "'n::euclidean_space set" assumes "\ affine_dependent s" "t \ s" "u \ s" "t \ u = {}" shows "(affine hull t) \ (affine hull u) = {}" proof - have "finite s" using assms by (simp add: aff_independent_finite) then have "finite t" "finite u" using assms finite_subset by blast+ { fix y assume yt: "y \ affine hull t" and yu: "y \ affine hull u" then obtain a b where a1 [simp]: "sum a t = 1" and [simp]: "sum (\v. a v *\<^sub>R v) t = y" and [simp]: "sum b u = 1" "sum (\v. b v *\<^sub>R v) u = y" by (auto simp: affine_hull_finite \finite t\ \finite u\) define c where "c x = (if x \ t then a x else if x \ u then -(b x) else 0)" for x have [simp]: "s \ t = t" "s \ - t \ u = u" using assms by auto have "sum c s = 0" by (simp add: c_def comm_monoid_add_class.sum.If_cases \finite s\ sum_negf) moreover have "\ (\v\s. c v = 0)" by (metis (no_types) IntD1 \s \ t = t\ a1 c_def sum.neutral zero_neq_one) moreover have "(\v\s. c v *\<^sub>R v) = 0" by (simp add: c_def if_smult sum_negf comm_monoid_add_class.sum.If_cases \finite s\) ultimately have False using assms \finite s\ by (auto simp: affine_dependent_explicit) } then show ?thesis by blast qed end \ No newline at end of file diff --git a/src/HOL/Analysis/Bochner_Integration.thy b/src/HOL/Analysis/Bochner_Integration.thy --- a/src/HOL/Analysis/Bochner_Integration.thy +++ b/src/HOL/Analysis/Bochner_Integration.thy @@ -1,3364 +1,3372 @@ (* Title: HOL/Analysis/Bochner_Integration.thy Author: Johannes Hölzl, TU München *) section \Bochner Integration for Vector-Valued Functions\ theory Bochner_Integration imports Finite_Product_Measure begin text \ In the following development of the Bochner integral we use second countable topologies instead of separable spaces. A second countable topology is also separable. \ proposition borel_measurable_implies_sequence_metric: fixes f :: "'a \ 'b :: {metric_space, second_countable_topology}" assumes [measurable]: "f \ borel_measurable M" shows "\F. (\i. simple_function M (F i)) \ (\x\space M. (\i. F i x) \ f x) \ (\i. \x\space M. dist (F i x) z \ 2 * dist (f x) z)" proof - obtain D :: "'b set" where "countable D" and D: "\X. open X \ X \ {} \ \d\D. d \ X" by (erule countable_dense_setE) define e where "e = from_nat_into D" { fix n x obtain d where "d \ D" and d: "d \ ball x (1 / Suc n)" using D[of "ball x (1 / Suc n)"] by auto from \d \ D\ D[of UNIV] \countable D\ obtain i where "d = e i" unfolding e_def by (auto dest: from_nat_into_surj) with d have "\i. dist x (e i) < 1 / Suc n" by auto } note e = this define A where [abs_def]: "A m n = {x\space M. dist (f x) (e n) < 1 / (Suc m) \ 1 / (Suc m) \ dist (f x) z}" for m n define B where [abs_def]: "B m = disjointed (A m)" for m define m where [abs_def]: "m N x = Max {m. m \ N \ x \ (\n\N. B m n)}" for N x define F where [abs_def]: "F N x = (if (\m\N. x \ (\n\N. B m n)) \ (\n\N. x \ B (m N x) n) then e (LEAST n. x \ B (m N x) n) else z)" for N x have B_imp_A[intro, simp]: "\x m n. x \ B m n \ x \ A m n" using disjointed_subset[of "A m" for m] unfolding B_def by auto { fix m have "\n. A m n \ sets M" by (auto simp: A_def) then have "\n. B m n \ sets M" using sets.range_disjointed_sets[of "A m" M] by (auto simp: B_def) } note this[measurable] { fix N i x assume "\m\N. x \ (\n\N. B m n)" then have "m N x \ {m::nat. m \ N \ x \ (\n\N. B m n)}" unfolding m_def by (intro Max_in) auto then have "m N x \ N" "\n\N. x \ B (m N x) n" by auto } note m = this { fix j N i x assume "j \ N" "i \ N" "x \ B j i" then have "j \ m N x" unfolding m_def by (intro Max_ge) auto } note m_upper = this show ?thesis unfolding simple_function_def proof (safe intro!: exI[of _ F]) have [measurable]: "\i. F i \ borel_measurable M" unfolding F_def m_def by measurable show "\x i. F i -` {x} \ space M \ sets M" by measurable { fix i { fix n x assume "x \ B (m i x) n" then have "(LEAST n. x \ B (m i x) n) \ n" by (intro Least_le) also assume "n \ i" finally have "(LEAST n. x \ B (m i x) n) \ i" . } then have "F i ` space M \ {z} \ e ` {.. i}" by (auto simp: F_def) then show "finite (F i ` space M)" by (rule finite_subset) auto } { fix N i n x assume "i \ N" "n \ N" "x \ B i n" then have 1: "\m\N. x \ (\n\N. B m n)" by auto from m[OF this] obtain n where n: "m N x \ N" "n \ N" "x \ B (m N x) n" by auto moreover define L where "L = (LEAST n. x \ B (m N x) n)" have "dist (f x) (e L) < 1 / Suc (m N x)" proof - have "x \ B (m N x) L" using n(3) unfolding L_def by (rule LeastI) then have "x \ A (m N x) L" by auto then show ?thesis unfolding A_def by simp qed ultimately have "dist (f x) (F N x) < 1 / Suc (m N x)" by (auto simp add: F_def L_def) } note * = this fix x assume "x \ space M" show "(\i. F i x) \ f x" proof cases assume "f x = z" then have "\i n. x \ A i n" unfolding A_def by auto then have "\i. F i x = z" by (auto simp: F_def) then show ?thesis using \f x = z\ by auto next assume "f x \ z" show ?thesis proof (rule tendstoI) fix e :: real assume "0 < e" with \f x \ z\ obtain n where "1 / Suc n < e" "1 / Suc n < dist (f x) z" by (metis dist_nz order_less_trans neq_iff nat_approx_posE) with \x\space M\ \f x \ z\ have "x \ (\i. B n i)" unfolding A_def B_def UN_disjointed_eq using e by auto then obtain i where i: "x \ B n i" by auto show "eventually (\i. dist (F i x) (f x) < e) sequentially" using eventually_ge_at_top[of "max n i"] proof eventually_elim fix j assume j: "max n i \ j" with i have "dist (f x) (F j x) < 1 / Suc (m j x)" by (intro *[OF _ _ i]) auto also have "\ \ 1 / Suc n" using j m_upper[OF _ _ i] by (auto simp: field_simps) also note \1 / Suc n < e\ finally show "dist (F j x) (f x) < e" by (simp add: less_imp_le dist_commute) qed qed qed fix i { fix n m assume "x \ A n m" then have "dist (e m) (f x) + dist (f x) z \ 2 * dist (f x) z" unfolding A_def by (auto simp: dist_commute) also have "dist (e m) z \ dist (e m) (f x) + dist (f x) z" by (rule dist_triangle) finally (xtrans) have "dist (e m) z \ 2 * dist (f x) z" . } then show "dist (F i x) z \ 2 * dist (f x) z" unfolding F_def apply auto apply (rule LeastI2) apply auto done qed qed lemma fixes f :: "'a \ 'b::semiring_1" assumes "finite A" shows sum_mult_indicator[simp]: "(\x \ A. f x * indicator (B x) (g x)) = (\x\{x\A. g x \ B x}. f x)" and sum_indicator_mult[simp]: "(\x \ A. indicator (B x) (g x) * f x) = (\x\{x\A. g x \ B x}. f x)" unfolding indicator_def using assms by (auto intro!: sum.mono_neutral_cong_right split: if_split_asm) lemma borel_measurable_induct_real[consumes 2, case_names set mult add seq]: fixes P :: "('a \ real) \ bool" assumes u: "u \ borel_measurable M" "\x. 0 \ u x" assumes set: "\A. A \ sets M \ P (indicator A)" assumes mult: "\u c. 0 \ c \ u \ borel_measurable M \ (\x. 0 \ u x) \ P u \ P (\x. c * u x)" assumes add: "\u v. u \ borel_measurable M \ (\x. 0 \ u x) \ P u \ v \ borel_measurable M \ (\x. 0 \ v x) \ (\x. x \ space M \ u x = 0 \ v x = 0) \ P v \ P (\x. v x + u x)" assumes seq: "\U. (\i. U i \ borel_measurable M) \ (\i x. 0 \ U i x) \ (\i. P (U i)) \ incseq U \ (\x. x \ space M \ (\i. U i x) \ u x) \ P u" shows "P u" proof - have "(\x. ennreal (u x)) \ borel_measurable M" using u by auto from borel_measurable_implies_simple_function_sequence'[OF this] obtain U where U: "\i. simple_function M (U i)" "incseq U" "\i x. U i x < top" and sup: "\x. (SUP i. U i x) = ennreal (u x)" by blast define U' where [abs_def]: "U' i x = indicator (space M) x * enn2real (U i x)" for i x then have U'_sf[measurable]: "\i. simple_function M (U' i)" using U by (auto intro!: simple_function_compose1[where g=enn2real]) show "P u" proof (rule seq) show U': "U' i \ borel_measurable M" "\x. 0 \ U' i x" for i using U by (auto intro: borel_measurable_simple_function intro!: borel_measurable_enn2real borel_measurable_times simp: U'_def zero_le_mult_iff) show "incseq U'" using U(2,3) by (auto simp: incseq_def le_fun_def image_iff eq_commute U'_def indicator_def enn2real_mono) fix x assume x: "x \ space M" have "(\i. U i x) \ (SUP i. U i x)" using U(2) by (intro LIMSEQ_SUP) (auto simp: incseq_def le_fun_def) moreover have "(\i. U i x) = (\i. ennreal (U' i x))" using x U(3) by (auto simp: fun_eq_iff U'_def image_iff eq_commute) moreover have "(SUP i. U i x) = ennreal (u x)" using sup u(2) by (simp add: max_def) ultimately show "(\i. U' i x) \ u x" using u U' by simp next fix i have "U' i ` space M \ enn2real ` (U i ` space M)" "finite (U i ` space M)" unfolding U'_def using U(1) by (auto dest: simple_functionD) then have fin: "finite (U' i ` space M)" by (metis finite_subset finite_imageI) moreover have "\z. {y. U' i z = y \ y \ U' i ` space M \ z \ space M} = (if z \ space M then {U' i z} else {})" by auto ultimately have U': "(\z. \y\U' i`space M. y * indicator {x\space M. U' i x = y} z) = U' i" by (simp add: U'_def fun_eq_iff) have "\x. x \ U' i ` space M \ 0 \ x" by (auto simp: U'_def) with fin have "P (\z. \y\U' i`space M. y * indicator {x\space M. U' i x = y} z)" proof induct case empty from set[of "{}"] show ?case by (simp add: indicator_def[abs_def]) next case (insert x F) - then show ?case - by (auto intro!: add mult set sum_nonneg split: split_indicator split_indicator_asm - simp del: sum_mult_indicator simp: sum_nonneg_eq_0_iff) + from insert.prems have nonneg: "x \ 0" "\y. y \ F \ y \ 0" + by simp_all + hence *: "P (\xa. x * indicat_real {x' \ space M. U' i x' = x} xa)" + by (intro mult set) auto + have "P (\z. x * indicat_real {x' \ space M. U' i x' = x} z + + (\y\F. y * indicat_real {x \ space M. U' i x = y} z))" + using insert(1-3) + by (intro add * sum_nonneg mult_nonneg_nonneg) + (auto simp: nonneg indicator_def sum_nonneg_eq_0_iff) + thus ?case + using insert.hyps by (subst sum.insert) auto qed with U' show "P (U' i)" by simp qed qed lemma scaleR_cong_right: fixes x :: "'a :: real_vector" shows "(x \ 0 \ r = p) \ r *\<^sub>R x = p *\<^sub>R x" by (cases "x = 0") auto inductive simple_bochner_integrable :: "'a measure \ ('a \ 'b::real_vector) \ bool" for M f where "simple_function M f \ emeasure M {y\space M. f y \ 0} \ \ \ simple_bochner_integrable M f" lemma simple_bochner_integrable_compose2: assumes p_0: "p 0 0 = 0" shows "simple_bochner_integrable M f \ simple_bochner_integrable M g \ simple_bochner_integrable M (\x. p (f x) (g x))" proof (safe intro!: simple_bochner_integrable.intros elim!: simple_bochner_integrable.cases del: notI) assume sf: "simple_function M f" "simple_function M g" then show "simple_function M (\x. p (f x) (g x))" by (rule simple_function_compose2) from sf have [measurable]: "f \ measurable M (count_space UNIV)" "g \ measurable M (count_space UNIV)" by (auto intro: measurable_simple_function) assume fin: "emeasure M {y \ space M. f y \ 0} \ \" "emeasure M {y \ space M. g y \ 0} \ \" have "emeasure M {x\space M. p (f x) (g x) \ 0} \ emeasure M ({x\space M. f x \ 0} \ {x\space M. g x \ 0})" by (intro emeasure_mono) (auto simp: p_0) also have "\ \ emeasure M {x\space M. f x \ 0} + emeasure M {x\space M. g x \ 0}" by (intro emeasure_subadditive) auto finally show "emeasure M {y \ space M. p (f y) (g y) \ 0} \ \" using fin by (auto simp: top_unique) qed lemma simple_function_finite_support: assumes f: "simple_function M f" and fin: "(\\<^sup>+x. f x \M) < \" and nn: "\x. 0 \ f x" shows "emeasure M {x\space M. f x \ 0} \ \" proof cases from f have meas[measurable]: "f \ borel_measurable M" by (rule borel_measurable_simple_function) assume non_empty: "\x\space M. f x \ 0" define m where "m = Min (f`space M - {0})" have "m \ f`space M - {0}" unfolding m_def using f non_empty by (intro Min_in) (auto simp: simple_function_def) then have m: "0 < m" using nn by (auto simp: less_le) from m have "m * emeasure M {x\space M. 0 \ f x} = (\\<^sup>+x. m * indicator {x\space M. 0 \ f x} x \M)" using f by (intro nn_integral_cmult_indicator[symmetric]) auto also have "\ \ (\\<^sup>+x. f x \M)" using AE_space proof (intro nn_integral_mono_AE, eventually_elim) fix x assume "x \ space M" with nn show "m * indicator {x \ space M. 0 \ f x} x \ f x" using f by (auto split: split_indicator simp: simple_function_def m_def) qed also note \\ < \\ finally show ?thesis using m by (auto simp: ennreal_mult_less_top) next assume "\ (\x\space M. f x \ 0)" with nn have *: "{x\space M. f x \ 0} = {}" by auto show ?thesis unfolding * by simp qed lemma simple_bochner_integrableI_bounded: assumes f: "simple_function M f" and fin: "(\\<^sup>+x. norm (f x) \M) < \" shows "simple_bochner_integrable M f" proof have "emeasure M {y \ space M. ennreal (norm (f y)) \ 0} \ \" proof (rule simple_function_finite_support) show "simple_function M (\x. ennreal (norm (f x)))" using f by (rule simple_function_compose1) show "(\\<^sup>+ y. ennreal (norm (f y)) \M) < \" by fact qed simp then show "emeasure M {y \ space M. f y \ 0} \ \" by simp qed fact definition\<^marker>\tag important\ simple_bochner_integral :: "'a measure \ ('a \ 'b::real_vector) \ 'b" where "simple_bochner_integral M f = (\y\f`space M. measure M {x\space M. f x = y} *\<^sub>R y)" proposition simple_bochner_integral_partition: assumes f: "simple_bochner_integrable M f" and g: "simple_function M g" assumes sub: "\x y. x \ space M \ y \ space M \ g x = g y \ f x = f y" assumes v: "\x. x \ space M \ f x = v (g x)" shows "simple_bochner_integral M f = (\y\g ` space M. measure M {x\space M. g x = y} *\<^sub>R v y)" (is "_ = ?r") proof - from f g have [simp]: "finite (f`space M)" "finite (g`space M)" by (auto simp: simple_function_def elim: simple_bochner_integrable.cases) from f have [measurable]: "f \ measurable M (count_space UNIV)" by (auto intro: measurable_simple_function elim: simple_bochner_integrable.cases) from g have [measurable]: "g \ measurable M (count_space UNIV)" by (auto intro: measurable_simple_function elim: simple_bochner_integrable.cases) { fix y assume "y \ space M" then have "f ` space M \ {i. \x\space M. i = f x \ g y = g x} = {v (g y)}" by (auto cong: sub simp: v[symmetric]) } note eq = this have "simple_bochner_integral M f = (\y\f`space M. (\z\g`space M. if \x\space M. y = f x \ z = g x then measure M {x\space M. g x = z} else 0) *\<^sub>R y)" unfolding simple_bochner_integral_def proof (safe intro!: sum.cong scaleR_cong_right) fix y assume y: "y \ space M" "f y \ 0" have [simp]: "g ` space M \ {z. \x\space M. f y = f x \ z = g x} = {z. \x\space M. f y = f x \ z = g x}" by auto have eq:"{x \ space M. f x = f y} = (\i\{z. \x\space M. f y = f x \ z = g x}. {x \ space M. g x = i})" by (auto simp: eq_commute cong: sub rev_conj_cong) have "finite (g`space M)" by simp then have "finite {z. \x\space M. f y = f x \ z = g x}" by (rule rev_finite_subset) auto moreover { fix x assume "x \ space M" "f x = f y" then have "x \ space M" "f x \ 0" using y by auto then have "emeasure M {y \ space M. g y = g x} \ emeasure M {y \ space M. f y \ 0}" by (auto intro!: emeasure_mono cong: sub) then have "emeasure M {xa \ space M. g xa = g x} < \" using f by (auto simp: simple_bochner_integrable.simps less_top) } ultimately show "measure M {x \ space M. f x = f y} = (\z\g ` space M. if \x\space M. f y = f x \ z = g x then measure M {x \ space M. g x = z} else 0)" apply (simp add: sum.If_cases eq) apply (subst measure_finite_Union[symmetric]) apply (auto simp: disjoint_family_on_def less_top) done qed also have "\ = (\y\f`space M. (\z\g`space M. if \x\space M. y = f x \ z = g x then measure M {x\space M. g x = z} *\<^sub>R y else 0))" by (auto intro!: sum.cong simp: scaleR_sum_left) also have "\ = ?r" by (subst sum.swap) (auto intro!: sum.cong simp: sum.If_cases scaleR_sum_right[symmetric] eq) finally show "simple_bochner_integral M f = ?r" . qed lemma simple_bochner_integral_add: assumes f: "simple_bochner_integrable M f" and g: "simple_bochner_integrable M g" shows "simple_bochner_integral M (\x. f x + g x) = simple_bochner_integral M f + simple_bochner_integral M g" proof - from f g have "simple_bochner_integral M (\x. f x + g x) = (\y\(\x. (f x, g x)) ` space M. measure M {x \ space M. (f x, g x) = y} *\<^sub>R (fst y + snd y))" by (intro simple_bochner_integral_partition) (auto simp: simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases) moreover from f g have "simple_bochner_integral M f = (\y\(\x. (f x, g x)) ` space M. measure M {x \ space M. (f x, g x) = y} *\<^sub>R fst y)" by (intro simple_bochner_integral_partition) (auto simp: simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases) moreover from f g have "simple_bochner_integral M g = (\y\(\x. (f x, g x)) ` space M. measure M {x \ space M. (f x, g x) = y} *\<^sub>R snd y)" by (intro simple_bochner_integral_partition) (auto simp: simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases) ultimately show ?thesis by (simp add: sum.distrib[symmetric] scaleR_add_right) qed lemma simple_bochner_integral_linear: assumes "linear f" assumes g: "simple_bochner_integrable M g" shows "simple_bochner_integral M (\x. f (g x)) = f (simple_bochner_integral M g)" proof - interpret linear f by fact from g have "simple_bochner_integral M (\x. f (g x)) = (\y\g ` space M. measure M {x \ space M. g x = y} *\<^sub>R f y)" by (intro simple_bochner_integral_partition) (auto simp: simple_bochner_integrable_compose2[where p="\x y. f x"] elim: simple_bochner_integrable.cases) also have "\ = f (simple_bochner_integral M g)" by (simp add: simple_bochner_integral_def sum scale) finally show ?thesis . qed lemma simple_bochner_integral_minus: assumes f: "simple_bochner_integrable M f" shows "simple_bochner_integral M (\x. - f x) = - simple_bochner_integral M f" proof - from linear_uminus f show ?thesis by (rule simple_bochner_integral_linear) qed lemma simple_bochner_integral_diff: assumes f: "simple_bochner_integrable M f" and g: "simple_bochner_integrable M g" shows "simple_bochner_integral M (\x. f x - g x) = simple_bochner_integral M f - simple_bochner_integral M g" unfolding diff_conv_add_uminus using f g by (subst simple_bochner_integral_add) (auto simp: simple_bochner_integral_minus simple_bochner_integrable_compose2[where p="\x y. - y"]) lemma simple_bochner_integral_norm_bound: assumes f: "simple_bochner_integrable M f" shows "norm (simple_bochner_integral M f) \ simple_bochner_integral M (\x. norm (f x))" proof - have "norm (simple_bochner_integral M f) \ (\y\f ` space M. norm (measure M {x \ space M. f x = y} *\<^sub>R y))" unfolding simple_bochner_integral_def by (rule norm_sum) also have "\ = (\y\f ` space M. measure M {x \ space M. f x = y} *\<^sub>R norm y)" by simp also have "\ = simple_bochner_integral M (\x. norm (f x))" using f by (intro simple_bochner_integral_partition[symmetric]) (auto intro: f simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases) finally show ?thesis . qed lemma simple_bochner_integral_nonneg[simp]: fixes f :: "'a \ real" shows "(\x. 0 \ f x) \ 0 \ simple_bochner_integral M f" by (force simp add: simple_bochner_integral_def intro: sum_nonneg) lemma simple_bochner_integral_eq_nn_integral: assumes f: "simple_bochner_integrable M f" "\x. 0 \ f x" shows "simple_bochner_integral M f = (\\<^sup>+x. f x \M)" proof - { fix x y z have "(x \ 0 \ y = z) \ ennreal x * y = ennreal x * z" by (cases "x = 0") (auto simp: zero_ennreal_def[symmetric]) } note ennreal_cong_mult = this have [measurable]: "f \ borel_measurable M" using f(1) by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) { fix y assume y: "y \ space M" "f y \ 0" have "ennreal (measure M {x \ space M. f x = f y}) = emeasure M {x \ space M. f x = f y}" proof (rule emeasure_eq_ennreal_measure[symmetric]) have "emeasure M {x \ space M. f x = f y} \ emeasure M {x \ space M. f x \ 0}" using y by (intro emeasure_mono) auto with f show "emeasure M {x \ space M. f x = f y} \ top" by (auto simp: simple_bochner_integrable.simps top_unique) qed moreover have "{x \ space M. f x = f y} = (\x. ennreal (f x)) -` {ennreal (f y)} \ space M" using f by auto ultimately have "ennreal (measure M {x \ space M. f x = f y}) = emeasure M ((\x. ennreal (f x)) -` {ennreal (f y)} \ space M)" by simp } with f have "simple_bochner_integral M f = (\\<^sup>Sx. f x \M)" unfolding simple_integral_def by (subst simple_bochner_integral_partition[OF f(1), where g="\x. ennreal (f x)" and v=enn2real]) (auto intro: f simple_function_compose1 elim: simple_bochner_integrable.cases intro!: sum.cong ennreal_cong_mult simp: ac_simps ennreal_mult simp flip: sum_ennreal) also have "\ = (\\<^sup>+x. f x \M)" using f by (intro nn_integral_eq_simple_integral[symmetric]) (auto simp: simple_function_compose1 simple_bochner_integrable.simps) finally show ?thesis . qed lemma simple_bochner_integral_bounded: fixes f :: "'a \ 'b::{real_normed_vector, second_countable_topology}" assumes f[measurable]: "f \ borel_measurable M" assumes s: "simple_bochner_integrable M s" and t: "simple_bochner_integrable M t" shows "ennreal (norm (simple_bochner_integral M s - simple_bochner_integral M t)) \ (\\<^sup>+ x. norm (f x - s x) \M) + (\\<^sup>+ x. norm (f x - t x) \M)" (is "ennreal (norm (?s - ?t)) \ ?S + ?T") proof - have [measurable]: "s \ borel_measurable M" "t \ borel_measurable M" using s t by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) have "ennreal (norm (?s - ?t)) = norm (simple_bochner_integral M (\x. s x - t x))" using s t by (subst simple_bochner_integral_diff) auto also have "\ \ simple_bochner_integral M (\x. norm (s x - t x))" using simple_bochner_integrable_compose2[of "(-)" M "s" "t"] s t by (auto intro!: simple_bochner_integral_norm_bound) also have "\ = (\\<^sup>+x. norm (s x - t x) \M)" using simple_bochner_integrable_compose2[of "\x y. norm (x - y)" M "s" "t"] s t by (auto intro!: simple_bochner_integral_eq_nn_integral) also have "\ \ (\\<^sup>+x. ennreal (norm (f x - s x)) + ennreal (norm (f x - t x)) \M)" by (auto intro!: nn_integral_mono simp flip: ennreal_plus) (metis (erased, hide_lams) add_diff_cancel_left add_diff_eq diff_add_eq order_trans norm_minus_commute norm_triangle_ineq4 order_refl) also have "\ = ?S + ?T" by (rule nn_integral_add) auto finally show ?thesis . qed inductive has_bochner_integral :: "'a measure \ ('a \ 'b) \ 'b::{real_normed_vector, second_countable_topology} \ bool" for M f x where "f \ borel_measurable M \ (\i. simple_bochner_integrable M (s i)) \ (\i. \\<^sup>+x. norm (f x - s i x) \M) \ 0 \ (\i. simple_bochner_integral M (s i)) \ x \ has_bochner_integral M f x" lemma has_bochner_integral_cong: assumes "M = N" "\x. x \ space N \ f x = g x" "x = y" shows "has_bochner_integral M f x \ has_bochner_integral N g y" unfolding has_bochner_integral.simps assms(1,3) using assms(2) by (simp cong: measurable_cong_simp nn_integral_cong_simp) lemma has_bochner_integral_cong_AE: "f \ borel_measurable M \ g \ borel_measurable M \ (AE x in M. f x = g x) \ has_bochner_integral M f x \ has_bochner_integral M g x" unfolding has_bochner_integral.simps by (intro arg_cong[where f=Ex] ext conj_cong rev_conj_cong refl arg_cong[where f="\x. x \ 0"] nn_integral_cong_AE) auto lemma borel_measurable_has_bochner_integral: "has_bochner_integral M f x \ f \ borel_measurable M" by (rule has_bochner_integral.cases) lemma borel_measurable_has_bochner_integral'[measurable_dest]: "has_bochner_integral M f x \ g \ measurable N M \ (\x. f (g x)) \ borel_measurable N" using borel_measurable_has_bochner_integral[measurable] by measurable lemma has_bochner_integral_simple_bochner_integrable: "simple_bochner_integrable M f \ has_bochner_integral M f (simple_bochner_integral M f)" by (rule has_bochner_integral.intros[where s="\_. f"]) (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases simp: zero_ennreal_def[symmetric]) lemma has_bochner_integral_real_indicator: assumes [measurable]: "A \ sets M" and A: "emeasure M A < \" shows "has_bochner_integral M (indicator A) (measure M A)" proof - have sbi: "simple_bochner_integrable M (indicator A::'a \ real)" proof have "{y \ space M. (indicator A y::real) \ 0} = A" using sets.sets_into_space[OF \A\sets M\] by (auto split: split_indicator) then show "emeasure M {y \ space M. (indicator A y::real) \ 0} \ \" using A by auto qed (rule simple_function_indicator assms)+ moreover have "simple_bochner_integral M (indicator A) = measure M A" using simple_bochner_integral_eq_nn_integral[OF sbi] A by (simp add: ennreal_indicator emeasure_eq_ennreal_measure) ultimately show ?thesis by (metis has_bochner_integral_simple_bochner_integrable) qed lemma has_bochner_integral_add[intro]: "has_bochner_integral M f x \ has_bochner_integral M g y \ has_bochner_integral M (\x. f x + g x) (x + y)" proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases) fix sf sg assume f_sf: "(\i. \\<^sup>+ x. norm (f x - sf i x) \M) \ 0" assume g_sg: "(\i. \\<^sup>+ x. norm (g x - sg i x) \M) \ 0" assume sf: "\i. simple_bochner_integrable M (sf i)" and sg: "\i. simple_bochner_integrable M (sg i)" then have [measurable]: "\i. sf i \ borel_measurable M" "\i. sg i \ borel_measurable M" by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) assume [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" show "\i. simple_bochner_integrable M (\x. sf i x + sg i x)" using sf sg by (simp add: simple_bochner_integrable_compose2) show "(\i. \\<^sup>+ x. (norm (f x + g x - (sf i x + sg i x))) \M) \ 0" (is "?f \ 0") proof (rule tendsto_sandwich) show "eventually (\n. 0 \ ?f n) sequentially" "(\_. 0) \ 0" by auto show "eventually (\i. ?f i \ (\\<^sup>+ x. (norm (f x - sf i x)) \M) + \\<^sup>+ x. (norm (g x - sg i x)) \M) sequentially" (is "eventually (\i. ?f i \ ?g i) sequentially") proof (intro always_eventually allI) fix i have "?f i \ (\\<^sup>+ x. (norm (f x - sf i x)) + ennreal (norm (g x - sg i x)) \M)" by (auto intro!: nn_integral_mono norm_diff_triangle_ineq simp flip: ennreal_plus) also have "\ = ?g i" by (intro nn_integral_add) auto finally show "?f i \ ?g i" . qed show "?g \ 0" using tendsto_add[OF f_sf g_sg] by simp qed qed (auto simp: simple_bochner_integral_add tendsto_add) lemma has_bochner_integral_bounded_linear: assumes "bounded_linear T" shows "has_bochner_integral M f x \ has_bochner_integral M (\x. T (f x)) (T x)" proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases) interpret T: bounded_linear T by fact have [measurable]: "T \ borel_measurable borel" by (intro borel_measurable_continuous_onI T.continuous_on continuous_on_id) assume [measurable]: "f \ borel_measurable M" then show "(\x. T (f x)) \ borel_measurable M" by auto fix s assume f_s: "(\i. \\<^sup>+ x. norm (f x - s i x) \M) \ 0" assume s: "\i. simple_bochner_integrable M (s i)" then show "\i. simple_bochner_integrable M (\x. T (s i x))" by (auto intro: simple_bochner_integrable_compose2 T.zero) have [measurable]: "\i. s i \ borel_measurable M" using s by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) obtain K where K: "K > 0" "\x i. norm (T (f x) - T (s i x)) \ norm (f x - s i x) * K" using T.pos_bounded by (auto simp: T.diff[symmetric]) show "(\i. \\<^sup>+ x. norm (T (f x) - T (s i x)) \M) \ 0" (is "?f \ 0") proof (rule tendsto_sandwich) show "eventually (\n. 0 \ ?f n) sequentially" "(\_. 0) \ 0" by auto show "eventually (\i. ?f i \ K * (\\<^sup>+ x. norm (f x - s i x) \M)) sequentially" (is "eventually (\i. ?f i \ ?g i) sequentially") proof (intro always_eventually allI) fix i have "?f i \ (\\<^sup>+ x. ennreal K * norm (f x - s i x) \M)" using K by (intro nn_integral_mono) (auto simp: ac_simps ennreal_mult[symmetric]) also have "\ = ?g i" using K by (intro nn_integral_cmult) auto finally show "?f i \ ?g i" . qed show "?g \ 0" using ennreal_tendsto_cmult[OF _ f_s] by simp qed assume "(\i. simple_bochner_integral M (s i)) \ x" with s show "(\i. simple_bochner_integral M (\x. T (s i x))) \ T x" by (auto intro!: T.tendsto simp: simple_bochner_integral_linear T.linear_axioms) qed lemma has_bochner_integral_zero[intro]: "has_bochner_integral M (\x. 0) 0" by (auto intro!: has_bochner_integral.intros[where s="\_ _. 0"] simp: zero_ennreal_def[symmetric] simple_bochner_integrable.simps simple_bochner_integral_def image_constant_conv) lemma has_bochner_integral_scaleR_left[intro]: "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. f x *\<^sub>R c) (x *\<^sub>R c)" by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_scaleR_left]) lemma has_bochner_integral_scaleR_right[intro]: "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. c *\<^sub>R f x) (c *\<^sub>R x)" by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_scaleR_right]) lemma has_bochner_integral_mult_left[intro]: fixes c :: "_::{real_normed_algebra,second_countable_topology}" shows "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. f x * c) (x * c)" by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_mult_left]) lemma has_bochner_integral_mult_right[intro]: fixes c :: "_::{real_normed_algebra,second_countable_topology}" shows "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. c * f x) (c * x)" by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_mult_right]) lemmas has_bochner_integral_divide = has_bochner_integral_bounded_linear[OF bounded_linear_divide] lemma has_bochner_integral_divide_zero[intro]: fixes c :: "_::{real_normed_field, field, second_countable_topology}" shows "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. f x / c) (x / c)" using has_bochner_integral_divide by (cases "c = 0") auto lemma has_bochner_integral_inner_left[intro]: "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. f x \ c) (x \ c)" by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_inner_left]) lemma has_bochner_integral_inner_right[intro]: "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. c \ f x) (c \ x)" by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_inner_right]) lemmas has_bochner_integral_minus = has_bochner_integral_bounded_linear[OF bounded_linear_minus[OF bounded_linear_ident]] lemmas has_bochner_integral_Re = has_bochner_integral_bounded_linear[OF bounded_linear_Re] lemmas has_bochner_integral_Im = has_bochner_integral_bounded_linear[OF bounded_linear_Im] lemmas has_bochner_integral_cnj = has_bochner_integral_bounded_linear[OF bounded_linear_cnj] lemmas has_bochner_integral_of_real = has_bochner_integral_bounded_linear[OF bounded_linear_of_real] lemmas has_bochner_integral_fst = has_bochner_integral_bounded_linear[OF bounded_linear_fst] lemmas has_bochner_integral_snd = has_bochner_integral_bounded_linear[OF bounded_linear_snd] lemma has_bochner_integral_indicator: "A \ sets M \ emeasure M A < \ \ has_bochner_integral M (\x. indicator A x *\<^sub>R c) (measure M A *\<^sub>R c)" by (intro has_bochner_integral_scaleR_left has_bochner_integral_real_indicator) lemma has_bochner_integral_diff: "has_bochner_integral M f x \ has_bochner_integral M g y \ has_bochner_integral M (\x. f x - g x) (x - y)" unfolding diff_conv_add_uminus by (intro has_bochner_integral_add has_bochner_integral_minus) lemma has_bochner_integral_sum: "(\i. i \ I \ has_bochner_integral M (f i) (x i)) \ has_bochner_integral M (\x. \i\I. f i x) (\i\I. x i)" by (induct I rule: infinite_finite_induct) auto proposition has_bochner_integral_implies_finite_norm: "has_bochner_integral M f x \ (\\<^sup>+x. norm (f x) \M) < \" proof (elim has_bochner_integral.cases) fix s v assume [measurable]: "f \ borel_measurable M" and s: "\i. simple_bochner_integrable M (s i)" and lim_0: "(\i. \\<^sup>+ x. ennreal (norm (f x - s i x)) \M) \ 0" from order_tendstoD[OF lim_0, of "\"] obtain i where f_s_fin: "(\\<^sup>+ x. ennreal (norm (f x - s i x)) \M) < \" by (auto simp: eventually_sequentially) have [measurable]: "\i. s i \ borel_measurable M" using s by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) define m where "m = (if space M = {} then 0 else Max ((\x. norm (s i x))`space M))" have "finite (s i ` space M)" using s by (auto simp: simple_function_def simple_bochner_integrable.simps) then have "finite (norm ` s i ` space M)" by (rule finite_imageI) then have "\x. x \ space M \ norm (s i x) \ m" "0 \ m" by (auto simp: m_def image_comp comp_def Max_ge_iff) then have "(\\<^sup>+x. norm (s i x) \M) \ (\\<^sup>+x. ennreal m * indicator {x\space M. s i x \ 0} x \M)" by (auto split: split_indicator intro!: Max_ge nn_integral_mono simp:) also have "\ < \" using s by (subst nn_integral_cmult_indicator) (auto simp: \0 \ m\ simple_bochner_integrable.simps ennreal_mult_less_top less_top) finally have s_fin: "(\\<^sup>+x. norm (s i x) \M) < \" . have "(\\<^sup>+ x. norm (f x) \M) \ (\\<^sup>+ x. ennreal (norm (f x - s i x)) + ennreal (norm (s i x)) \M)" by (auto intro!: nn_integral_mono simp flip: ennreal_plus) (metis add.commute norm_triangle_sub) also have "\ = (\\<^sup>+x. norm (f x - s i x) \M) + (\\<^sup>+x. norm (s i x) \M)" by (rule nn_integral_add) auto also have "\ < \" using s_fin f_s_fin by auto finally show "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" . qed proposition has_bochner_integral_norm_bound: assumes i: "has_bochner_integral M f x" shows "norm x \ (\\<^sup>+x. norm (f x) \M)" using assms proof fix s assume x: "(\i. simple_bochner_integral M (s i)) \ x" (is "?s \ x") and s[simp]: "\i. simple_bochner_integrable M (s i)" and lim: "(\i. \\<^sup>+ x. ennreal (norm (f x - s i x)) \M) \ 0" and f[measurable]: "f \ borel_measurable M" have [measurable]: "\i. s i \ borel_measurable M" using s by (auto simp: simple_bochner_integrable.simps intro: borel_measurable_simple_function) show "norm x \ (\\<^sup>+x. norm (f x) \M)" proof (rule LIMSEQ_le) show "(\i. ennreal (norm (?s i))) \ norm x" using x by (auto simp: tendsto_ennreal_iff intro: tendsto_intros) show "\N. \n\N. norm (?s n) \ (\\<^sup>+x. norm (f x - s n x) \M) + (\\<^sup>+x. norm (f x) \M)" (is "\N. \n\N. _ \ ?t n") proof (intro exI allI impI) fix n have "ennreal (norm (?s n)) \ simple_bochner_integral M (\x. norm (s n x))" by (auto intro!: simple_bochner_integral_norm_bound) also have "\ = (\\<^sup>+x. norm (s n x) \M)" by (intro simple_bochner_integral_eq_nn_integral) (auto intro: s simple_bochner_integrable_compose2) also have "\ \ (\\<^sup>+x. ennreal (norm (f x - s n x)) + norm (f x) \M)" by (auto intro!: nn_integral_mono simp flip: ennreal_plus) (metis add.commute norm_minus_commute norm_triangle_sub) also have "\ = ?t n" by (rule nn_integral_add) auto finally show "norm (?s n) \ ?t n" . qed have "?t \ 0 + (\\<^sup>+ x. ennreal (norm (f x)) \M)" using has_bochner_integral_implies_finite_norm[OF i] by (intro tendsto_add tendsto_const lim) then show "?t \ \\<^sup>+ x. ennreal (norm (f x)) \M" by simp qed qed lemma has_bochner_integral_eq: "has_bochner_integral M f x \ has_bochner_integral M f y \ x = y" proof (elim has_bochner_integral.cases) assume f[measurable]: "f \ borel_measurable M" fix s t assume "(\i. \\<^sup>+ x. norm (f x - s i x) \M) \ 0" (is "?S \ 0") assume "(\i. \\<^sup>+ x. norm (f x - t i x) \M) \ 0" (is "?T \ 0") assume s: "\i. simple_bochner_integrable M (s i)" assume t: "\i. simple_bochner_integrable M (t i)" have [measurable]: "\i. s i \ borel_measurable M" "\i. t i \ borel_measurable M" using s t by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) let ?s = "\i. simple_bochner_integral M (s i)" let ?t = "\i. simple_bochner_integral M (t i)" assume "?s \ x" "?t \ y" then have "(\i. norm (?s i - ?t i)) \ norm (x - y)" by (intro tendsto_intros) moreover have "(\i. ennreal (norm (?s i - ?t i))) \ ennreal 0" proof (rule tendsto_sandwich) show "eventually (\i. 0 \ ennreal (norm (?s i - ?t i))) sequentially" "(\_. 0) \ ennreal 0" by auto show "eventually (\i. norm (?s i - ?t i) \ ?S i + ?T i) sequentially" by (intro always_eventually allI simple_bochner_integral_bounded s t f) show "(\i. ?S i + ?T i) \ ennreal 0" using tendsto_add[OF \?S \ 0\ \?T \ 0\] by simp qed then have "(\i. norm (?s i - ?t i)) \ 0" by (simp flip: ennreal_0) ultimately have "norm (x - y) = 0" by (rule LIMSEQ_unique) then show "x = y" by simp qed lemma has_bochner_integralI_AE: assumes f: "has_bochner_integral M f x" and g: "g \ borel_measurable M" and ae: "AE x in M. f x = g x" shows "has_bochner_integral M g x" using f proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases) fix s assume "(\i. \\<^sup>+ x. ennreal (norm (f x - s i x)) \M) \ 0" also have "(\i. \\<^sup>+ x. ennreal (norm (f x - s i x)) \M) = (\i. \\<^sup>+ x. ennreal (norm (g x - s i x)) \M)" using ae by (intro ext nn_integral_cong_AE, eventually_elim) simp finally show "(\i. \\<^sup>+ x. ennreal (norm (g x - s i x)) \M) \ 0" . qed (auto intro: g) lemma has_bochner_integral_eq_AE: assumes f: "has_bochner_integral M f x" and g: "has_bochner_integral M g y" and ae: "AE x in M. f x = g x" shows "x = y" proof - from assms have "has_bochner_integral M g x" by (auto intro: has_bochner_integralI_AE) from this g show "x = y" by (rule has_bochner_integral_eq) qed lemma simple_bochner_integrable_restrict_space: fixes f :: "_ \ 'b::real_normed_vector" assumes \: "\ \ space M \ sets M" shows "simple_bochner_integrable (restrict_space M \) f \ simple_bochner_integrable M (\x. indicator \ x *\<^sub>R f x)" by (simp add: simple_bochner_integrable.simps space_restrict_space simple_function_restrict_space[OF \] emeasure_restrict_space[OF \] Collect_restrict indicator_eq_0_iff conj_left_commute) lemma simple_bochner_integral_restrict_space: fixes f :: "_ \ 'b::real_normed_vector" assumes \: "\ \ space M \ sets M" assumes f: "simple_bochner_integrable (restrict_space M \) f" shows "simple_bochner_integral (restrict_space M \) f = simple_bochner_integral M (\x. indicator \ x *\<^sub>R f x)" proof - have "finite ((\x. indicator \ x *\<^sub>R f x)`space M)" using f simple_bochner_integrable_restrict_space[OF \, of f] by (simp add: simple_bochner_integrable.simps simple_function_def) then show ?thesis by (auto simp: space_restrict_space measure_restrict_space[OF \(1)] le_infI2 simple_bochner_integral_def Collect_restrict split: split_indicator split_indicator_asm intro!: sum.mono_neutral_cong_left arg_cong2[where f=measure]) qed context notes [[inductive_internals]] begin inductive integrable for M f where "has_bochner_integral M f x \ integrable M f" end definition\<^marker>\tag important\ lebesgue_integral ("integral\<^sup>L") where "integral\<^sup>L M f = (if \x. has_bochner_integral M f x then THE x. has_bochner_integral M f x else 0)" syntax "_lebesgue_integral" :: "pttrn \ real \ 'a measure \ real" ("\((2 _./ _)/ \_)" [60,61] 110) translations "\ x. f \M" == "CONST lebesgue_integral M (\x. f)" syntax "_ascii_lebesgue_integral" :: "pttrn \ 'a measure \ real \ real" ("(3LINT (1_)/|(_)./ _)" [0,110,60] 60) translations "LINT x|M. f" == "CONST lebesgue_integral M (\x. f)" lemma has_bochner_integral_integral_eq: "has_bochner_integral M f x \ integral\<^sup>L M f = x" by (metis the_equality has_bochner_integral_eq lebesgue_integral_def) lemma has_bochner_integral_integrable: "integrable M f \ has_bochner_integral M f (integral\<^sup>L M f)" by (auto simp: has_bochner_integral_integral_eq integrable.simps) lemma has_bochner_integral_iff: "has_bochner_integral M f x \ integrable M f \ integral\<^sup>L M f = x" by (metis has_bochner_integral_integrable has_bochner_integral_integral_eq integrable.intros) lemma simple_bochner_integrable_eq_integral: "simple_bochner_integrable M f \ simple_bochner_integral M f = integral\<^sup>L M f" using has_bochner_integral_simple_bochner_integrable[of M f] by (simp add: has_bochner_integral_integral_eq) lemma not_integrable_integral_eq: "\ integrable M f \ integral\<^sup>L M f = 0" unfolding integrable.simps lebesgue_integral_def by (auto intro!: arg_cong[where f=The]) lemma integral_eq_cases: "integrable M f \ integrable N g \ (integrable M f \ integrable N g \ integral\<^sup>L M f = integral\<^sup>L N g) \ integral\<^sup>L M f = integral\<^sup>L N g" by (metis not_integrable_integral_eq) lemma borel_measurable_integrable[measurable_dest]: "integrable M f \ f \ borel_measurable M" by (auto elim: integrable.cases has_bochner_integral.cases) lemma borel_measurable_integrable'[measurable_dest]: "integrable M f \ g \ measurable N M \ (\x. f (g x)) \ borel_measurable N" using borel_measurable_integrable[measurable] by measurable lemma integrable_cong: "M = N \ (\x. x \ space N \ f x = g x) \ integrable M f \ integrable N g" by (simp cong: has_bochner_integral_cong add: integrable.simps) lemma integrable_cong_AE: "f \ borel_measurable M \ g \ borel_measurable M \ AE x in M. f x = g x \ integrable M f \ integrable M g" unfolding integrable.simps by (intro has_bochner_integral_cong_AE arg_cong[where f=Ex] ext) lemma integrable_cong_AE_imp: "integrable M g \ f \ borel_measurable M \ (AE x in M. g x = f x) \ integrable M f" using integrable_cong_AE[of f M g] by (auto simp: eq_commute) lemma integral_cong: "M = N \ (\x. x \ space N \ f x = g x) \ integral\<^sup>L M f = integral\<^sup>L N g" by (simp cong: has_bochner_integral_cong cong del: if_weak_cong add: lebesgue_integral_def) lemma integral_cong_AE: "f \ borel_measurable M \ g \ borel_measurable M \ AE x in M. f x = g x \ integral\<^sup>L M f = integral\<^sup>L M g" unfolding lebesgue_integral_def by (rule arg_cong[where x="has_bochner_integral M f"]) (intro has_bochner_integral_cong_AE ext) lemma integrable_add[simp, intro]: "integrable M f \ integrable M g \ integrable M (\x. f x + g x)" by (auto simp: integrable.simps) lemma integrable_zero[simp, intro]: "integrable M (\x. 0)" by (metis has_bochner_integral_zero integrable.simps) lemma integrable_sum[simp, intro]: "(\i. i \ I \ integrable M (f i)) \ integrable M (\x. \i\I. f i x)" by (metis has_bochner_integral_sum integrable.simps) lemma integrable_indicator[simp, intro]: "A \ sets M \ emeasure M A < \ \ integrable M (\x. indicator A x *\<^sub>R c)" by (metis has_bochner_integral_indicator integrable.simps) lemma integrable_real_indicator[simp, intro]: "A \ sets M \ emeasure M A < \ \ integrable M (indicator A :: 'a \ real)" by (metis has_bochner_integral_real_indicator integrable.simps) lemma integrable_diff[simp, intro]: "integrable M f \ integrable M g \ integrable M (\x. f x - g x)" by (auto simp: integrable.simps intro: has_bochner_integral_diff) lemma integrable_bounded_linear: "bounded_linear T \ integrable M f \ integrable M (\x. T (f x))" by (auto simp: integrable.simps intro: has_bochner_integral_bounded_linear) lemma integrable_scaleR_left[simp, intro]: "(c \ 0 \ integrable M f) \ integrable M (\x. f x *\<^sub>R c)" unfolding integrable.simps by fastforce lemma integrable_scaleR_right[simp, intro]: "(c \ 0 \ integrable M f) \ integrable M (\x. c *\<^sub>R f x)" unfolding integrable.simps by fastforce lemma integrable_mult_left[simp, intro]: fixes c :: "_::{real_normed_algebra,second_countable_topology}" shows "(c \ 0 \ integrable M f) \ integrable M (\x. f x * c)" unfolding integrable.simps by fastforce lemma integrable_mult_right[simp, intro]: fixes c :: "_::{real_normed_algebra,second_countable_topology}" shows "(c \ 0 \ integrable M f) \ integrable M (\x. c * f x)" unfolding integrable.simps by fastforce lemma integrable_divide_zero[simp, intro]: fixes c :: "_::{real_normed_field, field, second_countable_topology}" shows "(c \ 0 \ integrable M f) \ integrable M (\x. f x / c)" unfolding integrable.simps by fastforce lemma integrable_inner_left[simp, intro]: "(c \ 0 \ integrable M f) \ integrable M (\x. f x \ c)" unfolding integrable.simps by fastforce lemma integrable_inner_right[simp, intro]: "(c \ 0 \ integrable M f) \ integrable M (\x. c \ f x)" unfolding integrable.simps by fastforce lemmas integrable_minus[simp, intro] = integrable_bounded_linear[OF bounded_linear_minus[OF bounded_linear_ident]] lemmas integrable_divide[simp, intro] = integrable_bounded_linear[OF bounded_linear_divide] lemmas integrable_Re[simp, intro] = integrable_bounded_linear[OF bounded_linear_Re] lemmas integrable_Im[simp, intro] = integrable_bounded_linear[OF bounded_linear_Im] lemmas integrable_cnj[simp, intro] = integrable_bounded_linear[OF bounded_linear_cnj] lemmas integrable_of_real[simp, intro] = integrable_bounded_linear[OF bounded_linear_of_real] lemmas integrable_fst[simp, intro] = integrable_bounded_linear[OF bounded_linear_fst] lemmas integrable_snd[simp, intro] = integrable_bounded_linear[OF bounded_linear_snd] lemma integral_zero[simp]: "integral\<^sup>L M (\x. 0) = 0" by (intro has_bochner_integral_integral_eq has_bochner_integral_zero) lemma integral_add[simp]: "integrable M f \ integrable M g \ integral\<^sup>L M (\x. f x + g x) = integral\<^sup>L M f + integral\<^sup>L M g" by (intro has_bochner_integral_integral_eq has_bochner_integral_add has_bochner_integral_integrable) lemma integral_diff[simp]: "integrable M f \ integrable M g \ integral\<^sup>L M (\x. f x - g x) = integral\<^sup>L M f - integral\<^sup>L M g" by (intro has_bochner_integral_integral_eq has_bochner_integral_diff has_bochner_integral_integrable) lemma integral_sum: "(\i. i \ I \ integrable M (f i)) \ integral\<^sup>L M (\x. \i\I. f i x) = (\i\I. integral\<^sup>L M (f i))" by (intro has_bochner_integral_integral_eq has_bochner_integral_sum has_bochner_integral_integrable) lemma integral_sum'[simp]: "(\i. i \ I =simp=> integrable M (f i)) \ integral\<^sup>L M (\x. \i\I. f i x) = (\i\I. integral\<^sup>L M (f i))" unfolding simp_implies_def by (rule integral_sum) lemma integral_bounded_linear: "bounded_linear T \ integrable M f \ integral\<^sup>L M (\x. T (f x)) = T (integral\<^sup>L M f)" by (metis has_bochner_integral_bounded_linear has_bochner_integral_integrable has_bochner_integral_integral_eq) lemma integral_bounded_linear': assumes T: "bounded_linear T" and T': "bounded_linear T'" assumes *: "\ (\x. T x = 0) \ (\x. T' (T x) = x)" shows "integral\<^sup>L M (\x. T (f x)) = T (integral\<^sup>L M f)" proof cases assume "(\x. T x = 0)" then show ?thesis by simp next assume **: "\ (\x. T x = 0)" show ?thesis proof cases assume "integrable M f" with T show ?thesis by (rule integral_bounded_linear) next assume not: "\ integrable M f" moreover have "\ integrable M (\x. T (f x))" proof assume "integrable M (\x. T (f x))" from integrable_bounded_linear[OF T' this] not *[OF **] show False by auto qed ultimately show ?thesis using T by (simp add: not_integrable_integral_eq linear_simps) qed qed lemma integral_scaleR_left[simp]: "(c \ 0 \ integrable M f) \ (\ x. f x *\<^sub>R c \M) = integral\<^sup>L M f *\<^sub>R c" by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_scaleR_left) lemma integral_scaleR_right[simp]: "(\ x. c *\<^sub>R f x \M) = c *\<^sub>R integral\<^sup>L M f" by (rule integral_bounded_linear'[OF bounded_linear_scaleR_right bounded_linear_scaleR_right[of "1 / c"]]) simp lemma integral_mult_left[simp]: fixes c :: "_::{real_normed_algebra,second_countable_topology}" shows "(c \ 0 \ integrable M f) \ (\ x. f x * c \M) = integral\<^sup>L M f * c" by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_mult_left) lemma integral_mult_right[simp]: fixes c :: "_::{real_normed_algebra,second_countable_topology}" shows "(c \ 0 \ integrable M f) \ (\ x. c * f x \M) = c * integral\<^sup>L M f" by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_mult_right) lemma integral_mult_left_zero[simp]: fixes c :: "_::{real_normed_field,second_countable_topology}" shows "(\ x. f x * c \M) = integral\<^sup>L M f * c" by (rule integral_bounded_linear'[OF bounded_linear_mult_left bounded_linear_mult_left[of "1 / c"]]) simp lemma integral_mult_right_zero[simp]: fixes c :: "_::{real_normed_field,second_countable_topology}" shows "(\ x. c * f x \M) = c * integral\<^sup>L M f" by (rule integral_bounded_linear'[OF bounded_linear_mult_right bounded_linear_mult_right[of "1 / c"]]) simp lemma integral_inner_left[simp]: "(c \ 0 \ integrable M f) \ (\ x. f x \ c \M) = integral\<^sup>L M f \ c" by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_inner_left) lemma integral_inner_right[simp]: "(c \ 0 \ integrable M f) \ (\ x. c \ f x \M) = c \ integral\<^sup>L M f" by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_inner_right) lemma integral_divide_zero[simp]: fixes c :: "_::{real_normed_field, field, second_countable_topology}" shows "integral\<^sup>L M (\x. f x / c) = integral\<^sup>L M f / c" by (rule integral_bounded_linear'[OF bounded_linear_divide bounded_linear_mult_left[of c]]) simp lemma integral_minus[simp]: "integral\<^sup>L M (\x. - f x) = - integral\<^sup>L M f" by (rule integral_bounded_linear'[OF bounded_linear_minus[OF bounded_linear_ident] bounded_linear_minus[OF bounded_linear_ident]]) simp lemma integral_complex_of_real[simp]: "integral\<^sup>L M (\x. complex_of_real (f x)) = of_real (integral\<^sup>L M f)" by (rule integral_bounded_linear'[OF bounded_linear_of_real bounded_linear_Re]) simp lemma integral_cnj[simp]: "integral\<^sup>L M (\x. cnj (f x)) = cnj (integral\<^sup>L M f)" by (rule integral_bounded_linear'[OF bounded_linear_cnj bounded_linear_cnj]) simp lemmas integral_divide[simp] = integral_bounded_linear[OF bounded_linear_divide] lemmas integral_Re[simp] = integral_bounded_linear[OF bounded_linear_Re] lemmas integral_Im[simp] = integral_bounded_linear[OF bounded_linear_Im] lemmas integral_of_real[simp] = integral_bounded_linear[OF bounded_linear_of_real] lemmas integral_fst[simp] = integral_bounded_linear[OF bounded_linear_fst] lemmas integral_snd[simp] = integral_bounded_linear[OF bounded_linear_snd] lemma integral_norm_bound_ennreal: "integrable M f \ norm (integral\<^sup>L M f) \ (\\<^sup>+x. norm (f x) \M)" by (metis has_bochner_integral_integrable has_bochner_integral_norm_bound) lemma integrableI_sequence: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes f[measurable]: "f \ borel_measurable M" assumes s: "\i. simple_bochner_integrable M (s i)" assumes lim: "(\i. \\<^sup>+x. norm (f x - s i x) \M) \ 0" (is "?S \ 0") shows "integrable M f" proof - let ?s = "\n. simple_bochner_integral M (s n)" have "\x. ?s \ x" unfolding convergent_eq_Cauchy proof (rule metric_CauchyI) fix e :: real assume "0 < e" then have "0 < ennreal (e / 2)" by auto from order_tendstoD(2)[OF lim this] obtain M where M: "\n. M \ n \ ?S n < e / 2" by (auto simp: eventually_sequentially) show "\M. \m\M. \n\M. dist (?s m) (?s n) < e" proof (intro exI allI impI) fix m n assume m: "M \ m" and n: "M \ n" have "?S n \ \" using M[OF n] by auto have "norm (?s n - ?s m) \ ?S n + ?S m" by (intro simple_bochner_integral_bounded s f) also have "\ < ennreal (e / 2) + e / 2" by (intro add_strict_mono M n m) also have "\ = e" using \0 by (simp flip: ennreal_plus) finally show "dist (?s n) (?s m) < e" using \0 by (simp add: dist_norm ennreal_less_iff) qed qed then obtain x where "?s \ x" .. show ?thesis by (rule, rule) fact+ qed proposition nn_integral_dominated_convergence_norm: fixes u' :: "_ \ _::{real_normed_vector, second_countable_topology}" assumes [measurable]: "\i. u i \ borel_measurable M" "u' \ borel_measurable M" "w \ borel_measurable M" and bound: "\j. AE x in M. norm (u j x) \ w x" and w: "(\\<^sup>+x. w x \M) < \" and u': "AE x in M. (\i. u i x) \ u' x" shows "(\i. (\\<^sup>+x. norm (u' x - u i x) \M)) \ 0" proof - have "AE x in M. \j. norm (u j x) \ w x" unfolding AE_all_countable by rule fact with u' have bnd: "AE x in M. \j. norm (u' x - u j x) \ 2 * w x" proof (eventually_elim, intro allI) fix i x assume "(\i. u i x) \ u' x" "\j. norm (u j x) \ w x" "\j. norm (u j x) \ w x" then have "norm (u' x) \ w x" "norm (u i x) \ w x" by (auto intro: LIMSEQ_le_const2 tendsto_norm) then have "norm (u' x) + norm (u i x) \ 2 * w x" by simp also have "norm (u' x - u i x) \ norm (u' x) + norm (u i x)" by (rule norm_triangle_ineq4) finally (xtrans) show "norm (u' x - u i x) \ 2 * w x" . qed have w_nonneg: "AE x in M. 0 \ w x" using bound[of 0] by (auto intro: order_trans[OF norm_ge_zero]) have "(\i. (\\<^sup>+x. norm (u' x - u i x) \M)) \ (\\<^sup>+x. 0 \M)" proof (rule nn_integral_dominated_convergence) show "(\\<^sup>+x. 2 * w x \M) < \" by (rule nn_integral_mult_bounded_inf[OF _ w, of 2]) (insert w_nonneg, auto simp: ennreal_mult ) show "AE x in M. (\i. ennreal (norm (u' x - u i x))) \ 0" using u' proof eventually_elim fix x assume "(\i. u i x) \ u' x" from tendsto_diff[OF tendsto_const[of "u' x"] this] show "(\i. ennreal (norm (u' x - u i x))) \ 0" by (simp add: tendsto_norm_zero_iff flip: ennreal_0) qed qed (insert bnd w_nonneg, auto) then show ?thesis by simp qed proposition integrableI_bounded: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes f[measurable]: "f \ borel_measurable M" and fin: "(\\<^sup>+x. norm (f x) \M) < \" shows "integrable M f" proof - from borel_measurable_implies_sequence_metric[OF f, of 0] obtain s where s: "\i. simple_function M (s i)" and pointwise: "\x. x \ space M \ (\i. s i x) \ f x" and bound: "\i x. x \ space M \ norm (s i x) \ 2 * norm (f x)" by simp metis show ?thesis proof (rule integrableI_sequence) { fix i have "(\\<^sup>+x. norm (s i x) \M) \ (\\<^sup>+x. ennreal (2 * norm (f x)) \M)" by (intro nn_integral_mono) (simp add: bound) also have "\ = 2 * (\\<^sup>+x. ennreal (norm (f x)) \M)" by (simp add: ennreal_mult nn_integral_cmult) also have "\ < top" using fin by (simp add: ennreal_mult_less_top) finally have "(\\<^sup>+x. norm (s i x) \M) < \" by simp } note fin_s = this show "\i. simple_bochner_integrable M (s i)" by (rule simple_bochner_integrableI_bounded) fact+ show "(\i. \\<^sup>+ x. ennreal (norm (f x - s i x)) \M) \ 0" proof (rule nn_integral_dominated_convergence_norm) show "\j. AE x in M. norm (s j x) \ 2 * norm (f x)" using bound by auto show "\i. s i \ borel_measurable M" "(\x. 2 * norm (f x)) \ borel_measurable M" using s by (auto intro: borel_measurable_simple_function) show "(\\<^sup>+ x. ennreal (2 * norm (f x)) \M) < \" using fin by (simp add: nn_integral_cmult ennreal_mult ennreal_mult_less_top) show "AE x in M. (\i. s i x) \ f x" using pointwise by auto qed fact qed fact qed lemma integrableI_bounded_set: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes [measurable]: "A \ sets M" "f \ borel_measurable M" assumes finite: "emeasure M A < \" and bnd: "AE x in M. x \ A \ norm (f x) \ B" and null: "AE x in M. x \ A \ f x = 0" shows "integrable M f" proof (rule integrableI_bounded) { fix x :: 'b have "norm x \ B \ 0 \ B" using norm_ge_zero[of x] by arith } with bnd null have "(\\<^sup>+ x. ennreal (norm (f x)) \M) \ (\\<^sup>+ x. ennreal (max 0 B) * indicator A x \M)" by (intro nn_integral_mono_AE) (auto split: split_indicator split_max) also have "\ < \" using finite by (subst nn_integral_cmult_indicator) (auto simp: ennreal_mult_less_top) finally show "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" . qed simp lemma integrableI_bounded_set_indicator: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "A \ sets M \ f \ borel_measurable M \ emeasure M A < \ \ (AE x in M. x \ A \ norm (f x) \ B) \ integrable M (\x. indicator A x *\<^sub>R f x)" by (rule integrableI_bounded_set[where A=A]) auto lemma integrableI_nonneg: fixes f :: "'a \ real" assumes "f \ borel_measurable M" "AE x in M. 0 \ f x" "(\\<^sup>+x. f x \M) < \" shows "integrable M f" proof - have "(\\<^sup>+x. norm (f x) \M) = (\\<^sup>+x. f x \M)" using assms by (intro nn_integral_cong_AE) auto then show ?thesis using assms by (intro integrableI_bounded) auto qed lemma integrable_iff_bounded: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "integrable M f \ f \ borel_measurable M \ (\\<^sup>+x. norm (f x) \M) < \" using integrableI_bounded[of f M] has_bochner_integral_implies_finite_norm[of M f] unfolding integrable.simps has_bochner_integral.simps[abs_def] by auto lemma integrable_bound: fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ 'c::{banach, second_countable_topology}" shows "integrable M f \ g \ borel_measurable M \ (AE x in M. norm (g x) \ norm (f x)) \ integrable M g" unfolding integrable_iff_bounded proof safe assume "f \ borel_measurable M" "g \ borel_measurable M" assume "AE x in M. norm (g x) \ norm (f x)" then have "(\\<^sup>+ x. ennreal (norm (g x)) \M) \ (\\<^sup>+ x. ennreal (norm (f x)) \M)" by (intro nn_integral_mono_AE) auto also assume "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" finally show "(\\<^sup>+ x. ennreal (norm (g x)) \M) < \" . qed lemma integrable_mult_indicator: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "A \ sets M \ integrable M f \ integrable M (\x. indicator A x *\<^sub>R f x)" by (rule integrable_bound[of M f]) (auto split: split_indicator) lemma integrable_real_mult_indicator: fixes f :: "'a \ real" shows "A \ sets M \ integrable M f \ integrable M (\x. f x * indicator A x)" using integrable_mult_indicator[of A M f] by (simp add: mult_ac) lemma integrable_abs[simp, intro]: fixes f :: "'a \ real" assumes [measurable]: "integrable M f" shows "integrable M (\x. \f x\)" using assms by (rule integrable_bound) auto lemma integrable_norm[simp, intro]: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes [measurable]: "integrable M f" shows "integrable M (\x. norm (f x))" using assms by (rule integrable_bound) auto lemma integrable_norm_cancel: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes [measurable]: "integrable M (\x. norm (f x))" "f \ borel_measurable M" shows "integrable M f" using assms by (rule integrable_bound) auto lemma integrable_norm_iff: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "f \ borel_measurable M \ integrable M (\x. norm (f x)) \ integrable M f" by (auto intro: integrable_norm_cancel) lemma integrable_abs_cancel: fixes f :: "'a \ real" assumes [measurable]: "integrable M (\x. \f x\)" "f \ borel_measurable M" shows "integrable M f" using assms by (rule integrable_bound) auto lemma integrable_abs_iff: fixes f :: "'a \ real" shows "f \ borel_measurable M \ integrable M (\x. \f x\) \ integrable M f" by (auto intro: integrable_abs_cancel) lemma integrable_max[simp, intro]: fixes f :: "'a \ real" assumes fg[measurable]: "integrable M f" "integrable M g" shows "integrable M (\x. max (f x) (g x))" using integrable_add[OF integrable_norm[OF fg(1)] integrable_norm[OF fg(2)]] by (rule integrable_bound) auto lemma integrable_min[simp, intro]: fixes f :: "'a \ real" assumes fg[measurable]: "integrable M f" "integrable M g" shows "integrable M (\x. min (f x) (g x))" using integrable_add[OF integrable_norm[OF fg(1)] integrable_norm[OF fg(2)]] by (rule integrable_bound) auto lemma integral_minus_iff[simp]: "integrable M (\x. - f x ::'a::{banach, second_countable_topology}) \ integrable M f" unfolding integrable_iff_bounded by (auto) lemma integrable_indicator_iff: "integrable M (indicator A::_ \ real) \ A \ space M \ sets M \ emeasure M (A \ space M) < \" by (simp add: integrable_iff_bounded borel_measurable_indicator_iff ennreal_indicator nn_integral_indicator' cong: conj_cong) lemma integral_indicator[simp]: "integral\<^sup>L M (indicator A) = measure M (A \ space M)" proof cases assume *: "A \ space M \ sets M \ emeasure M (A \ space M) < \" have "integral\<^sup>L M (indicator A) = integral\<^sup>L M (indicator (A \ space M))" by (intro integral_cong) (auto split: split_indicator) also have "\ = measure M (A \ space M)" using * by (intro has_bochner_integral_integral_eq has_bochner_integral_real_indicator) auto finally show ?thesis . next assume *: "\ (A \ space M \ sets M \ emeasure M (A \ space M) < \)" have "integral\<^sup>L M (indicator A) = integral\<^sup>L M (indicator (A \ space M) :: _ \ real)" by (intro integral_cong) (auto split: split_indicator) also have "\ = 0" using * by (subst not_integrable_integral_eq) (auto simp: integrable_indicator_iff) also have "\ = measure M (A \ space M)" using * by (auto simp: measure_def emeasure_notin_sets not_less top_unique) finally show ?thesis . qed lemma integrable_discrete_difference: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes X: "countable X" assumes null: "\x. x \ X \ emeasure M {x} = 0" assumes sets: "\x. x \ X \ {x} \ sets M" assumes eq: "\x. x \ space M \ x \ X \ f x = g x" shows "integrable M f \ integrable M g" unfolding integrable_iff_bounded proof (rule conj_cong) { assume "f \ borel_measurable M" then have "g \ borel_measurable M" by (rule measurable_discrete_difference[where X=X]) (auto simp: assms) } moreover { assume "g \ borel_measurable M" then have "f \ borel_measurable M" by (rule measurable_discrete_difference[where X=X]) (auto simp: assms) } ultimately show "f \ borel_measurable M \ g \ borel_measurable M" .. next have "AE x in M. x \ X" by (rule AE_discrete_difference) fact+ then have "(\\<^sup>+ x. norm (f x) \M) = (\\<^sup>+ x. norm (g x) \M)" by (intro nn_integral_cong_AE) (auto simp: eq) then show "(\\<^sup>+ x. norm (f x) \M) < \ \ (\\<^sup>+ x. norm (g x) \M) < \" by simp qed lemma integral_discrete_difference: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes X: "countable X" assumes null: "\x. x \ X \ emeasure M {x} = 0" assumes sets: "\x. x \ X \ {x} \ sets M" assumes eq: "\x. x \ space M \ x \ X \ f x = g x" shows "integral\<^sup>L M f = integral\<^sup>L M g" proof (rule integral_eq_cases) show eq: "integrable M f \ integrable M g" by (rule integrable_discrete_difference[where X=X]) fact+ assume f: "integrable M f" show "integral\<^sup>L M f = integral\<^sup>L M g" proof (rule integral_cong_AE) show "f \ borel_measurable M" "g \ borel_measurable M" using f eq by (auto intro: borel_measurable_integrable) have "AE x in M. x \ X" by (rule AE_discrete_difference) fact+ with AE_space show "AE x in M. f x = g x" by eventually_elim fact qed qed lemma has_bochner_integral_discrete_difference: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes X: "countable X" assumes null: "\x. x \ X \ emeasure M {x} = 0" assumes sets: "\x. x \ X \ {x} \ sets M" assumes eq: "\x. x \ space M \ x \ X \ f x = g x" shows "has_bochner_integral M f x \ has_bochner_integral M g x" using integrable_discrete_difference[of X M f g, OF assms] using integral_discrete_difference[of X M f g, OF assms] by (metis has_bochner_integral_iff) lemma fixes f :: "'a \ 'b::{banach, second_countable_topology}" and w :: "'a \ real" assumes "f \ borel_measurable M" "\i. s i \ borel_measurable M" "integrable M w" assumes lim: "AE x in M. (\i. s i x) \ f x" assumes bound: "\i. AE x in M. norm (s i x) \ w x" shows integrable_dominated_convergence: "integrable M f" and integrable_dominated_convergence2: "\i. integrable M (s i)" and integral_dominated_convergence: "(\i. integral\<^sup>L M (s i)) \ integral\<^sup>L M f" proof - have w_nonneg: "AE x in M. 0 \ w x" using bound[of 0] by eventually_elim (auto intro: norm_ge_zero order_trans) then have "(\\<^sup>+x. w x \M) = (\\<^sup>+x. norm (w x) \M)" by (intro nn_integral_cong_AE) auto with \integrable M w\ have w: "w \ borel_measurable M" "(\\<^sup>+x. w x \M) < \" unfolding integrable_iff_bounded by auto show int_s: "\i. integrable M (s i)" unfolding integrable_iff_bounded proof fix i have "(\\<^sup>+ x. ennreal (norm (s i x)) \M) \ (\\<^sup>+x. w x \M)" using bound[of i] w_nonneg by (intro nn_integral_mono_AE) auto with w show "(\\<^sup>+ x. ennreal (norm (s i x)) \M) < \" by auto qed fact have all_bound: "AE x in M. \i. norm (s i x) \ w x" using bound unfolding AE_all_countable by auto show int_f: "integrable M f" unfolding integrable_iff_bounded proof have "(\\<^sup>+ x. ennreal (norm (f x)) \M) \ (\\<^sup>+x. w x \M)" using all_bound lim w_nonneg proof (intro nn_integral_mono_AE, eventually_elim) fix x assume "\i. norm (s i x) \ w x" "(\i. s i x) \ f x" "0 \ w x" then show "ennreal (norm (f x)) \ ennreal (w x)" by (intro LIMSEQ_le_const2[where X="\i. ennreal (norm (s i x))"]) (auto intro: tendsto_intros) qed with w show "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" by auto qed fact have "(\n. ennreal (norm (integral\<^sup>L M (s n) - integral\<^sup>L M f))) \ ennreal 0" (is "?d \ ennreal 0") proof (rule tendsto_sandwich) show "eventually (\n. ennreal 0 \ ?d n) sequentially" "(\_. ennreal 0) \ ennreal 0" by auto show "eventually (\n. ?d n \ (\\<^sup>+x. norm (s n x - f x) \M)) sequentially" proof (intro always_eventually allI) fix n have "?d n = norm (integral\<^sup>L M (\x. s n x - f x))" using int_f int_s by simp also have "\ \ (\\<^sup>+x. norm (s n x - f x) \M)" by (intro int_f int_s integrable_diff integral_norm_bound_ennreal) finally show "?d n \ (\\<^sup>+x. norm (s n x - f x) \M)" . qed show "(\n. \\<^sup>+x. norm (s n x - f x) \M) \ ennreal 0" unfolding ennreal_0 apply (subst norm_minus_commute) proof (rule nn_integral_dominated_convergence_norm[where w=w]) show "\n. s n \ borel_measurable M" using int_s unfolding integrable_iff_bounded by auto qed fact+ qed then have "(\n. integral\<^sup>L M (s n) - integral\<^sup>L M f) \ 0" by (simp add: tendsto_norm_zero_iff del: ennreal_0) from tendsto_add[OF this tendsto_const[of "integral\<^sup>L M f"]] show "(\i. integral\<^sup>L M (s i)) \ integral\<^sup>L M f" by simp qed context fixes s :: "real \ 'a \ 'b::{banach, second_countable_topology}" and w :: "'a \ real" and f :: "'a \ 'b" and M assumes "f \ borel_measurable M" "\t. s t \ borel_measurable M" "integrable M w" assumes lim: "AE x in M. ((\i. s i x) \ f x) at_top" assumes bound: "\\<^sub>F i in at_top. AE x in M. norm (s i x) \ w x" begin lemma integral_dominated_convergence_at_top: "((\t. integral\<^sup>L M (s t)) \ integral\<^sup>L M f) at_top" proof (rule tendsto_at_topI_sequentially) fix X :: "nat \ real" assume X: "filterlim X at_top sequentially" from filterlim_iff[THEN iffD1, OF this, rule_format, OF bound] obtain N where w: "\n. N \ n \ AE x in M. norm (s (X n) x) \ w x" by (auto simp: eventually_sequentially) show "(\n. integral\<^sup>L M (s (X n))) \ integral\<^sup>L M f" proof (rule LIMSEQ_offset, rule integral_dominated_convergence) show "AE x in M. norm (s (X (n + N)) x) \ w x" for n by (rule w) auto show "AE x in M. (\n. s (X (n + N)) x) \ f x" using lim proof eventually_elim fix x assume "((\i. s i x) \ f x) at_top" then show "(\n. s (X (n + N)) x) \ f x" by (intro LIMSEQ_ignore_initial_segment filterlim_compose[OF _ X]) qed qed fact+ qed lemma integrable_dominated_convergence_at_top: "integrable M f" proof - from bound obtain N where w: "\n. N \ n \ AE x in M. norm (s n x) \ w x" by (auto simp: eventually_at_top_linorder) show ?thesis proof (rule integrable_dominated_convergence) show "AE x in M. norm (s (N + i) x) \ w x" for i :: nat by (intro w) auto show "AE x in M. (\i. s (N + real i) x) \ f x" using lim proof eventually_elim fix x assume "((\i. s i x) \ f x) at_top" then show "(\n. s (N + n) x) \ f x" by (rule filterlim_compose) (auto intro!: filterlim_tendsto_add_at_top filterlim_real_sequentially) qed qed fact+ qed end lemma integrable_mult_left_iff [simp]: fixes f :: "'a \ real" shows "integrable M (\x. c * f x) \ c = 0 \ integrable M f" using integrable_mult_left[of c M f] integrable_mult_left[of "1 / c" M "\x. c * f x"] by (cases "c = 0") auto lemma integrable_mult_right_iff [simp]: fixes f :: "'a \ real" shows "integrable M (\x. f x * c) \ c = 0 \ integrable M f" using integrable_mult_left_iff [of M c f] by (simp add: mult.commute) lemma integrableI_nn_integral_finite: assumes [measurable]: "f \ borel_measurable M" and nonneg: "AE x in M. 0 \ f x" and finite: "(\\<^sup>+x. f x \M) = ennreal x" shows "integrable M f" proof (rule integrableI_bounded) have "(\\<^sup>+ x. ennreal (norm (f x)) \M) = (\\<^sup>+ x. ennreal (f x) \M)" using nonneg by (intro nn_integral_cong_AE) auto with finite show "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" by auto qed simp lemma integral_nonneg_AE: fixes f :: "'a \ real" assumes nonneg: "AE x in M. 0 \ f x" shows "0 \ integral\<^sup>L M f" proof cases assume f: "integrable M f" then have [measurable]: "f \ M \\<^sub>M borel" by auto have "(\x. max 0 (f x)) \ M \\<^sub>M borel" "\x. 0 \ max 0 (f x)" "integrable M (\x. max 0 (f x))" using f by auto from this have "0 \ integral\<^sup>L M (\x. max 0 (f x))" proof (induction rule: borel_measurable_induct_real) case (add f g) then have "integrable M f" "integrable M g" by (auto intro!: integrable_bound[OF add.prems]) with add show ?case by (simp add: nn_integral_add) next case (seq U) show ?case proof (rule LIMSEQ_le_const) have U_le: "x \ space M \ U i x \ max 0 (f x)" for x i using seq by (intro incseq_le) (auto simp: incseq_def le_fun_def) with seq nonneg show "(\i. integral\<^sup>L M (U i)) \ LINT x|M. max 0 (f x)" by (intro integral_dominated_convergence) auto have "integrable M (U i)" for i using seq.prems by (rule integrable_bound) (insert U_le seq, auto) with seq show "\N. \n\N. 0 \ integral\<^sup>L M (U n)" by auto qed qed (auto) also have "\ = integral\<^sup>L M f" using nonneg by (auto intro!: integral_cong_AE) finally show ?thesis . qed (simp add: not_integrable_integral_eq) lemma integral_nonneg[simp]: fixes f :: "'a \ real" shows "(\x. x \ space M \ 0 \ f x) \ 0 \ integral\<^sup>L M f" by (intro integral_nonneg_AE) auto proposition nn_integral_eq_integral: assumes f: "integrable M f" assumes nonneg: "AE x in M. 0 \ f x" shows "(\\<^sup>+ x. f x \M) = integral\<^sup>L M f" proof - { fix f :: "'a \ real" assume f: "f \ borel_measurable M" "\x. 0 \ f x" "integrable M f" then have "(\\<^sup>+ x. f x \M) = integral\<^sup>L M f" proof (induct rule: borel_measurable_induct_real) case (set A) then show ?case by (simp add: integrable_indicator_iff ennreal_indicator emeasure_eq_ennreal_measure) next case (mult f c) then show ?case by (auto simp add: nn_integral_cmult ennreal_mult integral_nonneg_AE) next case (add g f) then have "integrable M f" "integrable M g" by (auto intro!: integrable_bound[OF add.prems]) with add show ?case by (simp add: nn_integral_add integral_nonneg_AE) next case (seq U) show ?case proof (rule LIMSEQ_unique) have U_le_f: "x \ space M \ U i x \ f x" for x i using seq by (intro incseq_le) (auto simp: incseq_def le_fun_def) have int_U: "\i. integrable M (U i)" using seq f U_le_f by (intro integrable_bound[OF f(3)]) auto from U_le_f seq have "(\i. integral\<^sup>L M (U i)) \ integral\<^sup>L M f" by (intro integral_dominated_convergence) auto then show "(\i. ennreal (integral\<^sup>L M (U i))) \ ennreal (integral\<^sup>L M f)" using seq f int_U by (simp add: f integral_nonneg_AE) have "(\i. \\<^sup>+ x. U i x \M) \ \\<^sup>+ x. f x \M" using seq U_le_f f by (intro nn_integral_dominated_convergence[where w=f]) (auto simp: integrable_iff_bounded) then show "(\i. \x. U i x \M) \ \\<^sup>+x. f x \M" using seq int_U by simp qed qed } from this[of "\x. max 0 (f x)"] assms have "(\\<^sup>+ x. max 0 (f x) \M) = integral\<^sup>L M (\x. max 0 (f x))" by simp also have "\ = integral\<^sup>L M f" using assms by (auto intro!: integral_cong_AE simp: integral_nonneg_AE) also have "(\\<^sup>+ x. max 0 (f x) \M) = (\\<^sup>+ x. f x \M)" using assms by (auto intro!: nn_integral_cong_AE simp: max_def) finally show ?thesis . qed lemma nn_integral_eq_integrable: assumes f: "f \ M \\<^sub>M borel" "AE x in M. 0 \ f x" and "0 \ x" shows "(\\<^sup>+x. f x \M) = ennreal x \ (integrable M f \ integral\<^sup>L M f = x)" proof (safe intro!: nn_integral_eq_integral assms) assume *: "(\\<^sup>+x. f x \M) = ennreal x" with integrableI_nn_integral_finite[OF f this] nn_integral_eq_integral[of M f, OF _ f(2)] show "integrable M f" "integral\<^sup>L M f = x" by (simp_all add: * assms integral_nonneg_AE) qed lemma fixes f :: "_ \ _ \ 'a :: {banach, second_countable_topology}" assumes integrable[measurable]: "\i. integrable M (f i)" and summable: "AE x in M. summable (\i. norm (f i x))" and sums: "summable (\i. (\x. norm (f i x) \M))" shows integrable_suminf: "integrable M (\x. (\i. f i x))" (is "integrable M ?S") and sums_integral: "(\i. integral\<^sup>L M (f i)) sums (\x. (\i. f i x) \M)" (is "?f sums ?x") and integral_suminf: "(\x. (\i. f i x) \M) = (\i. integral\<^sup>L M (f i))" and summable_integral: "summable (\i. integral\<^sup>L M (f i))" proof - have 1: "integrable M (\x. \i. norm (f i x))" proof (rule integrableI_bounded) have "(\\<^sup>+ x. ennreal (norm (\i. norm (f i x))) \M) = (\\<^sup>+ x. (\i. ennreal (norm (f i x))) \M)" apply (intro nn_integral_cong_AE) using summable apply eventually_elim apply (simp add: suminf_nonneg ennreal_suminf_neq_top) done also have "\ = (\i. \\<^sup>+ x. norm (f i x) \M)" by (intro nn_integral_suminf) auto also have "\ = (\i. ennreal (\x. norm (f i x) \M))" by (intro arg_cong[where f=suminf] ext nn_integral_eq_integral integrable_norm integrable) auto finally show "(\\<^sup>+ x. ennreal (norm (\i. norm (f i x))) \M) < \" by (simp add: sums ennreal_suminf_neq_top less_top[symmetric] integral_nonneg_AE) qed simp have 2: "AE x in M. (\n. \i (\i. f i x)" using summable by eventually_elim (auto intro: summable_LIMSEQ summable_norm_cancel) have 3: "\j. AE x in M. norm (\i (\i. norm (f i x))" using summable proof eventually_elim fix j x assume [simp]: "summable (\i. norm (f i x))" have "norm (\i (\i \ (\i. norm (f i x))" using sum_le_suminf[of "\i. norm (f i x)"] unfolding sums_iff by auto finally show "norm (\i (\i. norm (f i x))" by simp qed note ibl = integrable_dominated_convergence[OF _ _ 1 2 3] note int = integral_dominated_convergence[OF _ _ 1 2 3] show "integrable M ?S" by (rule ibl) measurable show "?f sums ?x" unfolding sums_def using int by (simp add: integrable) then show "?x = suminf ?f" "summable ?f" unfolding sums_iff by auto qed proposition integral_norm_bound [simp]: fixes f :: "_ \ 'a :: {banach, second_countable_topology}" shows "norm (integral\<^sup>L M f) \ (\x. norm (f x) \M)" proof (cases "integrable M f") case True then show ?thesis using nn_integral_eq_integral[of M "\x. norm (f x)"] integral_norm_bound_ennreal[of M f] by (simp add: integral_nonneg_AE) next case False then have "norm (integral\<^sup>L M f) = 0" by (simp add: not_integrable_integral_eq) moreover have "(\x. norm (f x) \M) \ 0" by auto ultimately show ?thesis by simp qed proposition integral_abs_bound [simp]: fixes f :: "'a \ real" shows "abs (\x. f x \M) \ (\x. \f x\ \M)" using integral_norm_bound[of M f] by auto lemma integral_eq_nn_integral: assumes [measurable]: "f \ borel_measurable M" assumes nonneg: "AE x in M. 0 \ f x" shows "integral\<^sup>L M f = enn2real (\\<^sup>+ x. ennreal (f x) \M)" proof cases assume *: "(\\<^sup>+ x. ennreal (f x) \M) = \" also have "(\\<^sup>+ x. ennreal (f x) \M) = (\\<^sup>+ x. ennreal (norm (f x)) \M)" using nonneg by (intro nn_integral_cong_AE) auto finally have "\ integrable M f" by (auto simp: integrable_iff_bounded) then show ?thesis by (simp add: * not_integrable_integral_eq) next assume "(\\<^sup>+ x. ennreal (f x) \M) \ \" then have "integrable M f" by (cases "\\<^sup>+ x. ennreal (f x) \M" rule: ennreal_cases) (auto intro!: integrableI_nn_integral_finite assms) from nn_integral_eq_integral[OF this] nonneg show ?thesis by (simp add: integral_nonneg_AE) qed lemma enn2real_nn_integral_eq_integral: assumes eq: "AE x in M. f x = ennreal (g x)" and nn: "AE x in M. 0 \ g x" and fin: "(\\<^sup>+x. f x \M) < top" and [measurable]: "g \ M \\<^sub>M borel" shows "enn2real (\\<^sup>+x. f x \M) = (\x. g x \M)" proof - have "ennreal (enn2real (\\<^sup>+x. f x \M)) = (\\<^sup>+x. f x \M)" using fin by (intro ennreal_enn2real) auto also have "\ = (\\<^sup>+x. g x \M)" using eq by (rule nn_integral_cong_AE) also have "\ = (\x. g x \M)" proof (rule nn_integral_eq_integral) show "integrable M g" proof (rule integrableI_bounded) have "(\\<^sup>+ x. ennreal (norm (g x)) \M) = (\\<^sup>+ x. f x \M)" using eq nn by (auto intro!: nn_integral_cong_AE elim!: eventually_elim2) also note fin finally show "(\\<^sup>+ x. ennreal (norm (g x)) \M) < \" by simp qed simp qed fact finally show ?thesis using nn by (simp add: integral_nonneg_AE) qed lemma has_bochner_integral_nn_integral: assumes "f \ borel_measurable M" "AE x in M. 0 \ f x" "0 \ x" assumes "(\\<^sup>+x. f x \M) = ennreal x" shows "has_bochner_integral M f x" unfolding has_bochner_integral_iff using assms by (auto simp: assms integral_eq_nn_integral intro: integrableI_nn_integral_finite) lemma integrableI_simple_bochner_integrable: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "simple_bochner_integrable M f \ integrable M f" by (intro integrableI_sequence[where s="\_. f"] borel_measurable_simple_function) (auto simp: zero_ennreal_def[symmetric] simple_bochner_integrable.simps) proposition integrable_induct[consumes 1, case_names base add lim, induct pred: integrable]: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes "integrable M f" assumes base: "\A c. A \ sets M \ emeasure M A < \ \ P (\x. indicator A x *\<^sub>R c)" assumes add: "\f g. integrable M f \ P f \ integrable M g \ P g \ P (\x. f x + g x)" assumes lim: "\f s. (\i. integrable M (s i)) \ (\i. P (s i)) \ (\x. x \ space M \ (\i. s i x) \ f x) \ (\i x. x \ space M \ norm (s i x) \ 2 * norm (f x)) \ integrable M f \ P f" shows "P f" proof - from \integrable M f\ have f: "f \ borel_measurable M" "(\\<^sup>+x. norm (f x) \M) < \" unfolding integrable_iff_bounded by auto from borel_measurable_implies_sequence_metric[OF f(1)] obtain s where s: "\i. simple_function M (s i)" "\x. x \ space M \ (\i. s i x) \ f x" "\i x. x \ space M \ norm (s i x) \ 2 * norm (f x)" unfolding norm_conv_dist by metis { fix f A have [simp]: "P (\x. 0)" using base[of "{}" undefined] by simp have "(\i::'b. i \ A \ integrable M (f i::'a \ 'b)) \ (\i. i \ A \ P (f i)) \ P (\x. \i\A. f i x)" by (induct A rule: infinite_finite_induct) (auto intro!: add) } note sum = this define s' where [abs_def]: "s' i z = indicator (space M) z *\<^sub>R s i z" for i z then have s'_eq_s: "\i x. x \ space M \ s' i x = s i x" by simp have sf[measurable]: "\i. simple_function M (s' i)" unfolding s'_def using s(1) by (intro simple_function_compose2[where h="(*\<^sub>R)"] simple_function_indicator) auto { fix i have "\z. {y. s' i z = y \ y \ s' i ` space M \ y \ 0 \ z \ space M} = (if z \ space M \ s' i z \ 0 then {s' i z} else {})" by (auto simp add: s'_def split: split_indicator) then have "\z. s' i = (\z. \y\s' i`space M - {0}. indicator {x\space M. s' i x = y} z *\<^sub>R y)" using sf by (auto simp: fun_eq_iff simple_function_def s'_def) } note s'_eq = this show "P f" proof (rule lim) fix i have "(\\<^sup>+x. norm (s' i x) \M) \ (\\<^sup>+x. ennreal (2 * norm (f x)) \M)" using s by (intro nn_integral_mono) (auto simp: s'_eq_s) also have "\ < \" using f by (simp add: nn_integral_cmult ennreal_mult_less_top ennreal_mult) finally have sbi: "simple_bochner_integrable M (s' i)" using sf by (intro simple_bochner_integrableI_bounded) auto then show "integrable M (s' i)" by (rule integrableI_simple_bochner_integrable) { fix x assume"x \ space M" "s' i x \ 0" then have "emeasure M {y \ space M. s' i y = s' i x} \ emeasure M {y \ space M. s' i y \ 0}" by (intro emeasure_mono) auto also have "\ < \" using sbi by (auto elim: simple_bochner_integrable.cases simp: less_top) finally have "emeasure M {y \ space M. s' i y = s' i x} \ \" by simp } then show "P (s' i)" by (subst s'_eq) (auto intro!: sum base simp: less_top) fix x assume "x \ space M" with s show "(\i. s' i x) \ f x" by (simp add: s'_eq_s) show "norm (s' i x) \ 2 * norm (f x)" using \x \ space M\ s by (simp add: s'_eq_s) qed fact qed lemma integral_eq_zero_AE: "(AE x in M. f x = 0) \ integral\<^sup>L M f = 0" using integral_cong_AE[of f M "\_. 0"] by (cases "integrable M f") (simp_all add: not_integrable_integral_eq) lemma integral_nonneg_eq_0_iff_AE: fixes f :: "_ \ real" assumes f[measurable]: "integrable M f" and nonneg: "AE x in M. 0 \ f x" shows "integral\<^sup>L M f = 0 \ (AE x in M. f x = 0)" proof assume "integral\<^sup>L M f = 0" then have "integral\<^sup>N M f = 0" using nn_integral_eq_integral[OF f nonneg] by simp then have "AE x in M. ennreal (f x) \ 0" by (simp add: nn_integral_0_iff_AE) with nonneg show "AE x in M. f x = 0" by auto qed (auto simp add: integral_eq_zero_AE) lemma integral_mono_AE: fixes f :: "'a \ real" assumes "integrable M f" "integrable M g" "AE x in M. f x \ g x" shows "integral\<^sup>L M f \ integral\<^sup>L M g" proof - have "0 \ integral\<^sup>L M (\x. g x - f x)" using assms by (intro integral_nonneg_AE integrable_diff assms) auto also have "\ = integral\<^sup>L M g - integral\<^sup>L M f" by (intro integral_diff assms) finally show ?thesis by simp qed lemma integral_mono: fixes f :: "'a \ real" shows "integrable M f \ integrable M g \ (\x. x \ space M \ f x \ g x) \ integral\<^sup>L M f \ integral\<^sup>L M g" by (intro integral_mono_AE) auto lemma integral_norm_bound_integral: fixes f :: "'a::euclidean_space \ 'b::{banach,second_countable_topology}" assumes "integrable M f" "integrable M g" "\x. x \ space M \ norm(f x) \ g x" shows "norm (\x. f x \M) \ (\x. g x \M)" proof - have "norm (\x. f x \M) \ (\x. norm (f x) \M)" by (rule integral_norm_bound) also have "... \ (\x. g x \M)" using assms integrable_norm integral_mono by blast finally show ?thesis . qed lemma integral_abs_bound_integral: fixes f :: "'a::euclidean_space \ real" assumes "integrable M f" "integrable M g" "\x. x \ space M \ \f x\ \ g x" shows "\\x. f x \M\ \ (\x. g x \M)" by (metis integral_norm_bound_integral assms real_norm_def) text \The next two statements are useful to bound Lebesgue integrals, as they avoid one integrability assumption. The price to pay is that the upper function has to be nonnegative, but this is often true and easy to check in computations.\ lemma integral_mono_AE': fixes f::"_ \ real" assumes "integrable M f" "AE x in M. g x \ f x" "AE x in M. 0 \ f x" shows "(\x. g x \M) \ (\x. f x \M)" proof (cases "integrable M g") case True show ?thesis by (rule integral_mono_AE, auto simp add: assms True) next case False then have "(\x. g x \M) = 0" by (simp add: not_integrable_integral_eq) also have "... \ (\x. f x \M)" by (simp add: integral_nonneg_AE[OF assms(3)]) finally show ?thesis by simp qed lemma integral_mono': fixes f::"_ \ real" assumes "integrable M f" "\x. x \ space M \ g x \ f x" "\x. x \ space M \ 0 \ f x" shows "(\x. g x \M) \ (\x. f x \M)" by (rule integral_mono_AE', insert assms, auto) lemma (in finite_measure) integrable_measure: assumes I: "disjoint_family_on X I" "countable I" shows "integrable (count_space I) (\i. measure M (X i))" proof - have "(\\<^sup>+i. measure M (X i) \count_space I) = (\\<^sup>+i. measure M (if X i \ sets M then X i else {}) \count_space I)" by (auto intro!: nn_integral_cong measure_notin_sets) also have "\ = measure M (\i\I. if X i \ sets M then X i else {})" using I unfolding emeasure_eq_measure[symmetric] by (subst emeasure_UN_countable) (auto simp: disjoint_family_on_def) finally show ?thesis by (auto intro!: integrableI_bounded) qed lemma integrableI_real_bounded: assumes f: "f \ borel_measurable M" and ae: "AE x in M. 0 \ f x" and fin: "integral\<^sup>N M f < \" shows "integrable M f" proof (rule integrableI_bounded) have "(\\<^sup>+ x. ennreal (norm (f x)) \M) = \\<^sup>+ x. ennreal (f x) \M" using ae by (auto intro: nn_integral_cong_AE) also note fin finally show "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" . qed fact lemma nn_integral_nonneg_infinite: fixes f::"'a \ real" assumes "f \ borel_measurable M" "\ integrable M f" "AE x in M. f x \ 0" shows "(\\<^sup>+x. f x \M) = \" using assms integrableI_real_bounded less_top by auto lemma integral_real_bounded: assumes "0 \ r" "integral\<^sup>N M f \ ennreal r" shows "integral\<^sup>L M f \ r" proof cases assume [simp]: "integrable M f" have "integral\<^sup>L M (\x. max 0 (f x)) = integral\<^sup>N M (\x. max 0 (f x))" by (intro nn_integral_eq_integral[symmetric]) auto also have "\ = integral\<^sup>N M f" by (intro nn_integral_cong) (simp add: max_def ennreal_neg) also have "\ \ r" by fact finally have "integral\<^sup>L M (\x. max 0 (f x)) \ r" using \0 \ r\ by simp moreover have "integral\<^sup>L M f \ integral\<^sup>L M (\x. max 0 (f x))" by (rule integral_mono_AE) auto ultimately show ?thesis by simp next assume "\ integrable M f" then show ?thesis using \0 \ r\ by (simp add: not_integrable_integral_eq) qed lemma integrable_MIN: fixes f:: "_ \ _ \ real" shows "\ finite I; I \ {}; \i. i \ I \ integrable M (f i) \ \ integrable M (\x. MIN i\I. f i x)" by (induct rule: finite_ne_induct) simp+ lemma integrable_MAX: fixes f:: "_ \ _ \ real" shows "\ finite I; I \ {}; \i. i \ I \ integrable M (f i) \ \ integrable M (\x. MAX i\I. f i x)" by (induct rule: finite_ne_induct) simp+ theorem integral_Markov_inequality: assumes [measurable]: "integrable M u" and "AE x in M. 0 \ u x" "0 < (c::real)" shows "(emeasure M) {x\space M. u x \ c} \ (1/c) * (\x. u x \M)" proof - have "(\\<^sup>+ x. ennreal(u x) * indicator (space M) x \M) \ (\\<^sup>+ x. u x \M)" by (rule nn_integral_mono_AE, auto simp add: \c>0\ less_eq_real_def) also have "... = (\x. u x \M)" by (rule nn_integral_eq_integral, auto simp add: assms) finally have *: "(\\<^sup>+ x. ennreal(u x) * indicator (space M) x \M) \ (\x. u x \M)" by simp have "{x \ space M. u x \ c} = {x \ space M. ennreal(1/c) * u x \ 1} \ (space M)" using \c>0\ by (auto simp: ennreal_mult'[symmetric]) then have "emeasure M {x \ space M. u x \ c} = emeasure M ({x \ space M. ennreal(1/c) * u x \ 1} \ (space M))" by simp also have "... \ ennreal(1/c) * (\\<^sup>+ x. ennreal(u x) * indicator (space M) x \M)" by (rule nn_integral_Markov_inequality) (auto simp add: assms) also have "... \ ennreal(1/c) * (\x. u x \M)" apply (rule mult_left_mono) using * \c>0\ by auto finally show ?thesis using \0 by (simp add: ennreal_mult'[symmetric]) qed lemma integral_ineq_eq_0_then_AE: fixes f::"_ \ real" assumes "AE x in M. f x \ g x" "integrable M f" "integrable M g" "(\x. f x \M) = (\x. g x \M)" shows "AE x in M. f x = g x" proof - define h where "h = (\x. g x - f x)" have "AE x in M. h x = 0" apply (subst integral_nonneg_eq_0_iff_AE[symmetric]) unfolding h_def using assms by auto then show ?thesis unfolding h_def by auto qed lemma not_AE_zero_int_E: fixes f::"'a \ real" assumes "AE x in M. f x \ 0" "(\x. f x \M) > 0" and [measurable]: "f \ borel_measurable M" shows "\A e. A \ sets M \ e>0 \ emeasure M A > 0 \ (\x \ A. f x \ e)" proof (rule not_AE_zero_E, auto simp add: assms) assume *: "AE x in M. f x = 0" have "(\x. f x \M) = (\x. 0 \M)" by (rule integral_cong_AE, auto simp add: *) then have "(\x. f x \M) = 0" by simp then show False using assms(2) by simp qed proposition tendsto_L1_int: fixes u :: "_ \ _ \ 'b::{banach, second_countable_topology}" assumes [measurable]: "\n. integrable M (u n)" "integrable M f" and "((\n. (\\<^sup>+x. norm(u n x - f x) \M)) \ 0) F" shows "((\n. (\x. u n x \M)) \ (\x. f x \M)) F" proof - have "((\n. norm((\x. u n x \M) - (\x. f x \M))) \ (0::ennreal)) F" proof (rule tendsto_sandwich[of "\_. 0", where ?h = "\n. (\\<^sup>+x. norm(u n x - f x) \M)"], auto simp add: assms) { fix n have "(\x. u n x \M) - (\x. f x \M) = (\x. u n x - f x \M)" apply (rule Bochner_Integration.integral_diff[symmetric]) using assms by auto then have "norm((\x. u n x \M) - (\x. f x \M)) = norm (\x. u n x - f x \M)" by auto also have "... \ (\x. norm(u n x - f x) \M)" by (rule integral_norm_bound) finally have "ennreal(norm((\x. u n x \M) - (\x. f x \M))) \ (\x. norm(u n x - f x) \M)" by simp also have "... = (\\<^sup>+x. norm(u n x - f x) \M)" apply (rule nn_integral_eq_integral[symmetric]) using assms by auto finally have "norm((\x. u n x \M) - (\x. f x \M)) \ (\\<^sup>+x. norm(u n x - f x) \M)" by simp } then show "eventually (\n. norm((\x. u n x \M) - (\x. f x \M)) \ (\\<^sup>+x. norm(u n x - f x) \M)) F" by auto qed then have "((\n. norm((\x. u n x \M) - (\x. f x \M))) \ 0) F" by (simp flip: ennreal_0) then have "((\n. ((\x. u n x \M) - (\x. f x \M))) \ 0) F" using tendsto_norm_zero_iff by blast then show ?thesis using Lim_null by auto qed text \The next lemma asserts that, if a sequence of functions converges in \L\<^sup>1\, then it admits a subsequence that converges almost everywhere.\ proposition tendsto_L1_AE_subseq: fixes u :: "nat \ 'a \ 'b::{banach, second_countable_topology}" assumes [measurable]: "\n. integrable M (u n)" and "(\n. (\x. norm(u n x) \M)) \ 0" shows "\r::nat\nat. strict_mono r \ (AE x in M. (\n. u (r n) x) \ 0)" proof - { fix k have "eventually (\n. (\x. norm(u n x) \M) < (1/2)^k) sequentially" using order_tendstoD(2)[OF assms(2)] by auto with eventually_elim2[OF eventually_gt_at_top[of k] this] have "\n>k. (\x. norm(u n x) \M) < (1/2)^k" by (metis eventually_False_sequentially) } then have "\r. \n. True \ (r (Suc n) > r n \ (\x. norm(u (r (Suc n)) x) \M) < (1/2)^(r n))" by (intro dependent_nat_choice, auto) then obtain r0 where r0: "strict_mono r0" "\n. (\x. norm(u (r0 (Suc n)) x) \M) < (1/2)^(r0 n)" by (auto simp: strict_mono_Suc_iff) define r where "r = (\n. r0(n+1))" have "strict_mono r" unfolding r_def using r0(1) by (simp add: strict_mono_Suc_iff) have I: "(\\<^sup>+x. norm(u (r n) x) \M) < ennreal((1/2)^n)" for n proof - have "r0 n \ n" using \strict_mono r0\ by (simp add: seq_suble) have "(1/2::real)^(r0 n) \ (1/2)^n" by (rule power_decreasing[OF \r0 n \ n\], auto) then have "(\x. norm(u (r0 (Suc n)) x) \M) < (1/2)^n" using r0(2) less_le_trans by blast then have "(\x. norm(u (r n) x) \M) < (1/2)^n" unfolding r_def by auto moreover have "(\\<^sup>+x. norm(u (r n) x) \M) = (\x. norm(u (r n) x) \M)" by (rule nn_integral_eq_integral, auto simp add: integrable_norm[OF assms(1)[of "r n"]]) ultimately show ?thesis by (auto intro: ennreal_lessI) qed have "AE x in M. limsup (\n. ennreal (norm(u (r n) x))) \ 0" proof (rule AE_upper_bound_inf_ennreal) fix e::real assume "e > 0" define A where "A = (\n. {x \ space M. norm(u (r n) x) \ e})" have A_meas [measurable]: "\n. A n \ sets M" unfolding A_def by auto have A_bound: "emeasure M (A n) < (1/e) * ennreal((1/2)^n)" for n proof - have *: "indicator (A n) x \ (1/e) * ennreal(norm(u (r n) x))" for x apply (cases "x \ A n") unfolding A_def using \0 < e\ by (auto simp: ennreal_mult[symmetric]) have "emeasure M (A n) = (\\<^sup>+x. indicator (A n) x \M)" by auto also have "... \ (\\<^sup>+x. (1/e) * ennreal(norm(u (r n) x)) \M)" apply (rule nn_integral_mono) using * by auto also have "... = (1/e) * (\\<^sup>+x. norm(u (r n) x) \M)" apply (rule nn_integral_cmult) using \e > 0\ by auto also have "... < (1/e) * ennreal((1/2)^n)" using I[of n] \e > 0\ by (intro ennreal_mult_strict_left_mono) auto finally show ?thesis by simp qed have A_fin: "emeasure M (A n) < \" for n using \e > 0\ A_bound[of n] by (auto simp add: ennreal_mult less_top[symmetric]) have A_sum: "summable (\n. measure M (A n))" proof (rule summable_comparison_test'[of "\n. (1/e) * (1/2)^n" 0]) have "summable (\n. (1/(2::real))^n)" by (simp add: summable_geometric) then show "summable (\n. (1/e) * (1/2)^n)" using summable_mult by blast fix n::nat assume "n \ 0" have "norm(measure M (A n)) = measure M (A n)" by simp also have "... = enn2real(emeasure M (A n))" unfolding measure_def by simp also have "... < enn2real((1/e) * (1/2)^n)" using A_bound[of n] \emeasure M (A n) < \\ \0 < e\ by (auto simp: emeasure_eq_ennreal_measure ennreal_mult[symmetric] ennreal_less_iff) also have "... = (1/e) * (1/2)^n" using \0 by auto finally show "norm(measure M (A n)) \ (1/e) * (1/2)^n" by simp qed have "AE x in M. eventually (\n. x \ space M - A n) sequentially" by (rule borel_cantelli_AE1[OF A_meas A_fin A_sum]) moreover { fix x assume "eventually (\n. x \ space M - A n) sequentially" moreover have "norm(u (r n) x) \ ennreal e" if "x \ space M - A n" for n using that unfolding A_def by (auto intro: ennreal_leI) ultimately have "eventually (\n. norm(u (r n) x) \ ennreal e) sequentially" by (simp add: eventually_mono) then have "limsup (\n. ennreal (norm(u (r n) x))) \ e" by (simp add: Limsup_bounded) } ultimately show "AE x in M. limsup (\n. ennreal (norm(u (r n) x))) \ 0 + ennreal e" by auto qed moreover { fix x assume "limsup (\n. ennreal (norm(u (r n) x))) \ 0" moreover then have "liminf (\n. ennreal (norm(u (r n) x))) \ 0" by (rule order_trans[rotated]) (auto intro: Liminf_le_Limsup) ultimately have "(\n. ennreal (norm(u (r n) x))) \ 0" using tendsto_Limsup[of sequentially "\n. ennreal (norm(u (r n) x))"] by auto then have "(\n. norm(u (r n) x)) \ 0" by (simp flip: ennreal_0) then have "(\n. u (r n) x) \ 0" by (simp add: tendsto_norm_zero_iff) } ultimately have "AE x in M. (\n. u (r n) x) \ 0" by auto then show ?thesis using \strict_mono r\ by auto qed subsection \Restricted measure spaces\ lemma integrable_restrict_space: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes \[simp]: "\ \ space M \ sets M" shows "integrable (restrict_space M \) f \ integrable M (\x. indicator \ x *\<^sub>R f x)" unfolding integrable_iff_bounded borel_measurable_restrict_space_iff[OF \] nn_integral_restrict_space[OF \] by (simp add: ac_simps ennreal_indicator ennreal_mult) lemma integral_restrict_space: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes \[simp]: "\ \ space M \ sets M" shows "integral\<^sup>L (restrict_space M \) f = integral\<^sup>L M (\x. indicator \ x *\<^sub>R f x)" proof (rule integral_eq_cases) assume "integrable (restrict_space M \) f" then show ?thesis proof induct case (base A c) then show ?case by (simp add: indicator_inter_arith[symmetric] sets_restrict_space_iff emeasure_restrict_space Int_absorb1 measure_restrict_space) next case (add g f) then show ?case by (simp add: scaleR_add_right integrable_restrict_space) next case (lim f s) show ?case proof (rule LIMSEQ_unique) show "(\i. integral\<^sup>L (restrict_space M \) (s i)) \ integral\<^sup>L (restrict_space M \) f" using lim by (intro integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) simp_all show "(\i. integral\<^sup>L (restrict_space M \) (s i)) \ (\ x. indicator \ x *\<^sub>R f x \M)" unfolding lim using lim by (intro integral_dominated_convergence[where w="\x. 2 * norm (indicator \ x *\<^sub>R f x)"]) (auto simp add: space_restrict_space integrable_restrict_space simp del: norm_scaleR split: split_indicator) qed qed qed (simp add: integrable_restrict_space) lemma integral_empty: assumes "space M = {}" shows "integral\<^sup>L M f = 0" proof - have "(\ x. f x \M) = (\ x. 0 \M)" by(rule integral_cong)(simp_all add: assms) thus ?thesis by simp qed subsection \Measure spaces with an associated density\ lemma integrable_density: fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ real" assumes [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" and nn: "AE x in M. 0 \ g x" shows "integrable (density M g) f \ integrable M (\x. g x *\<^sub>R f x)" unfolding integrable_iff_bounded using nn apply (simp add: nn_integral_density less_top[symmetric]) apply (intro arg_cong2[where f="(=)"] refl nn_integral_cong_AE) apply (auto simp: ennreal_mult) done lemma integral_density: fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ real" assumes f: "f \ borel_measurable M" and g[measurable]: "g \ borel_measurable M" "AE x in M. 0 \ g x" shows "integral\<^sup>L (density M g) f = integral\<^sup>L M (\x. g x *\<^sub>R f x)" proof (rule integral_eq_cases) assume "integrable (density M g) f" then show ?thesis proof induct case (base A c) then have [measurable]: "A \ sets M" by auto have int: "integrable M (\x. g x * indicator A x)" using g base integrable_density[of "indicator A :: 'a \ real" M g] by simp then have "integral\<^sup>L M (\x. g x * indicator A x) = (\\<^sup>+ x. ennreal (g x * indicator A x) \M)" using g by (subst nn_integral_eq_integral) auto also have "\ = (\\<^sup>+ x. ennreal (g x) * indicator A x \M)" by (intro nn_integral_cong) (auto split: split_indicator) also have "\ = emeasure (density M g) A" by (rule emeasure_density[symmetric]) auto also have "\ = ennreal (measure (density M g) A)" using base by (auto intro: emeasure_eq_ennreal_measure) also have "\ = integral\<^sup>L (density M g) (indicator A)" using base by simp finally show ?case using base g apply (simp add: int integral_nonneg_AE) apply (subst (asm) ennreal_inj) apply (auto intro!: integral_nonneg_AE) done next case (add f h) then have [measurable]: "f \ borel_measurable M" "h \ borel_measurable M" by (auto dest!: borel_measurable_integrable) from add g show ?case by (simp add: scaleR_add_right integrable_density) next case (lim f s) have [measurable]: "f \ borel_measurable M" "\i. s i \ borel_measurable M" using lim(1,5)[THEN borel_measurable_integrable] by auto show ?case proof (rule LIMSEQ_unique) show "(\i. integral\<^sup>L M (\x. g x *\<^sub>R s i x)) \ integral\<^sup>L M (\x. g x *\<^sub>R f x)" proof (rule integral_dominated_convergence) show "integrable M (\x. 2 * norm (g x *\<^sub>R f x))" by (intro integrable_mult_right integrable_norm integrable_density[THEN iffD1] lim g) auto show "AE x in M. (\i. g x *\<^sub>R s i x) \ g x *\<^sub>R f x" using lim(3) by (auto intro!: tendsto_scaleR AE_I2[of M]) show "\i. AE x in M. norm (g x *\<^sub>R s i x) \ 2 * norm (g x *\<^sub>R f x)" using lim(4) g by (auto intro!: AE_I2[of M] mult_left_mono simp: field_simps) qed auto show "(\i. integral\<^sup>L M (\x. g x *\<^sub>R s i x)) \ integral\<^sup>L (density M g) f" unfolding lim(2)[symmetric] by (rule integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) (insert lim(3-5), auto) qed qed qed (simp add: f g integrable_density) lemma fixes g :: "'a \ real" assumes "f \ borel_measurable M" "AE x in M. 0 \ f x" "g \ borel_measurable M" shows integral_real_density: "integral\<^sup>L (density M f) g = (\ x. f x * g x \M)" and integrable_real_density: "integrable (density M f) g \ integrable M (\x. f x * g x)" using assms integral_density[of g M f] integrable_density[of g M f] by auto lemma has_bochner_integral_density: fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ real" shows "f \ borel_measurable M \ g \ borel_measurable M \ (AE x in M. 0 \ g x) \ has_bochner_integral M (\x. g x *\<^sub>R f x) x \ has_bochner_integral (density M g) f x" by (simp add: has_bochner_integral_iff integrable_density integral_density) subsection \Distributions\ lemma integrable_distr_eq: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes [measurable]: "g \ measurable M N" "f \ borel_measurable N" shows "integrable (distr M N g) f \ integrable M (\x. f (g x))" unfolding integrable_iff_bounded by (simp_all add: nn_integral_distr) lemma integrable_distr: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "T \ measurable M M' \ integrable (distr M M' T) f \ integrable M (\x. f (T x))" by (subst integrable_distr_eq[symmetric, where g=T]) (auto dest: borel_measurable_integrable) lemma integral_distr: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes g[measurable]: "g \ measurable M N" and f: "f \ borel_measurable N" shows "integral\<^sup>L (distr M N g) f = integral\<^sup>L M (\x. f (g x))" proof (rule integral_eq_cases) assume "integrable (distr M N g) f" then show ?thesis proof induct case (base A c) then have [measurable]: "A \ sets N" by auto from base have int: "integrable (distr M N g) (\a. indicator A a *\<^sub>R c)" by (intro integrable_indicator) have "integral\<^sup>L (distr M N g) (\a. indicator A a *\<^sub>R c) = measure (distr M N g) A *\<^sub>R c" using base by auto also have "\ = measure M (g -` A \ space M) *\<^sub>R c" by (subst measure_distr) auto also have "\ = integral\<^sup>L M (\a. indicator (g -` A \ space M) a *\<^sub>R c)" using base by (auto simp: emeasure_distr) also have "\ = integral\<^sup>L M (\a. indicator A (g a) *\<^sub>R c)" using int base by (intro integral_cong_AE) (auto simp: emeasure_distr split: split_indicator) finally show ?case . next case (add f h) then have [measurable]: "f \ borel_measurable N" "h \ borel_measurable N" by (auto dest!: borel_measurable_integrable) from add g show ?case by (simp add: scaleR_add_right integrable_distr_eq) next case (lim f s) have [measurable]: "f \ borel_measurable N" "\i. s i \ borel_measurable N" using lim(1,5)[THEN borel_measurable_integrable] by auto show ?case proof (rule LIMSEQ_unique) show "(\i. integral\<^sup>L M (\x. s i (g x))) \ integral\<^sup>L M (\x. f (g x))" proof (rule integral_dominated_convergence) show "integrable M (\x. 2 * norm (f (g x)))" using lim by (auto simp: integrable_distr_eq) show "AE x in M. (\i. s i (g x)) \ f (g x)" using lim(3) g[THEN measurable_space] by auto show "\i. AE x in M. norm (s i (g x)) \ 2 * norm (f (g x))" using lim(4) g[THEN measurable_space] by auto qed auto show "(\i. integral\<^sup>L M (\x. s i (g x))) \ integral\<^sup>L (distr M N g) f" unfolding lim(2)[symmetric] by (rule integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) (insert lim(3-5), auto) qed qed qed (simp add: f g integrable_distr_eq) lemma has_bochner_integral_distr: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "f \ borel_measurable N \ g \ measurable M N \ has_bochner_integral M (\x. f (g x)) x \ has_bochner_integral (distr M N g) f x" by (simp add: has_bochner_integral_iff integrable_distr_eq integral_distr) subsection \Lebesgue integration on \<^const>\count_space\\ lemma integrable_count_space: fixes f :: "'a \ 'b::{banach,second_countable_topology}" shows "finite X \ integrable (count_space X) f" by (auto simp: nn_integral_count_space integrable_iff_bounded) lemma measure_count_space[simp]: "B \ A \ finite B \ measure (count_space A) B = card B" unfolding measure_def by (subst emeasure_count_space ) auto lemma lebesgue_integral_count_space_finite_support: assumes f: "finite {a\A. f a \ 0}" shows "(\x. f x \count_space A) = (\a | a \ A \ f a \ 0. f a)" proof - have eq: "\x. x \ A \ (\a | x = a \ a \ A \ f a \ 0. f a) = (\x\{x}. f x)" by (intro sum.mono_neutral_cong_left) auto have "(\x. f x \count_space A) = (\x. (\a | a \ A \ f a \ 0. indicator {a} x *\<^sub>R f a) \count_space A)" by (intro integral_cong refl) (simp add: f eq) also have "\ = (\a | a \ A \ f a \ 0. measure (count_space A) {a} *\<^sub>R f a)" by (subst integral_sum) (auto intro!: sum.cong) finally show ?thesis by auto qed lemma lebesgue_integral_count_space_finite: "finite A \ (\x. f x \count_space A) = (\a\A. f a)" by (subst lebesgue_integral_count_space_finite_support) (auto intro!: sum.mono_neutral_cong_left) lemma integrable_count_space_nat_iff: fixes f :: "nat \ _::{banach,second_countable_topology}" shows "integrable (count_space UNIV) f \ summable (\x. norm (f x))" by (auto simp add: integrable_iff_bounded nn_integral_count_space_nat ennreal_suminf_neq_top intro: summable_suminf_not_top) lemma sums_integral_count_space_nat: fixes f :: "nat \ _::{banach,second_countable_topology}" assumes *: "integrable (count_space UNIV) f" shows "f sums (integral\<^sup>L (count_space UNIV) f)" proof - let ?f = "\n i. indicator {n} i *\<^sub>R f i" have f': "\n i. ?f n i = indicator {n} i *\<^sub>R f n" by (auto simp: fun_eq_iff split: split_indicator) have "(\i. \n. ?f i n \count_space UNIV) sums \ n. (\i. ?f i n) \count_space UNIV" proof (rule sums_integral) show "\i. integrable (count_space UNIV) (?f i)" using * by (intro integrable_mult_indicator) auto show "AE n in count_space UNIV. summable (\i. norm (?f i n))" using summable_finite[of "{n}" "\i. norm (?f i n)" for n] by simp show "summable (\i. \ n. norm (?f i n) \count_space UNIV)" using * by (subst f') (simp add: integrable_count_space_nat_iff) qed also have "(\ n. (\i. ?f i n) \count_space UNIV) = (\n. f n \count_space UNIV)" using suminf_finite[of "{n}" "\i. ?f i n" for n] by (auto intro!: integral_cong) also have "(\i. \n. ?f i n \count_space UNIV) = f" by (subst f') simp finally show ?thesis . qed lemma integral_count_space_nat: fixes f :: "nat \ _::{banach,second_countable_topology}" shows "integrable (count_space UNIV) f \ integral\<^sup>L (count_space UNIV) f = (\x. f x)" using sums_integral_count_space_nat by (rule sums_unique) lemma integrable_bij_count_space: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes g: "bij_betw g A B" shows "integrable (count_space A) (\x. f (g x)) \ integrable (count_space B) f" unfolding integrable_iff_bounded by (subst nn_integral_bij_count_space[OF g]) auto lemma integral_bij_count_space: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes g: "bij_betw g A B" shows "integral\<^sup>L (count_space A) (\x. f (g x)) = integral\<^sup>L (count_space B) f" using g[THEN bij_betw_imp_funcset] apply (subst distr_bij_count_space[OF g, symmetric]) apply (intro integral_distr[symmetric]) apply auto done lemma has_bochner_integral_count_space_nat: fixes f :: "nat \ _::{banach,second_countable_topology}" shows "has_bochner_integral (count_space UNIV) f x \ f sums x" unfolding has_bochner_integral_iff by (auto intro!: sums_integral_count_space_nat) subsection \Point measure\ lemma lebesgue_integral_point_measure_finite: fixes g :: "'a \ 'b::{banach, second_countable_topology}" shows "finite A \ (\a. a \ A \ 0 \ f a) \ integral\<^sup>L (point_measure A f) g = (\a\A. f a *\<^sub>R g a)" by (simp add: lebesgue_integral_count_space_finite AE_count_space integral_density point_measure_def) proposition integrable_point_measure_finite: fixes g :: "'a \ 'b::{banach, second_countable_topology}" and f :: "'a \ real" shows "finite A \ integrable (point_measure A f) g" unfolding point_measure_def apply (subst density_cong[where f'="\x. ennreal (max 0 (f x))"]) apply (auto split: split_max simp: ennreal_neg) apply (subst integrable_density) apply (auto simp: AE_count_space integrable_count_space) done subsection \Lebesgue integration on \<^const>\null_measure\\ lemma has_bochner_integral_null_measure_iff[iff]: "has_bochner_integral (null_measure M) f 0 \ f \ borel_measurable M" by (auto simp add: has_bochner_integral.simps simple_bochner_integral_def[abs_def] intro!: exI[of _ "\n x. 0"] simple_bochner_integrable.intros) lemma integrable_null_measure_iff[iff]: "integrable (null_measure M) f \ f \ borel_measurable M" by (auto simp add: integrable.simps) lemma integral_null_measure[simp]: "integral\<^sup>L (null_measure M) f = 0" by (cases "integrable (null_measure M) f") (auto simp add: not_integrable_integral_eq has_bochner_integral_integral_eq) subsection \Legacy lemmas for the real-valued Lebesgue integral\ theorem real_lebesgue_integral_def: assumes f[measurable]: "integrable M f" shows "integral\<^sup>L M f = enn2real (\\<^sup>+x. f x \M) - enn2real (\\<^sup>+x. ennreal (- f x) \M)" proof - have "integral\<^sup>L M f = integral\<^sup>L M (\x. max 0 (f x) - max 0 (- f x))" by (auto intro!: arg_cong[where f="integral\<^sup>L M"]) also have "\ = integral\<^sup>L M (\x. max 0 (f x)) - integral\<^sup>L M (\x. max 0 (- f x))" by (intro integral_diff integrable_max integrable_minus integrable_zero f) also have "integral\<^sup>L M (\x. max 0 (f x)) = enn2real (\\<^sup>+x. ennreal (f x) \M)" by (subst integral_eq_nn_integral) (auto intro!: arg_cong[where f=enn2real] nn_integral_cong simp: max_def ennreal_neg) also have "integral\<^sup>L M (\x. max 0 (- f x)) = enn2real (\\<^sup>+x. ennreal (- f x) \M)" by (subst integral_eq_nn_integral) (auto intro!: arg_cong[where f=enn2real] nn_integral_cong simp: max_def ennreal_neg) finally show ?thesis . qed theorem real_integrable_def: "integrable M f \ f \ borel_measurable M \ (\\<^sup>+ x. ennreal (f x) \M) \ \ \ (\\<^sup>+ x. ennreal (- f x) \M) \ \" unfolding integrable_iff_bounded proof (safe del: notI) assume *: "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" have "(\\<^sup>+ x. ennreal (f x) \M) \ (\\<^sup>+ x. ennreal (norm (f x)) \M)" by (intro nn_integral_mono) auto also note * finally show "(\\<^sup>+ x. ennreal (f x) \M) \ \" by simp have "(\\<^sup>+ x. ennreal (- f x) \M) \ (\\<^sup>+ x. ennreal (norm (f x)) \M)" by (intro nn_integral_mono) auto also note * finally show "(\\<^sup>+ x. ennreal (- f x) \M) \ \" by simp next assume [measurable]: "f \ borel_measurable M" assume fin: "(\\<^sup>+ x. ennreal (f x) \M) \ \" "(\\<^sup>+ x. ennreal (- f x) \M) \ \" have "(\\<^sup>+ x. norm (f x) \M) = (\\<^sup>+ x. ennreal (f x) + ennreal (- f x) \M)" by (intro nn_integral_cong) (auto simp: abs_real_def ennreal_neg) also have"\ = (\\<^sup>+ x. ennreal (f x) \M) + (\\<^sup>+ x. ennreal (- f x) \M)" by (intro nn_integral_add) auto also have "\ < \" using fin by (auto simp: less_top) finally show "(\\<^sup>+ x. norm (f x) \M) < \" . qed lemma integrableD[dest]: assumes "integrable M f" shows "f \ borel_measurable M" "(\\<^sup>+ x. ennreal (f x) \M) \ \" "(\\<^sup>+ x. ennreal (- f x) \M) \ \" using assms unfolding real_integrable_def by auto lemma integrableE: assumes "integrable M f" obtains r q where "0 \ r" "0 \ q" "(\\<^sup>+x. ennreal (f x)\M) = ennreal r" "(\\<^sup>+x. ennreal (-f x)\M) = ennreal q" "f \ borel_measurable M" "integral\<^sup>L M f = r - q" using assms unfolding real_integrable_def real_lebesgue_integral_def[OF assms] by (cases rule: ennreal2_cases[of "(\\<^sup>+x. ennreal (-f x)\M)" "(\\<^sup>+x. ennreal (f x)\M)"]) auto lemma integral_monotone_convergence_nonneg: fixes f :: "nat \ 'a \ real" assumes i: "\i. integrable M (f i)" and mono: "AE x in M. mono (\n. f n x)" and pos: "\i. AE x in M. 0 \ f i x" and lim: "AE x in M. (\i. f i x) \ u x" and ilim: "(\i. integral\<^sup>L M (f i)) \ x" and u: "u \ borel_measurable M" shows "integrable M u" and "integral\<^sup>L M u = x" proof - have nn: "AE x in M. \i. 0 \ f i x" using pos unfolding AE_all_countable by auto with lim have u_nn: "AE x in M. 0 \ u x" by eventually_elim (auto intro: LIMSEQ_le_const) have [simp]: "0 \ x" by (intro LIMSEQ_le_const[OF ilim] allI exI impI integral_nonneg_AE pos) have "(\\<^sup>+ x. ennreal (u x) \M) = (SUP n. (\\<^sup>+ x. ennreal (f n x) \M))" proof (subst nn_integral_monotone_convergence_SUP_AE[symmetric]) fix i from mono nn show "AE x in M. ennreal (f i x) \ ennreal (f (Suc i) x)" by eventually_elim (auto simp: mono_def) show "(\x. ennreal (f i x)) \ borel_measurable M" using i by auto next show "(\\<^sup>+ x. ennreal (u x) \M) = \\<^sup>+ x. (SUP i. ennreal (f i x)) \M" apply (rule nn_integral_cong_AE) using lim mono nn u_nn apply eventually_elim apply (simp add: LIMSEQ_unique[OF _ LIMSEQ_SUP] incseq_def) done qed also have "\ = ennreal x" using mono i nn unfolding nn_integral_eq_integral[OF i pos] by (subst LIMSEQ_unique[OF LIMSEQ_SUP]) (auto simp: mono_def integral_nonneg_AE pos intro!: integral_mono_AE ilim) finally have "(\\<^sup>+ x. ennreal (u x) \M) = ennreal x" . moreover have "(\\<^sup>+ x. ennreal (- u x) \M) = 0" using u u_nn by (subst nn_integral_0_iff_AE) (auto simp add: ennreal_neg) ultimately show "integrable M u" "integral\<^sup>L M u = x" by (auto simp: real_integrable_def real_lebesgue_integral_def u) qed lemma fixes f :: "nat \ 'a \ real" assumes f: "\i. integrable M (f i)" and mono: "AE x in M. mono (\n. f n x)" and lim: "AE x in M. (\i. f i x) \ u x" and ilim: "(\i. integral\<^sup>L M (f i)) \ x" and u: "u \ borel_measurable M" shows integrable_monotone_convergence: "integrable M u" and integral_monotone_convergence: "integral\<^sup>L M u = x" and has_bochner_integral_monotone_convergence: "has_bochner_integral M u x" proof - have 1: "\i. integrable M (\x. f i x - f 0 x)" using f by auto have 2: "AE x in M. mono (\n. f n x - f 0 x)" using mono by (auto simp: mono_def le_fun_def) have 3: "\n. AE x in M. 0 \ f n x - f 0 x" using mono by (auto simp: field_simps mono_def le_fun_def) have 4: "AE x in M. (\i. f i x - f 0 x) \ u x - f 0 x" using lim by (auto intro!: tendsto_diff) have 5: "(\i. (\x. f i x - f 0 x \M)) \ x - integral\<^sup>L M (f 0)" using f ilim by (auto intro!: tendsto_diff) have 6: "(\x. u x - f 0 x) \ borel_measurable M" using f[of 0] u by auto note diff = integral_monotone_convergence_nonneg[OF 1 2 3 4 5 6] have "integrable M (\x. (u x - f 0 x) + f 0 x)" using diff(1) f by (rule integrable_add) with diff(2) f show "integrable M u" "integral\<^sup>L M u = x" by auto then show "has_bochner_integral M u x" by (metis has_bochner_integral_integrable) qed lemma integral_norm_eq_0_iff: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes f[measurable]: "integrable M f" shows "(\x. norm (f x) \M) = 0 \ emeasure M {x\space M. f x \ 0} = 0" proof - have "(\\<^sup>+x. norm (f x) \M) = (\x. norm (f x) \M)" using f by (intro nn_integral_eq_integral integrable_norm) auto then have "(\x. norm (f x) \M) = 0 \ (\\<^sup>+x. norm (f x) \M) = 0" by simp also have "\ \ emeasure M {x\space M. ennreal (norm (f x)) \ 0} = 0" by (intro nn_integral_0_iff) auto finally show ?thesis by simp qed lemma integral_0_iff: fixes f :: "'a \ real" shows "integrable M f \ (\x. \f x\ \M) = 0 \ emeasure M {x\space M. f x \ 0} = 0" using integral_norm_eq_0_iff[of M f] by simp lemma (in finite_measure) integrable_const[intro!, simp]: "integrable M (\x. a)" using integrable_indicator[of "space M" M a] by (simp cong: integrable_cong add: less_top[symmetric]) lemma lebesgue_integral_const[simp]: fixes a :: "'a :: {banach, second_countable_topology}" shows "(\x. a \M) = measure M (space M) *\<^sub>R a" proof - { assume "emeasure M (space M) = \" "a \ 0" then have ?thesis by (auto simp add: not_integrable_integral_eq ennreal_mult_less_top measure_def integrable_iff_bounded) } moreover { assume "a = 0" then have ?thesis by simp } moreover { assume "emeasure M (space M) \ \" interpret finite_measure M proof qed fact have "(\x. a \M) = (\x. indicator (space M) x *\<^sub>R a \M)" by (intro integral_cong) auto also have "\ = measure M (space M) *\<^sub>R a" by (simp add: less_top[symmetric]) finally have ?thesis . } ultimately show ?thesis by blast qed lemma (in finite_measure) integrable_const_bound: fixes f :: "'a \ 'b::{banach,second_countable_topology}" shows "AE x in M. norm (f x) \ B \ f \ borel_measurable M \ integrable M f" apply (rule integrable_bound[OF integrable_const[of B], of f]) apply assumption apply (cases "0 \ B") apply auto done lemma (in finite_measure) integral_bounded_eq_bound_then_AE: assumes "AE x in M. f x \ (c::real)" "integrable M f" "(\x. f x \M) = c * measure M (space M)" shows "AE x in M. f x = c" apply (rule integral_ineq_eq_0_then_AE) using assms by auto lemma integral_indicator_finite_real: fixes f :: "'a \ real" assumes [simp]: "finite A" assumes [measurable]: "\a. a \ A \ {a} \ sets M" assumes finite: "\a. a \ A \ emeasure M {a} < \" shows "(\x. f x * indicator A x \M) = (\a\A. f a * measure M {a})" proof - have "(\x. f x * indicator A x \M) = (\x. (\a\A. f a * indicator {a} x) \M)" proof (intro integral_cong refl) fix x show "f x * indicator A x = (\a\A. f a * indicator {a} x)" by (auto split: split_indicator simp: eq_commute[of x] cong: conj_cong) qed also have "\ = (\a\A. f a * measure M {a})" using finite by (subst integral_sum) (auto) finally show ?thesis . qed lemma (in finite_measure) ennreal_integral_real: assumes [measurable]: "f \ borel_measurable M" assumes ae: "AE x in M. f x \ ennreal B" "0 \ B" shows "ennreal (\x. enn2real (f x) \M) = (\\<^sup>+x. f x \M)" proof (subst nn_integral_eq_integral[symmetric]) show "integrable M (\x. enn2real (f x))" using ae by (intro integrable_const_bound[where B=B]) (auto simp: enn2real_leI) show "(\\<^sup>+ x. ennreal (enn2real (f x)) \M) = integral\<^sup>N M f" using ae by (intro nn_integral_cong_AE) (auto simp: le_less_trans[OF _ ennreal_less_top]) qed auto lemma (in finite_measure) integral_less_AE: fixes X Y :: "'a \ real" assumes int: "integrable M X" "integrable M Y" assumes A: "(emeasure M) A \ 0" "A \ sets M" "AE x in M. x \ A \ X x \ Y x" assumes gt: "AE x in M. X x \ Y x" shows "integral\<^sup>L M X < integral\<^sup>L M Y" proof - have "integral\<^sup>L M X \ integral\<^sup>L M Y" using gt int by (intro integral_mono_AE) auto moreover have "integral\<^sup>L M X \ integral\<^sup>L M Y" proof assume eq: "integral\<^sup>L M X = integral\<^sup>L M Y" have "integral\<^sup>L M (\x. \Y x - X x\) = integral\<^sup>L M (\x. Y x - X x)" using gt int by (intro integral_cong_AE) auto also have "\ = 0" using eq int by simp finally have "(emeasure M) {x \ space M. Y x - X x \ 0} = 0" using int by (simp add: integral_0_iff) moreover have "(\\<^sup>+x. indicator A x \M) \ (\\<^sup>+x. indicator {x \ space M. Y x - X x \ 0} x \M)" using A by (intro nn_integral_mono_AE) auto then have "(emeasure M) A \ (emeasure M) {x \ space M. Y x - X x \ 0}" using int A by (simp add: integrable_def) ultimately have "emeasure M A = 0" by simp with \(emeasure M) A \ 0\ show False by auto qed ultimately show ?thesis by auto qed lemma (in finite_measure) integral_less_AE_space: fixes X Y :: "'a \ real" assumes int: "integrable M X" "integrable M Y" assumes gt: "AE x in M. X x < Y x" "emeasure M (space M) \ 0" shows "integral\<^sup>L M X < integral\<^sup>L M Y" using gt by (intro integral_less_AE[OF int, where A="space M"]) auto lemma tendsto_integral_at_top: fixes f :: "real \ 'a::{banach, second_countable_topology}" assumes [measurable_cong]: "sets M = sets borel" and f[measurable]: "integrable M f" shows "((\y. \ x. indicator {.. y} x *\<^sub>R f x \M) \ \ x. f x \M) at_top" proof (rule tendsto_at_topI_sequentially) fix X :: "nat \ real" assume "filterlim X at_top sequentially" show "(\n. \x. indicator {..X n} x *\<^sub>R f x \M) \ integral\<^sup>L M f" proof (rule integral_dominated_convergence) show "integrable M (\x. norm (f x))" by (rule integrable_norm) fact show "AE x in M. (\n. indicator {..X n} x *\<^sub>R f x) \ f x" proof fix x from \filterlim X at_top sequentially\ have "eventually (\n. x \ X n) sequentially" unfolding filterlim_at_top_ge[where c=x] by auto then show "(\n. indicator {..X n} x *\<^sub>R f x) \ f x" by (intro tendsto_eventually) (auto split: split_indicator elim!: eventually_mono) qed fix n show "AE x in M. norm (indicator {..X n} x *\<^sub>R f x) \ norm (f x)" by (auto split: split_indicator) qed auto qed lemma fixes f :: "real \ real" assumes M: "sets M = sets borel" assumes nonneg: "AE x in M. 0 \ f x" assumes borel: "f \ borel_measurable borel" assumes int: "\y. integrable M (\x. f x * indicator {.. y} x)" assumes conv: "((\y. \ x. f x * indicator {.. y} x \M) \ x) at_top" shows has_bochner_integral_monotone_convergence_at_top: "has_bochner_integral M f x" and integrable_monotone_convergence_at_top: "integrable M f" and integral_monotone_convergence_at_top:"integral\<^sup>L M f = x" proof - from nonneg have "AE x in M. mono (\n::nat. f x * indicator {..real n} x)" by (auto split: split_indicator intro!: monoI) { fix x have "eventually (\n. f x * indicator {..real n} x = f x) sequentially" by (rule eventually_sequentiallyI[of "nat \x\"]) (auto split: split_indicator simp: nat_le_iff ceiling_le_iff) } from filterlim_cong[OF refl refl this] have "AE x in M. (\i. f x * indicator {..real i} x) \ f x" by simp have "(\i. \ x. f x * indicator {..real i} x \M) \ x" using conv filterlim_real_sequentially by (rule filterlim_compose) have M_measure[simp]: "borel_measurable M = borel_measurable borel" using M by (simp add: sets_eq_imp_space_eq measurable_def) have "f \ borel_measurable M" using borel by simp show "has_bochner_integral M f x" by (rule has_bochner_integral_monotone_convergence) fact+ then show "integrable M f" "integral\<^sup>L M f = x" by (auto simp: _has_bochner_integral_iff) qed subsection \Product measure\ lemma (in sigma_finite_measure) borel_measurable_lebesgue_integrable[measurable (raw)]: fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" assumes [measurable]: "case_prod f \ borel_measurable (N \\<^sub>M M)" shows "Measurable.pred N (\x. integrable M (f x))" proof - have [simp]: "\x. x \ space N \ integrable M (f x) \ (\\<^sup>+y. norm (f x y) \M) < \" unfolding integrable_iff_bounded by simp show ?thesis by (simp cong: measurable_cong) qed lemma (in sigma_finite_measure) measurable_measure[measurable (raw)]: "(\x. x \ space N \ A x \ space M) \ {x \ space (N \\<^sub>M M). snd x \ A (fst x)} \ sets (N \\<^sub>M M) \ (\x. measure M (A x)) \ borel_measurable N" unfolding measure_def by (intro measurable_emeasure borel_measurable_enn2real) auto proposition (in sigma_finite_measure) borel_measurable_lebesgue_integral[measurable (raw)]: fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" assumes f[measurable]: "case_prod f \ borel_measurable (N \\<^sub>M M)" shows "(\x. \y. f x y \M) \ borel_measurable N" proof - from borel_measurable_implies_sequence_metric[OF f, of 0] guess s .. then have s: "\i. simple_function (N \\<^sub>M M) (s i)" "\x y. x \ space N \ y \ space M \ (\i. s i (x, y)) \ f x y" "\i x y. x \ space N \ y \ space M \ norm (s i (x, y)) \ 2 * norm (f x y)" by (auto simp: space_pair_measure) have [measurable]: "\i. s i \ borel_measurable (N \\<^sub>M M)" by (rule borel_measurable_simple_function) fact have "\i. s i \ measurable (N \\<^sub>M M) (count_space UNIV)" by (rule measurable_simple_function) fact define f' where [abs_def]: "f' i x = (if integrable M (f x) then simple_bochner_integral M (\y. s i (x, y)) else 0)" for i x { fix i x assume "x \ space N" then have "simple_bochner_integral M (\y. s i (x, y)) = (\z\s i ` (space N \ space M). measure M {y \ space M. s i (x, y) = z} *\<^sub>R z)" using s(1)[THEN simple_functionD(1)] unfolding simple_bochner_integral_def by (intro sum.mono_neutral_cong_left) (auto simp: eq_commute space_pair_measure image_iff cong: conj_cong) } note eq = this show ?thesis proof (rule borel_measurable_LIMSEQ_metric) fix i show "f' i \ borel_measurable N" unfolding f'_def by (simp_all add: eq cong: measurable_cong if_cong) next fix x assume x: "x \ space N" { assume int_f: "integrable M (f x)" have int_2f: "integrable M (\y. 2 * norm (f x y))" by (intro integrable_norm integrable_mult_right int_f) have "(\i. integral\<^sup>L M (\y. s i (x, y))) \ integral\<^sup>L M (f x)" proof (rule integral_dominated_convergence) from int_f show "f x \ borel_measurable M" by auto show "\i. (\y. s i (x, y)) \ borel_measurable M" using x by simp show "AE xa in M. (\i. s i (x, xa)) \ f x xa" using x s(2) by auto show "\i. AE xa in M. norm (s i (x, xa)) \ 2 * norm (f x xa)" using x s(3) by auto qed fact moreover { fix i have "simple_bochner_integrable M (\y. s i (x, y))" proof (rule simple_bochner_integrableI_bounded) have "(\y. s i (x, y)) ` space M \ s i ` (space N \ space M)" using x by auto then show "simple_function M (\y. s i (x, y))" using simple_functionD(1)[OF s(1), of i] x by (intro simple_function_borel_measurable) (auto simp: space_pair_measure dest: finite_subset) have "(\\<^sup>+ y. ennreal (norm (s i (x, y))) \M) \ (\\<^sup>+ y. 2 * norm (f x y) \M)" using x s by (intro nn_integral_mono) auto also have "(\\<^sup>+ y. 2 * norm (f x y) \M) < \" using int_2f unfolding integrable_iff_bounded by simp finally show "(\\<^sup>+ xa. ennreal (norm (s i (x, xa))) \M) < \" . qed then have "integral\<^sup>L M (\y. s i (x, y)) = simple_bochner_integral M (\y. s i (x, y))" by (rule simple_bochner_integrable_eq_integral[symmetric]) } ultimately have "(\i. simple_bochner_integral M (\y. s i (x, y))) \ integral\<^sup>L M (f x)" by simp } then show "(\i. f' i x) \ integral\<^sup>L M (f x)" unfolding f'_def by (cases "integrable M (f x)") (simp_all add: not_integrable_integral_eq) qed qed lemma (in pair_sigma_finite) integrable_product_swap: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes "integrable (M1 \\<^sub>M M2) f" shows "integrable (M2 \\<^sub>M M1) (\(x,y). f (y,x))" proof - interpret Q: pair_sigma_finite M2 M1 .. have *: "(\(x,y). f (y,x)) = (\x. f (case x of (x,y)\(y,x)))" by (auto simp: fun_eq_iff) show ?thesis unfolding * by (rule integrable_distr[OF measurable_pair_swap']) (simp add: distr_pair_swap[symmetric] assms) qed lemma (in pair_sigma_finite) integrable_product_swap_iff: fixes f :: "_ \ _::{banach, second_countable_topology}" shows "integrable (M2 \\<^sub>M M1) (\(x,y). f (y,x)) \ integrable (M1 \\<^sub>M M2) f" proof - interpret Q: pair_sigma_finite M2 M1 .. from Q.integrable_product_swap[of "\(x,y). f (y,x)"] integrable_product_swap[of f] show ?thesis by auto qed lemma (in pair_sigma_finite) integral_product_swap: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f: "f \ borel_measurable (M1 \\<^sub>M M2)" shows "(\(x,y). f (y,x) \(M2 \\<^sub>M M1)) = integral\<^sup>L (M1 \\<^sub>M M2) f" proof - have *: "(\(x,y). f (y,x)) = (\x. f (case x of (x,y)\(y,x)))" by (auto simp: fun_eq_iff) show ?thesis unfolding * by (simp add: integral_distr[symmetric, OF measurable_pair_swap' f] distr_pair_swap[symmetric]) qed theorem (in pair_sigma_finite) Fubini_integrable: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f[measurable]: "f \ borel_measurable (M1 \\<^sub>M M2)" and integ1: "integrable M1 (\x. \ y. norm (f (x, y)) \M2)" and integ2: "AE x in M1. integrable M2 (\y. f (x, y))" shows "integrable (M1 \\<^sub>M M2) f" proof (rule integrableI_bounded) have "(\\<^sup>+ p. norm (f p) \(M1 \\<^sub>M M2)) = (\\<^sup>+ x. (\\<^sup>+ y. norm (f (x, y)) \M2) \M1)" by (simp add: M2.nn_integral_fst [symmetric]) also have "\ = (\\<^sup>+ x. \\y. norm (f (x, y)) \M2\ \M1)" apply (intro nn_integral_cong_AE) using integ2 proof eventually_elim fix x assume "integrable M2 (\y. f (x, y))" then have f: "integrable M2 (\y. norm (f (x, y)))" by simp then have "(\\<^sup>+y. ennreal (norm (f (x, y))) \M2) = ennreal (LINT y|M2. norm (f (x, y)))" by (rule nn_integral_eq_integral) simp also have "\ = ennreal \LINT y|M2. norm (f (x, y))\" using f by simp finally show "(\\<^sup>+y. ennreal (norm (f (x, y))) \M2) = ennreal \LINT y|M2. norm (f (x, y))\" . qed also have "\ < \" using integ1 by (simp add: integrable_iff_bounded integral_nonneg_AE) finally show "(\\<^sup>+ p. norm (f p) \(M1 \\<^sub>M M2)) < \" . qed fact lemma (in pair_sigma_finite) emeasure_pair_measure_finite: assumes A: "A \ sets (M1 \\<^sub>M M2)" and finite: "emeasure (M1 \\<^sub>M M2) A < \" shows "AE x in M1. emeasure M2 {y\space M2. (x, y) \ A} < \" proof - from M2.emeasure_pair_measure_alt[OF A] finite have "(\\<^sup>+ x. emeasure M2 (Pair x -` A) \M1) \ \" by simp then have "AE x in M1. emeasure M2 (Pair x -` A) \ \" by (rule nn_integral_PInf_AE[rotated]) (intro M2.measurable_emeasure_Pair A) moreover have "\x. x \ space M1 \ Pair x -` A = {y\space M2. (x, y) \ A}" using sets.sets_into_space[OF A] by (auto simp: space_pair_measure) ultimately show ?thesis by (auto simp: less_top) qed lemma (in pair_sigma_finite) AE_integrable_fst': fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f[measurable]: "integrable (M1 \\<^sub>M M2) f" shows "AE x in M1. integrable M2 (\y. f (x, y))" proof - have "(\\<^sup>+x. (\\<^sup>+y. norm (f (x, y)) \M2) \M1) = (\\<^sup>+x. norm (f x) \(M1 \\<^sub>M M2))" by (rule M2.nn_integral_fst) simp also have "(\\<^sup>+x. norm (f x) \(M1 \\<^sub>M M2)) \ \" using f unfolding integrable_iff_bounded by simp finally have "AE x in M1. (\\<^sup>+y. norm (f (x, y)) \M2) \ \" by (intro nn_integral_PInf_AE M2.borel_measurable_nn_integral ) (auto simp: measurable_split_conv) with AE_space show ?thesis by eventually_elim (auto simp: integrable_iff_bounded measurable_compose[OF _ borel_measurable_integrable[OF f]] less_top) qed lemma (in pair_sigma_finite) integrable_fst': fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f[measurable]: "integrable (M1 \\<^sub>M M2) f" shows "integrable M1 (\x. \y. f (x, y) \M2)" unfolding integrable_iff_bounded proof show "(\x. \ y. f (x, y) \M2) \ borel_measurable M1" by (rule M2.borel_measurable_lebesgue_integral) simp have "(\\<^sup>+ x. ennreal (norm (\ y. f (x, y) \M2)) \M1) \ (\\<^sup>+x. (\\<^sup>+y. norm (f (x, y)) \M2) \M1)" using AE_integrable_fst'[OF f] by (auto intro!: nn_integral_mono_AE integral_norm_bound_ennreal) also have "(\\<^sup>+x. (\\<^sup>+y. norm (f (x, y)) \M2) \M1) = (\\<^sup>+x. norm (f x) \(M1 \\<^sub>M M2))" by (rule M2.nn_integral_fst) simp also have "(\\<^sup>+x. norm (f x) \(M1 \\<^sub>M M2)) < \" using f unfolding integrable_iff_bounded by simp finally show "(\\<^sup>+ x. ennreal (norm (\ y. f (x, y) \M2)) \M1) < \" . qed proposition (in pair_sigma_finite) integral_fst': fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f: "integrable (M1 \\<^sub>M M2) f" shows "(\x. (\y. f (x, y) \M2) \M1) = integral\<^sup>L (M1 \\<^sub>M M2) f" using f proof induct case (base A c) have A[measurable]: "A \ sets (M1 \\<^sub>M M2)" by fact have eq: "\x y. x \ space M1 \ indicator A (x, y) = indicator {y\space M2. (x, y) \ A} y" using sets.sets_into_space[OF A] by (auto split: split_indicator simp: space_pair_measure) have int_A: "integrable (M1 \\<^sub>M M2) (indicator A :: _ \ real)" using base by (rule integrable_real_indicator) have "(\ x. \ y. indicator A (x, y) *\<^sub>R c \M2 \M1) = (\x. measure M2 {y\space M2. (x, y) \ A} *\<^sub>R c \M1)" proof (intro integral_cong_AE, simp, simp) from AE_integrable_fst'[OF int_A] AE_space show "AE x in M1. (\y. indicator A (x, y) *\<^sub>R c \M2) = measure M2 {y\space M2. (x, y) \ A} *\<^sub>R c" by eventually_elim (simp add: eq integrable_indicator_iff) qed also have "\ = measure (M1 \\<^sub>M M2) A *\<^sub>R c" proof (subst integral_scaleR_left) have "(\\<^sup>+x. ennreal (measure M2 {y \ space M2. (x, y) \ A}) \M1) = (\\<^sup>+x. emeasure M2 {y \ space M2. (x, y) \ A} \M1)" using emeasure_pair_measure_finite[OF base] by (intro nn_integral_cong_AE, eventually_elim) (simp add: emeasure_eq_ennreal_measure) also have "\ = emeasure (M1 \\<^sub>M M2) A" using sets.sets_into_space[OF A] by (subst M2.emeasure_pair_measure_alt) (auto intro!: nn_integral_cong arg_cong[where f="emeasure M2"] simp: space_pair_measure) finally have *: "(\\<^sup>+x. ennreal (measure M2 {y \ space M2. (x, y) \ A}) \M1) = emeasure (M1 \\<^sub>M M2) A" . from base * show "integrable M1 (\x. measure M2 {y \ space M2. (x, y) \ A})" by (simp add: integrable_iff_bounded) then have "(\x. measure M2 {y \ space M2. (x, y) \ A} \M1) = (\\<^sup>+x. ennreal (measure M2 {y \ space M2. (x, y) \ A}) \M1)" by (rule nn_integral_eq_integral[symmetric]) simp also note * finally show "(\x. measure M2 {y \ space M2. (x, y) \ A} \M1) *\<^sub>R c = measure (M1 \\<^sub>M M2) A *\<^sub>R c" using base by (simp add: emeasure_eq_ennreal_measure) qed also have "\ = (\ a. indicator A a *\<^sub>R c \(M1 \\<^sub>M M2))" using base by simp finally show ?case . next case (add f g) then have [measurable]: "f \ borel_measurable (M1 \\<^sub>M M2)" "g \ borel_measurable (M1 \\<^sub>M M2)" by auto have "(\ x. \ y. f (x, y) + g (x, y) \M2 \M1) = (\ x. (\ y. f (x, y) \M2) + (\ y. g (x, y) \M2) \M1)" apply (rule integral_cong_AE) apply simp_all using AE_integrable_fst'[OF add(1)] AE_integrable_fst'[OF add(3)] apply eventually_elim apply simp done also have "\ = (\ x. f x \(M1 \\<^sub>M M2)) + (\ x. g x \(M1 \\<^sub>M M2))" using integrable_fst'[OF add(1)] integrable_fst'[OF add(3)] add(2,4) by simp finally show ?case using add by simp next case (lim f s) then have [measurable]: "f \ borel_measurable (M1 \\<^sub>M M2)" "\i. s i \ borel_measurable (M1 \\<^sub>M M2)" by auto show ?case proof (rule LIMSEQ_unique) show "(\i. integral\<^sup>L (M1 \\<^sub>M M2) (s i)) \ integral\<^sup>L (M1 \\<^sub>M M2) f" proof (rule integral_dominated_convergence) show "integrable (M1 \\<^sub>M M2) (\x. 2 * norm (f x))" using lim(5) by auto qed (insert lim, auto) have "(\i. \ x. \ y. s i (x, y) \M2 \M1) \ \ x. \ y. f (x, y) \M2 \M1" proof (rule integral_dominated_convergence) have "AE x in M1. \i. integrable M2 (\y. s i (x, y))" unfolding AE_all_countable using AE_integrable_fst'[OF lim(1)] .. with AE_space AE_integrable_fst'[OF lim(5)] show "AE x in M1. (\i. \ y. s i (x, y) \M2) \ \ y. f (x, y) \M2" proof eventually_elim fix x assume x: "x \ space M1" and s: "\i. integrable M2 (\y. s i (x, y))" and f: "integrable M2 (\y. f (x, y))" show "(\i. \ y. s i (x, y) \M2) \ \ y. f (x, y) \M2" proof (rule integral_dominated_convergence) show "integrable M2 (\y. 2 * norm (f (x, y)))" using f by auto show "AE xa in M2. (\i. s i (x, xa)) \ f (x, xa)" using x lim(3) by (auto simp: space_pair_measure) show "\i. AE xa in M2. norm (s i (x, xa)) \ 2 * norm (f (x, xa))" using x lim(4) by (auto simp: space_pair_measure) qed (insert x, measurable) qed show "integrable M1 (\x. (\ y. 2 * norm (f (x, y)) \M2))" by (intro integrable_mult_right integrable_norm integrable_fst' lim) fix i show "AE x in M1. norm (\ y. s i (x, y) \M2) \ (\ y. 2 * norm (f (x, y)) \M2)" using AE_space AE_integrable_fst'[OF lim(1), of i] AE_integrable_fst'[OF lim(5)] proof eventually_elim fix x assume x: "x \ space M1" and s: "integrable M2 (\y. s i (x, y))" and f: "integrable M2 (\y. f (x, y))" from s have "norm (\ y. s i (x, y) \M2) \ (\\<^sup>+y. norm (s i (x, y)) \M2)" by (rule integral_norm_bound_ennreal) also have "\ \ (\\<^sup>+y. 2 * norm (f (x, y)) \M2)" using x lim by (auto intro!: nn_integral_mono simp: space_pair_measure) also have "\ = (\y. 2 * norm (f (x, y)) \M2)" using f by (intro nn_integral_eq_integral) auto finally show "norm (\ y. s i (x, y) \M2) \ (\ y. 2 * norm (f (x, y)) \M2)" by simp qed qed simp_all then show "(\i. integral\<^sup>L (M1 \\<^sub>M M2) (s i)) \ \ x. \ y. f (x, y) \M2 \M1" using lim by simp qed qed lemma (in pair_sigma_finite) fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" assumes f: "integrable (M1 \\<^sub>M M2) (case_prod f)" shows AE_integrable_fst: "AE x in M1. integrable M2 (\y. f x y)" (is "?AE") and integrable_fst: "integrable M1 (\x. \y. f x y \M2)" (is "?INT") and integral_fst: "(\x. (\y. f x y \M2) \M1) = integral\<^sup>L (M1 \\<^sub>M M2) (\(x, y). f x y)" (is "?EQ") using AE_integrable_fst'[OF f] integrable_fst'[OF f] integral_fst'[OF f] by auto lemma (in pair_sigma_finite) fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" assumes f[measurable]: "integrable (M1 \\<^sub>M M2) (case_prod f)" shows AE_integrable_snd: "AE y in M2. integrable M1 (\x. f x y)" (is "?AE") and integrable_snd: "integrable M2 (\y. \x. f x y \M1)" (is "?INT") and integral_snd: "(\y. (\x. f x y \M1) \M2) = integral\<^sup>L (M1 \\<^sub>M M2) (case_prod f)" (is "?EQ") proof - interpret Q: pair_sigma_finite M2 M1 .. have Q_int: "integrable (M2 \\<^sub>M M1) (\(x, y). f y x)" using f unfolding integrable_product_swap_iff[symmetric] by simp show ?AE using Q.AE_integrable_fst'[OF Q_int] by simp show ?INT using Q.integrable_fst'[OF Q_int] by simp show ?EQ using Q.integral_fst'[OF Q_int] using integral_product_swap[of "case_prod f"] by simp qed proposition (in pair_sigma_finite) Fubini_integral: fixes f :: "_ \ _ \ _ :: {banach, second_countable_topology}" assumes f: "integrable (M1 \\<^sub>M M2) (case_prod f)" shows "(\y. (\x. f x y \M1) \M2) = (\x. (\y. f x y \M2) \M1)" unfolding integral_snd[OF assms] integral_fst[OF assms] .. lemma (in product_sigma_finite) product_integral_singleton: fixes f :: "_ \ _::{banach, second_countable_topology}" shows "f \ borel_measurable (M i) \ (\x. f (x i) \Pi\<^sub>M {i} M) = integral\<^sup>L (M i) f" apply (subst distr_singleton[symmetric]) apply (subst integral_distr) apply simp_all done lemma (in product_sigma_finite) product_integral_fold: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes IJ[simp]: "I \ J = {}" and fin: "finite I" "finite J" and f: "integrable (Pi\<^sub>M (I \ J) M) f" shows "integral\<^sup>L (Pi\<^sub>M (I \ J) M) f = (\x. (\y. f (merge I J (x, y)) \Pi\<^sub>M J M) \Pi\<^sub>M I M)" proof - interpret I: finite_product_sigma_finite M I by standard fact interpret J: finite_product_sigma_finite M J by standard fact have "finite (I \ J)" using fin by auto interpret IJ: finite_product_sigma_finite M "I \ J" by standard fact interpret P: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" .. let ?M = "merge I J" let ?f = "\x. f (?M x)" from f have f_borel: "f \ borel_measurable (Pi\<^sub>M (I \ J) M)" by auto have P_borel: "(\x. f (merge I J x)) \ borel_measurable (Pi\<^sub>M I M \\<^sub>M Pi\<^sub>M J M)" using measurable_comp[OF measurable_merge f_borel] by (simp add: comp_def) have f_int: "integrable (Pi\<^sub>M I M \\<^sub>M Pi\<^sub>M J M) ?f" by (rule integrable_distr[OF measurable_merge]) (simp add: distr_merge[OF IJ fin] f) show ?thesis apply (subst distr_merge[symmetric, OF IJ fin]) apply (subst integral_distr[OF measurable_merge f_borel]) apply (subst P.integral_fst'[symmetric, OF f_int]) apply simp done qed lemma (in product_sigma_finite) product_integral_insert: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes I: "finite I" "i \ I" and f: "integrable (Pi\<^sub>M (insert i I) M) f" shows "integral\<^sup>L (Pi\<^sub>M (insert i I) M) f = (\x. (\y. f (x(i:=y)) \M i) \Pi\<^sub>M I M)" proof - have "integral\<^sup>L (Pi\<^sub>M (insert i I) M) f = integral\<^sup>L (Pi\<^sub>M (I \ {i}) M) f" by simp also have "\ = (\x. (\y. f (merge I {i} (x,y)) \Pi\<^sub>M {i} M) \Pi\<^sub>M I M)" using f I by (intro product_integral_fold) auto also have "\ = (\x. (\y. f (x(i := y)) \M i) \Pi\<^sub>M I M)" proof (rule integral_cong[OF refl], subst product_integral_singleton[symmetric]) fix x assume x: "x \ space (Pi\<^sub>M I M)" have f_borel: "f \ borel_measurable (Pi\<^sub>M (insert i I) M)" using f by auto show "(\y. f (x(i := y))) \ borel_measurable (M i)" using measurable_comp[OF measurable_component_update f_borel, OF x \i \ I\] unfolding comp_def . from x I show "(\ y. f (merge I {i} (x,y)) \Pi\<^sub>M {i} M) = (\ xa. f (x(i := xa i)) \Pi\<^sub>M {i} M)" by (auto intro!: integral_cong arg_cong[where f=f] simp: merge_def space_PiM extensional_def PiE_def) qed finally show ?thesis . qed lemma (in product_sigma_finite) product_integrable_prod: fixes f :: "'i \ 'a \ _::{real_normed_field,banach,second_countable_topology}" assumes [simp]: "finite I" and integrable: "\i. i \ I \ integrable (M i) (f i)" shows "integrable (Pi\<^sub>M I M) (\x. (\i\I. f i (x i)))" (is "integrable _ ?f") proof (unfold integrable_iff_bounded, intro conjI) interpret finite_product_sigma_finite M I by standard fact show "?f \ borel_measurable (Pi\<^sub>M I M)" using assms by simp have "(\\<^sup>+ x. ennreal (norm (\i\I. f i (x i))) \Pi\<^sub>M I M) = (\\<^sup>+ x. (\i\I. ennreal (norm (f i (x i)))) \Pi\<^sub>M I M)" by (simp add: prod_norm prod_ennreal) also have "\ = (\i\I. \\<^sup>+ x. ennreal (norm (f i x)) \M i)" using assms by (intro product_nn_integral_prod) auto also have "\ < \" using integrable by (simp add: less_top[symmetric] ennreal_prod_eq_top integrable_iff_bounded) finally show "(\\<^sup>+ x. ennreal (norm (\i\I. f i (x i))) \Pi\<^sub>M I M) < \" . qed lemma (in product_sigma_finite) product_integral_prod: fixes f :: "'i \ 'a \ _::{real_normed_field,banach,second_countable_topology}" assumes "finite I" and integrable: "\i. i \ I \ integrable (M i) (f i)" shows "(\x. (\i\I. f i (x i)) \Pi\<^sub>M I M) = (\i\I. integral\<^sup>L (M i) (f i))" using assms proof induct case empty interpret finite_measure "Pi\<^sub>M {} M" by rule (simp add: space_PiM) show ?case by (simp add: space_PiM measure_def) next case (insert i I) then have iI: "finite (insert i I)" by auto then have prod: "\J. J \ insert i I \ integrable (Pi\<^sub>M J M) (\x. (\i\J. f i (x i)))" by (intro product_integrable_prod insert(4)) (auto intro: finite_subset) interpret I: finite_product_sigma_finite M I by standard fact have *: "\x y. (\j\I. f j (if j = i then y else x j)) = (\j\I. f j (x j))" using \i \ I\ by (auto intro!: prod.cong) show ?case unfolding product_integral_insert[OF insert(1,2) prod[OF subset_refl]] by (simp add: * insert prod subset_insertI) qed lemma integrable_subalgebra: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes borel: "f \ borel_measurable N" and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ emeasure N A = emeasure M A" shows "integrable N f \ integrable M f" (is ?P) proof - have "f \ borel_measurable M" using assms by (auto simp: measurable_def) with assms show ?thesis using assms by (auto simp: integrable_iff_bounded nn_integral_subalgebra) qed lemma integral_subalgebra: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes borel: "f \ borel_measurable N" and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ emeasure N A = emeasure M A" shows "integral\<^sup>L N f = integral\<^sup>L M f" proof cases assume "integrable N f" then show ?thesis proof induct case base with assms show ?case by (auto simp: subset_eq measure_def) next case (add f g) then have "(\ a. f a + g a \N) = integral\<^sup>L M f + integral\<^sup>L M g" by simp also have "\ = (\ a. f a + g a \M)" using add integrable_subalgebra[OF _ N, of f] integrable_subalgebra[OF _ N, of g] by simp finally show ?case . next case (lim f s) then have M: "\i. integrable M (s i)" "integrable M f" using integrable_subalgebra[OF _ N, of f] integrable_subalgebra[OF _ N, of "s i" for i] by simp_all show ?case proof (intro LIMSEQ_unique) show "(\i. integral\<^sup>L N (s i)) \ integral\<^sup>L N f" apply (rule integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) using lim apply auto done show "(\i. integral\<^sup>L N (s i)) \ integral\<^sup>L M f" unfolding lim apply (rule integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) using lim M N(2) apply auto done qed qed qed (simp add: not_integrable_integral_eq integrable_subalgebra[OF assms]) hide_const (open) simple_bochner_integral hide_const (open) simple_bochner_integrable end diff --git a/src/HOL/Analysis/Finite_Cartesian_Product.thy b/src/HOL/Analysis/Finite_Cartesian_Product.thy --- a/src/HOL/Analysis/Finite_Cartesian_Product.thy +++ b/src/HOL/Analysis/Finite_Cartesian_Product.thy @@ -1,1284 +1,1284 @@ (* Title: HOL/Analysis/Finite_Cartesian_Product.thy Author: Amine Chaieb, University of Cambridge *) section \Definition of Finite Cartesian Product Type\ theory Finite_Cartesian_Product imports Euclidean_Space L2_Norm "HOL-Library.Numeral_Type" "HOL-Library.Countable_Set" "HOL-Library.FuncSet" begin subsection\<^marker>\tag unimportant\ \Finite Cartesian products, with indexing and lambdas\ typedef ('a, 'b) vec = "UNIV :: ('b::finite \ 'a) set" morphisms vec_nth vec_lambda .. declare vec_lambda_inject [simplified, simp] bundle vec_syntax begin notation vec_nth (infixl "$" 90) and vec_lambda (binder "\" 10) end bundle no_vec_syntax begin no_notation vec_nth (infixl "$" 90) and vec_lambda (binder "\" 10) end unbundle vec_syntax text \ Concrete syntax for \('a, 'b) vec\: \<^item> \'a^'b\ becomes \('a, 'b::finite) vec\ \<^item> \'a^'b::_\ becomes \('a, 'b) vec\ without extra sort-constraint \ syntax "_vec_type" :: "type \ type \ type" (infixl "^" 15) parse_translation \ let fun vec t u = Syntax.const \<^type_syntax>\vec\ $ t $ u; fun finite_vec_tr [t, u] = (case Term_Position.strip_positions u of v as Free (x, _) => if Lexicon.is_tid x then vec t (Syntax.const \<^syntax_const>\_ofsort\ $ v $ Syntax.const \<^class_syntax>\finite\) else vec t u | _ => vec t u) in [(\<^syntax_const>\_vec_type\, K finite_vec_tr)] end \ lemma vec_eq_iff: "(x = y) \ (\i. x$i = y$i)" by (simp add: vec_nth_inject [symmetric] fun_eq_iff) lemma vec_lambda_beta [simp]: "vec_lambda g $ i = g i" by (simp add: vec_lambda_inverse) lemma vec_lambda_unique: "(\i. f$i = g i) \ vec_lambda g = f" by (auto simp add: vec_eq_iff) lemma vec_lambda_eta [simp]: "(\ i. (g$i)) = g" by (simp add: vec_eq_iff) subsection \Cardinality of vectors\ instance vec :: (finite, finite) finite proof show "finite (UNIV :: ('a, 'b) vec set)" proof (subst bij_betw_finite) show "bij_betw vec_nth UNIV (Pi (UNIV :: 'b set) (\_. UNIV :: 'a set))" by (intro bij_betwI[of _ _ _ vec_lambda]) (auto simp: vec_eq_iff) have "finite (PiE (UNIV :: 'b set) (\_. UNIV :: 'a set))" by (intro finite_PiE) auto also have "(PiE (UNIV :: 'b set) (\_. UNIV :: 'a set)) = Pi UNIV (\_. UNIV)" by auto finally show "finite \" . qed qed lemma countable_PiE: "finite I \ (\i. i \ I \ countable (F i)) \ countable (Pi\<^sub>E I F)" by (induct I arbitrary: F rule: finite_induct) (auto simp: PiE_insert_eq) instance vec :: (countable, finite) countable proof have "countable (UNIV :: ('a, 'b) vec set)" proof (rule countableI_bij2) show "bij_betw vec_nth UNIV (Pi (UNIV :: 'b set) (\_. UNIV :: 'a set))" by (intro bij_betwI[of _ _ _ vec_lambda]) (auto simp: vec_eq_iff) have "countable (PiE (UNIV :: 'b set) (\_. UNIV :: 'a set))" by (intro countable_PiE) auto also have "(PiE (UNIV :: 'b set) (\_. UNIV :: 'a set)) = Pi UNIV (\_. UNIV)" by auto finally show "countable \" . qed thus "\t::('a, 'b) vec \ nat. inj t" by (auto elim!: countableE) qed lemma infinite_UNIV_vec: assumes "infinite (UNIV :: 'a set)" shows "infinite (UNIV :: ('a^'b) set)" proof (subst bij_betw_finite) show "bij_betw vec_nth UNIV (Pi (UNIV :: 'b set) (\_. UNIV :: 'a set))" by (intro bij_betwI[of _ _ _ vec_lambda]) (auto simp: vec_eq_iff) have "infinite (PiE (UNIV :: 'b set) (\_. UNIV :: 'a set))" (is "infinite ?A") proof assume "finite ?A" hence "finite ((\f. f undefined) ` ?A)" by (rule finite_imageI) also have "(\f. f undefined) ` ?A = UNIV" by auto finally show False using \infinite (UNIV :: 'a set)\ by contradiction qed also have "?A = Pi UNIV (\_. UNIV)" by auto finally show "infinite (Pi (UNIV :: 'b set) (\_. UNIV :: 'a set))" . qed proposition CARD_vec [simp]: "CARD('a^'b) = CARD('a) ^ CARD('b)" proof (cases "finite (UNIV :: 'a set)") case True show ?thesis proof (subst bij_betw_same_card) show "bij_betw vec_nth UNIV (Pi (UNIV :: 'b set) (\_. UNIV :: 'a set))" by (intro bij_betwI[of _ _ _ vec_lambda]) (auto simp: vec_eq_iff) have "CARD('a) ^ CARD('b) = card (PiE (UNIV :: 'b set) (\_. UNIV :: 'a set))" (is "_ = card ?A") by (subst card_PiE) (auto) also have "?A = Pi UNIV (\_. UNIV)" by auto finally show "card \ = CARD('a) ^ CARD('b)" .. qed qed (simp_all add: infinite_UNIV_vec) lemma countable_vector: fixes B:: "'n::finite \ 'a set" assumes "\i. countable (B i)" shows "countable {V. \i::'n::finite. V $ i \ B i}" proof - have "f \ ($) ` {V. \i. V $ i \ B i}" if "f \ Pi\<^sub>E UNIV B" for f proof - have "\W. (\i. W $ i \ B i) \ ($) W = f" by (metis that PiE_iff UNIV_I vec_lambda_inverse) then show "f \ ($) ` {v. \i. v $ i \ B i}" by blast qed then have "Pi\<^sub>E UNIV B = vec_nth ` {V. \i::'n. V $ i \ B i}" by blast then have "countable (vec_nth ` {V. \i. V $ i \ B i})" by (metis finite_class.finite_UNIV countable_PiE assms) then have "countable (vec_lambda ` vec_nth ` {V. \i. V $ i \ B i})" by auto then show ?thesis by (simp add: image_comp o_def vec_nth_inverse) qed subsection\<^marker>\tag unimportant\ \Group operations and class instances\ instantiation vec :: (zero, finite) zero begin definition "0 \ (\ i. 0)" instance .. end instantiation vec :: (plus, finite) plus begin definition "(+) \ (\ x y. (\ i. x$i + y$i))" instance .. end instantiation vec :: (minus, finite) minus begin definition "(-) \ (\ x y. (\ i. x$i - y$i))" instance .. end instantiation vec :: (uminus, finite) uminus begin definition "uminus \ (\ x. (\ i. - (x$i)))" instance .. end lemma zero_index [simp]: "0 $ i = 0" unfolding zero_vec_def by simp lemma vector_add_component [simp]: "(x + y)$i = x$i + y$i" unfolding plus_vec_def by simp lemma vector_minus_component [simp]: "(x - y)$i = x$i - y$i" unfolding minus_vec_def by simp lemma vector_uminus_component [simp]: "(- x)$i = - (x$i)" unfolding uminus_vec_def by simp instance vec :: (semigroup_add, finite) semigroup_add by standard (simp add: vec_eq_iff add.assoc) instance vec :: (ab_semigroup_add, finite) ab_semigroup_add by standard (simp add: vec_eq_iff add.commute) instance vec :: (monoid_add, finite) monoid_add by standard (simp_all add: vec_eq_iff) instance vec :: (comm_monoid_add, finite) comm_monoid_add by standard (simp add: vec_eq_iff) instance vec :: (cancel_semigroup_add, finite) cancel_semigroup_add by standard (simp_all add: vec_eq_iff) instance vec :: (cancel_ab_semigroup_add, finite) cancel_ab_semigroup_add by standard (simp_all add: vec_eq_iff diff_diff_eq) instance vec :: (cancel_comm_monoid_add, finite) cancel_comm_monoid_add .. instance vec :: (group_add, finite) group_add by standard (simp_all add: vec_eq_iff) instance vec :: (ab_group_add, finite) ab_group_add by standard (simp_all add: vec_eq_iff) subsection\<^marker>\tag unimportant\\Basic componentwise operations on vectors\ instantiation vec :: (times, finite) times begin definition "(*) \ (\ x y. (\ i. (x$i) * (y$i)))" instance .. end instantiation vec :: (one, finite) one begin definition "1 \ (\ i. 1)" instance .. end instantiation vec :: (ord, finite) ord begin definition "x \ y \ (\i. x$i \ y$i)" definition "x < (y::'a^'b) \ x \ y \ \ y \ x" instance .. end text\The ordering on one-dimensional vectors is linear.\ instance vec:: (order, finite) order by standard (auto simp: less_eq_vec_def less_vec_def vec_eq_iff intro: order.trans order.antisym order.strict_implies_order) instance vec :: (linorder, CARD_1) linorder proof obtain a :: 'b where all: "\P. (\i. P i) \ P a" proof - have "CARD ('b) = 1" by (rule CARD_1) then obtain b :: 'b where "UNIV = {b}" by (auto iff: card_Suc_eq) then have "\P. (\i\UNIV. P i) \ P b" by auto then show thesis by (auto intro: that) qed fix x y :: "'a^'b::CARD_1" note [simp] = less_eq_vec_def less_vec_def all vec_eq_iff field_simps show "x \ y \ y \ x" by auto qed text\Constant Vectors\ definition "vec x = (\ i. x)" text\Also the scalar-vector multiplication.\ definition vector_scalar_mult:: "'a::times \ 'a ^ 'n \ 'a ^ 'n" (infixl "*s" 70) where "c *s x = (\ i. c * (x$i))" text \scalar product\ definition scalar_product :: "'a :: semiring_1 ^ 'n \ 'a ^ 'n \ 'a" where "scalar_product v w = (\ i \ UNIV. v $ i * w $ i)" subsection \Real vector space\ instantiation\<^marker>\tag unimportant\ vec :: (real_vector, finite) real_vector begin definition\<^marker>\tag important\ "scaleR \ (\ r x. (\ i. scaleR r (x$i)))" lemma vector_scaleR_component [simp]: "(scaleR r x)$i = scaleR r (x$i)" unfolding scaleR_vec_def by simp instance\<^marker>\tag unimportant\ by standard (simp_all add: vec_eq_iff scaleR_left_distrib scaleR_right_distrib) end subsection \Topological space\ instantiation\<^marker>\tag unimportant\ vec :: (topological_space, finite) topological_space begin definition\<^marker>\tag important\ [code del]: "open (S :: ('a ^ 'b) set) \ (\x\S. \A. (\i. open (A i) \ x$i \ A i) \ (\y. (\i. y$i \ A i) \ y \ S))" instance\<^marker>\tag unimportant\ proof show "open (UNIV :: ('a ^ 'b) set)" unfolding open_vec_def by auto next fix S T :: "('a ^ 'b) set" assume "open S" "open T" thus "open (S \ T)" unfolding open_vec_def apply clarify apply (drule (1) bspec)+ apply (clarify, rename_tac Sa Ta) apply (rule_tac x="\i. Sa i \ Ta i" in exI) apply (simp add: open_Int) done next fix K :: "('a ^ 'b) set set" assume "\S\K. open S" thus "open (\K)" unfolding open_vec_def apply clarify apply (drule (1) bspec) apply (drule (1) bspec) apply clarify apply (rule_tac x=A in exI) apply fast done qed end lemma open_vector_box: "\i. open (S i) \ open {x. \i. x $ i \ S i}" unfolding open_vec_def by auto lemma open_vimage_vec_nth: "open S \ open ((\x. x $ i) -` S)" unfolding open_vec_def apply clarify apply (rule_tac x="\k. if k = i then S else UNIV" in exI, simp) done lemma closed_vimage_vec_nth: "closed S \ closed ((\x. x $ i) -` S)" unfolding closed_open vimage_Compl [symmetric] by (rule open_vimage_vec_nth) lemma closed_vector_box: "\i. closed (S i) \ closed {x. \i. x $ i \ S i}" proof - have "{x. \i. x $ i \ S i} = (\i. (\x. x $ i) -` S i)" by auto thus "\i. closed (S i) \ closed {x. \i. x $ i \ S i}" by (simp add: closed_INT closed_vimage_vec_nth) qed lemma tendsto_vec_nth [tendsto_intros]: assumes "((\x. f x) \ a) net" shows "((\x. f x $ i) \ a $ i) net" proof (rule topological_tendstoI) fix S assume "open S" "a $ i \ S" then have "open ((\y. y $ i) -` S)" "a \ ((\y. y $ i) -` S)" by (simp_all add: open_vimage_vec_nth) with assms have "eventually (\x. f x \ (\y. y $ i) -` S) net" by (rule topological_tendstoD) then show "eventually (\x. f x $ i \ S) net" by simp qed lemma isCont_vec_nth [simp]: "isCont f a \ isCont (\x. f x $ i) a" unfolding isCont_def by (rule tendsto_vec_nth) lemma vec_tendstoI: assumes "\i. ((\x. f x $ i) \ a $ i) net" shows "((\x. f x) \ a) net" proof (rule topological_tendstoI) fix S assume "open S" and "a \ S" then obtain A where A: "\i. open (A i)" "\i. a $ i \ A i" and S: "\y. \i. y $ i \ A i \ y \ S" unfolding open_vec_def by metis have "\i. eventually (\x. f x $ i \ A i) net" using assms A by (rule topological_tendstoD) hence "eventually (\x. \i. f x $ i \ A i) net" by (rule eventually_all_finite) thus "eventually (\x. f x \ S) net" by (rule eventually_mono, simp add: S) qed lemma tendsto_vec_lambda [tendsto_intros]: assumes "\i. ((\x. f x i) \ a i) net" shows "((\x. \ i. f x i) \ (\ i. a i)) net" using assms by (simp add: vec_tendstoI) lemma open_image_vec_nth: assumes "open S" shows "open ((\x. x $ i) ` S)" proof (rule openI) fix a assume "a \ (\x. x $ i) ` S" then obtain z where "a = z $ i" and "z \ S" .. then obtain A where A: "\i. open (A i) \ z $ i \ A i" and S: "\y. (\i. y $ i \ A i) \ y \ S" using \open S\ unfolding open_vec_def by auto hence "A i \ (\x. x $ i) ` S" by (clarsimp, rule_tac x="\ j. if j = i then x else z $ j" in image_eqI, simp_all) hence "open (A i) \ a \ A i \ A i \ (\x. x $ i) ` S" using A \a = z $ i\ by simp then show "\T. open T \ a \ T \ T \ (\x. x $ i) ` S" by - (rule exI) qed instance\<^marker>\tag unimportant\ vec :: (perfect_space, finite) perfect_space proof fix x :: "'a ^ 'b" show "\ open {x}" proof assume "open {x}" hence "\i. open ((\x. x $ i) ` {x})" by (fast intro: open_image_vec_nth) hence "\i. open {x $ i}" by simp thus "False" by (simp add: not_open_singleton) qed qed subsection \Metric space\ (* TODO: Product of uniform spaces and compatibility with metric_spaces! *) instantiation\<^marker>\tag unimportant\ vec :: (metric_space, finite) dist begin definition\<^marker>\tag important\ "dist x y = L2_set (\i. dist (x$i) (y$i)) UNIV" instance .. end instantiation\<^marker>\tag unimportant\ vec :: (metric_space, finite) uniformity_dist begin definition\<^marker>\tag important\ [code del]: "(uniformity :: (('a^'b::_) \ ('a^'b::_)) filter) = (INF e\{0 <..}. principal {(x, y). dist x y < e})" instance\<^marker>\tag unimportant\ by standard (rule uniformity_vec_def) end declare uniformity_Abort[where 'a="'a :: metric_space ^ 'b :: finite", code] instantiation\<^marker>\tag unimportant\ vec :: (metric_space, finite) metric_space begin proposition dist_vec_nth_le: "dist (x $ i) (y $ i) \ dist x y" unfolding dist_vec_def by (rule member_le_L2_set) simp_all instance proof fix x y :: "'a ^ 'b" show "dist x y = 0 \ x = y" unfolding dist_vec_def by (simp add: L2_set_eq_0_iff vec_eq_iff) next fix x y z :: "'a ^ 'b" show "dist x y \ dist x z + dist y z" unfolding dist_vec_def apply (rule order_trans [OF _ L2_set_triangle_ineq]) apply (simp add: L2_set_mono dist_triangle2) done next fix S :: "('a ^ 'b) set" have *: "open S \ (\x\S. \e>0. \y. dist y x < e \ y \ S)" proof assume "open S" show "\x\S. \e>0. \y. dist y x < e \ y \ S" proof fix x assume "x \ S" obtain A where A: "\i. open (A i)" "\i. x $ i \ A i" and S: "\y. (\i. y $ i \ A i) \ y \ S" using \open S\ and \x \ S\ unfolding open_vec_def by metis have "\i\UNIV. \r>0. \y. dist y (x $ i) < r \ y \ A i" using A unfolding open_dist by simp hence "\r. \i\UNIV. 0 < r i \ (\y. dist y (x $ i) < r i \ y \ A i)" by (rule finite_set_choice [OF finite]) then obtain r where r1: "\i. 0 < r i" and r2: "\i y. dist y (x $ i) < r i \ y \ A i" by fast have "0 < Min (range r) \ (\y. dist y x < Min (range r) \ y \ S)" by (simp add: r1 r2 S le_less_trans [OF dist_vec_nth_le]) thus "\e>0. \y. dist y x < e \ y \ S" .. qed next assume *: "\x\S. \e>0. \y. dist y x < e \ y \ S" show "open S" proof (unfold open_vec_def, rule) fix x assume "x \ S" then obtain e where "0 < e" and S: "\y. dist y x < e \ y \ S" using * by fast define r where [abs_def]: "r i = e / sqrt (of_nat CARD('b))" for i :: 'b from \0 < e\ have r: "\i. 0 < r i" unfolding r_def by simp_all from \0 < e\ have e: "e = L2_set r UNIV" unfolding r_def by (simp add: L2_set_constant) define A where "A i = {y. dist (x $ i) y < r i}" for i have "\i. open (A i) \ x $ i \ A i" unfolding A_def by (simp add: open_ball r) moreover have "\y. (\i. y $ i \ A i) \ y \ S" by (simp add: A_def S dist_vec_def e L2_set_strict_mono dist_commute) ultimately show "\A. (\i. open (A i) \ x $ i \ A i) \ (\y. (\i. y $ i \ A i) \ y \ S)" by metis qed qed show "open S = (\x\S. \\<^sub>F (x', y) in uniformity. x' = x \ y \ S)" unfolding * eventually_uniformity_metric by (simp del: split_paired_All add: dist_vec_def dist_commute) qed end lemma Cauchy_vec_nth: "Cauchy (\n. X n) \ Cauchy (\n. X n $ i)" unfolding Cauchy_def by (fast intro: le_less_trans [OF dist_vec_nth_le]) lemma vec_CauchyI: fixes X :: "nat \ 'a::metric_space ^ 'n" assumes X: "\i. Cauchy (\n. X n $ i)" shows "Cauchy (\n. X n)" proof (rule metric_CauchyI) fix r :: real assume "0 < r" hence "0 < r / of_nat CARD('n)" (is "0 < ?s") by simp define N where "N i = (LEAST N. \m\N. \n\N. dist (X m $ i) (X n $ i) < ?s)" for i define M where "M = Max (range N)" have "\i. \N. \m\N. \n\N. dist (X m $ i) (X n $ i) < ?s" using X \0 < ?s\ by (rule metric_CauchyD) hence "\i. \m\N i. \n\N i. dist (X m $ i) (X n $ i) < ?s" unfolding N_def by (rule LeastI_ex) hence M: "\i. \m\M. \n\M. dist (X m $ i) (X n $ i) < ?s" unfolding M_def by simp { fix m n :: nat assume "M \ m" "M \ n" have "dist (X m) (X n) = L2_set (\i. dist (X m $ i) (X n $ i)) UNIV" unfolding dist_vec_def .. also have "\ \ sum (\i. dist (X m $ i) (X n $ i)) UNIV" by (rule L2_set_le_sum [OF zero_le_dist]) also have "\ < sum (\i::'n. ?s) UNIV" by (rule sum_strict_mono, simp_all add: M \M \ m\ \M \ n\) also have "\ = r" by simp finally have "dist (X m) (X n) < r" . } hence "\m\M. \n\M. dist (X m) (X n) < r" by simp then show "\M. \m\M. \n\M. dist (X m) (X n) < r" .. qed instance\<^marker>\tag unimportant\ vec :: (complete_space, finite) complete_space proof fix X :: "nat \ 'a ^ 'b" assume "Cauchy X" have "\i. (\n. X n $ i) \ lim (\n. X n $ i)" using Cauchy_vec_nth [OF \Cauchy X\] by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) hence "X \ vec_lambda (\i. lim (\n. X n $ i))" by (simp add: vec_tendstoI) then show "convergent X" by (rule convergentI) qed subsection \Normed vector space\ instantiation\<^marker>\tag unimportant\ vec :: (real_normed_vector, finite) real_normed_vector begin definition\<^marker>\tag important\ "norm x = L2_set (\i. norm (x$i)) UNIV" definition\<^marker>\tag important\ "sgn (x::'a^'b) = scaleR (inverse (norm x)) x" instance\<^marker>\tag unimportant\ proof fix a :: real and x y :: "'a ^ 'b" show "norm x = 0 \ x = 0" unfolding norm_vec_def by (simp add: L2_set_eq_0_iff vec_eq_iff) show "norm (x + y) \ norm x + norm y" unfolding norm_vec_def apply (rule order_trans [OF _ L2_set_triangle_ineq]) apply (simp add: L2_set_mono norm_triangle_ineq) done show "norm (scaleR a x) = \a\ * norm x" unfolding norm_vec_def by (simp add: L2_set_right_distrib) show "sgn x = scaleR (inverse (norm x)) x" by (rule sgn_vec_def) show "dist x y = norm (x - y)" unfolding dist_vec_def norm_vec_def by (simp add: dist_norm) qed end lemma norm_nth_le: "norm (x $ i) \ norm x" unfolding norm_vec_def by (rule member_le_L2_set) simp_all lemma norm_le_componentwise_cart: fixes x :: "'a::real_normed_vector^'n" assumes "\i. norm(x$i) \ norm(y$i)" shows "norm x \ norm y" unfolding norm_vec_def by (rule L2_set_mono) (auto simp: assms) lemma component_le_norm_cart: "\x$i\ \ norm x" apply (simp add: norm_vec_def) apply (rule member_le_L2_set, simp_all) done lemma norm_bound_component_le_cart: "norm x \ e ==> \x$i\ \ e" by (metis component_le_norm_cart order_trans) lemma norm_bound_component_lt_cart: "norm x < e ==> \x$i\ < e" by (metis component_le_norm_cart le_less_trans) lemma norm_le_l1_cart: "norm x \ sum(\i. \x$i\) UNIV" by (simp add: norm_vec_def L2_set_le_sum) lemma bounded_linear_vec_nth[intro]: "bounded_linear (\x. x $ i)" apply standard apply (rule vector_add_component) apply (rule vector_scaleR_component) apply (rule_tac x="1" in exI, simp add: norm_nth_le) done instance vec :: (banach, finite) banach .. subsection \Inner product space\ instantiation\<^marker>\tag unimportant\ vec :: (real_inner, finite) real_inner begin definition\<^marker>\tag important\ "inner x y = sum (\i. inner (x$i) (y$i)) UNIV" instance\<^marker>\tag unimportant\ proof fix r :: real and x y z :: "'a ^ 'b" show "inner x y = inner y x" unfolding inner_vec_def by (simp add: inner_commute) show "inner (x + y) z = inner x z + inner y z" unfolding inner_vec_def by (simp add: inner_add_left sum.distrib) show "inner (scaleR r x) y = r * inner x y" unfolding inner_vec_def by (simp add: sum_distrib_left) show "0 \ inner x x" unfolding inner_vec_def by (simp add: sum_nonneg) show "inner x x = 0 \ x = 0" unfolding inner_vec_def by (simp add: vec_eq_iff sum_nonneg_eq_0_iff) show "norm x = sqrt (inner x x)" unfolding inner_vec_def norm_vec_def L2_set_def by (simp add: power2_norm_eq_inner) qed end subsection \Euclidean space\ text \Vectors pointing along a single axis.\ definition\<^marker>\tag important\ "axis k x = (\ i. if i = k then x else 0)" lemma axis_nth [simp]: "axis i x $ i = x" unfolding axis_def by simp lemma axis_eq_axis: "axis i x = axis j y \ x = y \ i = j \ x = 0 \ y = 0" unfolding axis_def vec_eq_iff by auto lemma inner_axis_axis: "inner (axis i x) (axis j y) = (if i = j then inner x y else 0)" unfolding inner_vec_def apply (cases "i = j") apply clarsimp apply (subst sum.remove [of _ j], simp_all) apply (rule sum.neutral, simp add: axis_def) apply (rule sum.neutral, simp add: axis_def) done lemma inner_axis: "inner x (axis i y) = inner (x $ i) y" by (simp add: inner_vec_def axis_def sum.remove [where x=i]) lemma inner_axis': "inner(axis i y) x = inner y (x $ i)" by (simp add: inner_axis inner_commute) instantiation\<^marker>\tag unimportant\ vec :: (euclidean_space, finite) euclidean_space begin definition\<^marker>\tag important\ "Basis = (\i. \u\Basis. {axis i u})" instance\<^marker>\tag unimportant\ proof show "(Basis :: ('a ^ 'b) set) \ {}" unfolding Basis_vec_def by simp next show "finite (Basis :: ('a ^ 'b) set)" unfolding Basis_vec_def by simp next fix u v :: "'a ^ 'b" assume "u \ Basis" and "v \ Basis" thus "inner u v = (if u = v then 1 else 0)" unfolding Basis_vec_def by (auto simp add: inner_axis_axis axis_eq_axis inner_Basis) next fix x :: "'a ^ 'b" show "(\u\Basis. inner x u = 0) \ x = 0" unfolding Basis_vec_def by (simp add: inner_axis euclidean_all_zero_iff vec_eq_iff) qed proposition DIM_cart [simp]: "DIM('a^'b) = CARD('b) * DIM('a)" proof - have "card (\i::'b. \u::'a\Basis. {axis i u}) = (\i::'b\UNIV. card (\u::'a\Basis. {axis i u}))" by (rule card_UN_disjoint) (auto simp: axis_eq_axis) also have "... = CARD('b) * DIM('a)" by (subst card_UN_disjoint) (auto simp: axis_eq_axis) finally show ?thesis by (simp add: Basis_vec_def) qed end lemma norm_axis_1 [simp]: "norm (axis m (1::real)) = 1" by (simp add: inner_axis' norm_eq_1) lemma sum_norm_allsubsets_bound_cart: fixes f:: "'a \ real ^'n" assumes fP: "finite P" and fPs: "\Q. Q \ P \ norm (sum f Q) \ e" shows "sum (\x. norm (f x)) P \ 2 * real CARD('n) * e" using sum_norm_allsubsets_bound[OF assms] by simp lemma cart_eq_inner_axis: "a $ i = inner a (axis i 1)" by (simp add: inner_axis) lemma axis_eq_0_iff [simp]: shows "axis m x = 0 \ x = 0" by (simp add: axis_def vec_eq_iff) lemma axis_in_Basis_iff [simp]: "axis i a \ Basis \ a \ Basis" by (auto simp: Basis_vec_def axis_eq_axis) text\Mapping each basis element to the corresponding finite index\ definition axis_index :: "('a::comm_ring_1)^'n \ 'n" where "axis_index v \ SOME i. v = axis i 1" lemma axis_inverse: fixes v :: "real^'n" assumes "v \ Basis" shows "\i. v = axis i 1" proof - have "v \ (\n. \r\Basis. {axis n r})" using assms Basis_vec_def by blast then show ?thesis by (force simp add: vec_eq_iff) qed lemma axis_index: fixes v :: "real^'n" assumes "v \ Basis" shows "v = axis (axis_index v) 1" by (metis (mono_tags) assms axis_inverse axis_index_def someI_ex) lemma axis_index_axis [simp]: fixes UU :: "real^'n" shows "(axis_index (axis u 1 :: real^'n)) = (u::'n)" by (simp add: axis_eq_axis axis_index_def) subsection\<^marker>\tag unimportant\ \A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space\ lemma sum_cong_aux: "(\x. x \ A \ f x = g x) \ sum f A = sum g A" by (auto intro: sum.cong) hide_fact (open) sum_cong_aux method_setup vector = \ let val ss1 = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps [@{thm sum.distrib} RS sym, @{thm sum_subtractf} RS sym, @{thm sum_distrib_left}, @{thm sum_distrib_right}, @{thm sum_negf} RS sym]) val ss2 = simpset_of (\<^context> addsimps [@{thm plus_vec_def}, @{thm times_vec_def}, @{thm minus_vec_def}, @{thm uminus_vec_def}, @{thm one_vec_def}, @{thm zero_vec_def}, @{thm vec_def}, @{thm scaleR_vec_def}, @{thm vector_scalar_mult_def}]) fun vector_arith_tac ctxt ths = simp_tac (put_simpset ss1 ctxt) THEN' (fn i => resolve_tac ctxt @{thms Finite_Cartesian_Product.sum_cong_aux} i ORELSE resolve_tac ctxt @{thms sum.neutral} i ORELSE simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm vec_eq_iff}]) i) (* THEN' TRY o clarify_tac HOL_cs THEN' (TRY o rtac @{thm iffI}) *) THEN' asm_full_simp_tac (put_simpset ss2 ctxt addsimps ths) in Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (vector_arith_tac ctxt ths)) end \ "lift trivial vector statements to real arith statements" lemma vec_0[simp]: "vec 0 = 0" by vector lemma vec_1[simp]: "vec 1 = 1" by vector lemma vec_inj[simp]: "vec x = vec y \ x = y" by vector lemma vec_in_image_vec: "vec x \ (vec ` S) \ x \ S" by auto lemma vec_add: "vec(x + y) = vec x + vec y" by vector lemma vec_sub: "vec(x - y) = vec x - vec y" by vector lemma vec_cmul: "vec(c * x) = c *s vec x " by vector lemma vec_neg: "vec(- x) = - vec x " by vector lemma vec_scaleR: "vec(c * x) = c *\<^sub>R vec x" by vector lemma vec_sum: assumes "finite S" shows "vec(sum f S) = sum (vec \ f) S" using assms proof induct case empty then show ?case by simp next case insert then show ?case by (auto simp add: vec_add) qed text\Obvious "component-pushing".\ lemma vec_component [simp]: "vec x $ i = x" by vector lemma vector_mult_component [simp]: "(x * y)$i = x$i * y$i" by vector lemma vector_smult_component [simp]: "(c *s y)$i = c * (y$i)" by vector lemma cond_component: "(if b then x else y)$i = (if b then x$i else y$i)" by vector lemmas\<^marker>\tag unimportant\ vector_component = vec_component vector_add_component vector_mult_component vector_smult_component vector_minus_component vector_uminus_component vector_scaleR_component cond_component subsection\<^marker>\tag unimportant\ \Some frequently useful arithmetic lemmas over vectors\ instance vec :: (semigroup_mult, finite) semigroup_mult by standard (vector mult.assoc) instance vec :: (monoid_mult, finite) monoid_mult by standard vector+ instance vec :: (ab_semigroup_mult, finite) ab_semigroup_mult by standard (vector mult.commute) instance vec :: (comm_monoid_mult, finite) comm_monoid_mult by standard vector instance vec :: (semiring, finite) semiring by standard (vector field_simps)+ instance vec :: (semiring_0, finite) semiring_0 by standard (vector field_simps)+ instance vec :: (semiring_1, finite) semiring_1 by standard vector instance vec :: (comm_semiring, finite) comm_semiring by standard (vector field_simps)+ instance vec :: (comm_semiring_0, finite) comm_semiring_0 .. instance vec :: (semiring_0_cancel, finite) semiring_0_cancel .. instance vec :: (comm_semiring_0_cancel, finite) comm_semiring_0_cancel .. instance vec :: (ring, finite) ring .. instance vec :: (semiring_1_cancel, finite) semiring_1_cancel .. instance vec :: (comm_semiring_1, finite) comm_semiring_1 .. instance vec :: (ring_1, finite) ring_1 .. instance vec :: (real_algebra, finite) real_algebra by standard (simp_all add: vec_eq_iff) instance vec :: (real_algebra_1, finite) real_algebra_1 .. lemma of_nat_index: "(of_nat n :: 'a::semiring_1 ^'n)$i = of_nat n" proof (induct n) case 0 then show ?case by vector next case Suc then show ?case by vector qed lemma one_index [simp]: "(1 :: 'a :: one ^ 'n) $ i = 1" by vector lemma neg_one_index [simp]: "(- 1 :: 'a :: {one, uminus} ^ 'n) $ i = - 1" by vector instance vec :: (semiring_char_0, finite) semiring_char_0 proof fix m n :: nat show "inj (of_nat :: nat \ 'a ^ 'b)" by (auto intro!: injI simp add: vec_eq_iff of_nat_index) qed instance vec :: (numeral, finite) numeral .. instance vec :: (semiring_numeral, finite) semiring_numeral .. lemma numeral_index [simp]: "numeral w $ i = numeral w" by (induct w) (simp_all only: numeral.simps vector_add_component one_index) lemma neg_numeral_index [simp]: "- numeral w $ i = - numeral w" by (simp only: vector_uminus_component numeral_index) instance vec :: (comm_ring_1, finite) comm_ring_1 .. instance vec :: (ring_char_0, finite) ring_char_0 .. lemma vector_smult_assoc: "a *s (b *s x) = ((a::'a::semigroup_mult) * b) *s x" by (vector mult.assoc) lemma vector_sadd_rdistrib: "((a::'a::semiring) + b) *s x = a *s x + b *s x" by (vector field_simps) lemma vector_add_ldistrib: "(c::'a::semiring) *s (x + y) = c *s x + c *s y" by (vector field_simps) lemma vector_smult_lzero[simp]: "(0::'a::mult_zero) *s x = 0" by vector lemma vector_smult_lid[simp]: "(1::'a::monoid_mult) *s x = x" by vector lemma vector_ssub_ldistrib: "(c::'a::ring) *s (x - y) = c *s x - c *s y" by (vector field_simps) lemma vector_smult_rneg: "(c::'a::ring) *s -x = -(c *s x)" by vector lemma vector_smult_lneg: "- (c::'a::ring) *s x = -(c *s x)" by vector lemma vector_sneg_minus1: "-x = (-1::'a::ring_1) *s x" by vector lemma vector_smult_rzero[simp]: "c *s 0 = (0::'a::mult_zero ^ 'n)" by vector lemma vector_sub_rdistrib: "((a::'a::ring) - b) *s x = a *s x - b *s x" by (vector field_simps) lemma vec_eq[simp]: "(vec m = vec n) \ (m = n)" by (simp add: vec_eq_iff) lemma Vector_Spaces_linear_vec [simp]: "Vector_Spaces.linear (*) vector_scalar_mult vec" by unfold_locales (vector algebra_simps)+ lemma vector_mul_eq_0[simp]: "(a *s x = 0) \ a = (0::'a::idom) \ x = 0" by vector lemma vector_mul_lcancel[simp]: "a *s x = a *s y \ a = (0::'a::field) \ x = y" by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_ssub_ldistrib) lemma vector_mul_rcancel[simp]: "a *s x = b *s x \ (a::'a::field) = b \ x = 0" - by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_sub_rdistrib) + by (subst eq_iff_diff_eq_0, subst vector_sub_rdistrib [symmetric]) simp lemma scalar_mult_eq_scaleR [abs_def]: "c *s x = c *\<^sub>R x" unfolding scaleR_vec_def vector_scalar_mult_def by simp lemma dist_mul[simp]: "dist (c *s x) (c *s y) = \c\ * dist x y" unfolding dist_norm scalar_mult_eq_scaleR unfolding scaleR_right_diff_distrib[symmetric] by simp lemma sum_component [simp]: fixes f:: " 'a \ ('b::comm_monoid_add) ^'n" shows "(sum f S)$i = sum (\x. (f x)$i) S" proof (cases "finite S") case True then show ?thesis by induct simp_all next case False then show ?thesis by simp qed lemma sum_eq: "sum f S = (\ i. sum (\x. (f x)$i ) S)" by (simp add: vec_eq_iff) lemma sum_cmul: fixes f:: "'c \ ('a::semiring_1)^'n" shows "sum (\x. c *s f x) S = c *s sum f S" by (simp add: vec_eq_iff sum_distrib_left) lemma linear_vec [simp]: "linear vec" using Vector_Spaces_linear_vec apply (auto ) by unfold_locales (vector algebra_simps)+ subsection \Matrix operations\ text\Matrix notation. NB: an MxN matrix is of type \<^typ>\'a^'n^'m\, not \<^typ>\'a^'m^'n\\ definition\<^marker>\tag important\ map_matrix::"('a \ 'b) \ (('a, 'i::finite)vec, 'j::finite) vec \ (('b, 'i)vec, 'j) vec" where "map_matrix f x = (\ i j. f (x $ i $ j))" lemma nth_map_matrix[simp]: "map_matrix f x $ i $ j = f (x $ i $ j)" by (simp add: map_matrix_def) definition\<^marker>\tag important\ matrix_matrix_mult :: "('a::semiring_1) ^'n^'m \ 'a ^'p^'n \ 'a ^ 'p ^'m" (infixl "**" 70) where "m ** m' == (\ i j. sum (\k. ((m$i)$k) * ((m'$k)$j)) (UNIV :: 'n set)) ::'a ^ 'p ^'m" definition\<^marker>\tag important\ matrix_vector_mult :: "('a::semiring_1) ^'n^'m \ 'a ^'n \ 'a ^ 'm" (infixl "*v" 70) where "m *v x \ (\ i. sum (\j. ((m$i)$j) * (x$j)) (UNIV ::'n set)) :: 'a^'m" definition\<^marker>\tag important\ vector_matrix_mult :: "'a ^ 'm \ ('a::semiring_1) ^'n^'m \ 'a ^'n " (infixl "v*" 70) where "v v* m == (\ j. sum (\i. ((m$i)$j) * (v$i)) (UNIV :: 'm set)) :: 'a^'n" definition\<^marker>\tag unimportant\ "(mat::'a::zero => 'a ^'n^'n) k = (\ i j. if i = j then k else 0)" definition\<^marker>\tag unimportant\ transpose where "(transpose::'a^'n^'m \ 'a^'m^'n) A = (\ i j. ((A$j)$i))" definition\<^marker>\tag unimportant\ "(row::'m => 'a ^'n^'m \ 'a ^'n) i A = (\ j. ((A$i)$j))" definition\<^marker>\tag unimportant\ "(column::'n =>'a^'n^'m =>'a^'m) j A = (\ i. ((A$i)$j))" definition\<^marker>\tag unimportant\ "rows(A::'a^'n^'m) = { row i A | i. i \ (UNIV :: 'm set)}" definition\<^marker>\tag unimportant\ "columns(A::'a^'n^'m) = { column i A | i. i \ (UNIV :: 'n set)}" lemma times0_left [simp]: "(0::'a::semiring_1^'n^'m) ** (A::'a ^'p^'n) = 0" by (simp add: matrix_matrix_mult_def zero_vec_def) lemma times0_right [simp]: "(A::'a::semiring_1^'n^'m) ** (0::'a ^'p^'n) = 0" by (simp add: matrix_matrix_mult_def zero_vec_def) lemma mat_0[simp]: "mat 0 = 0" by (vector mat_def) lemma matrix_add_ldistrib: "(A ** (B + C)) = (A ** B) + (A ** C)" by (vector matrix_matrix_mult_def sum.distrib[symmetric] field_simps) lemma matrix_mul_lid [simp]: fixes A :: "'a::semiring_1 ^ 'm ^ 'n" shows "mat 1 ** A = A" apply (simp add: matrix_matrix_mult_def mat_def) apply vector apply (auto simp only: if_distrib if_distribR sum.delta'[OF finite] mult_1_left mult_zero_left if_True UNIV_I) done lemma matrix_mul_rid [simp]: fixes A :: "'a::semiring_1 ^ 'm ^ 'n" shows "A ** mat 1 = A" apply (simp add: matrix_matrix_mult_def mat_def) apply vector apply (auto simp only: if_distrib if_distribR sum.delta[OF finite] mult_1_right mult_zero_right if_True UNIV_I cong: if_cong) done proposition matrix_mul_assoc: "A ** (B ** C) = (A ** B) ** C" apply (vector matrix_matrix_mult_def sum_distrib_left sum_distrib_right mult.assoc) apply (subst sum.swap) apply simp done proposition matrix_vector_mul_assoc: "A *v (B *v x) = (A ** B) *v x" apply (vector matrix_matrix_mult_def matrix_vector_mult_def sum_distrib_left sum_distrib_right mult.assoc) apply (subst sum.swap) apply simp done proposition scalar_matrix_assoc: fixes A :: "('a::real_algebra_1)^'m^'n" shows "k *\<^sub>R (A ** B) = (k *\<^sub>R A) ** B" by (simp add: matrix_matrix_mult_def sum_distrib_left mult_ac vec_eq_iff scaleR_sum_right) proposition matrix_scalar_ac: fixes A :: "('a::real_algebra_1)^'m^'n" shows "A ** (k *\<^sub>R B) = k *\<^sub>R A ** B" by (simp add: matrix_matrix_mult_def sum_distrib_left mult_ac vec_eq_iff) lemma matrix_vector_mul_lid [simp]: "mat 1 *v x = (x::'a::semiring_1 ^ 'n)" apply (vector matrix_vector_mult_def mat_def) apply (simp add: if_distrib if_distribR cong del: if_weak_cong) done lemma matrix_transpose_mul: "transpose(A ** B) = transpose B ** transpose (A::'a::comm_semiring_1^_^_)" by (simp add: matrix_matrix_mult_def transpose_def vec_eq_iff mult.commute) lemma matrix_mult_transpose_dot_column: shows "transpose A ** A = (\ i j. inner (column i A) (column j A))" by (simp add: matrix_matrix_mult_def vec_eq_iff transpose_def column_def inner_vec_def) lemma matrix_mult_transpose_dot_row: shows "A ** transpose A = (\ i j. inner (row i A) (row j A))" by (simp add: matrix_matrix_mult_def vec_eq_iff transpose_def row_def inner_vec_def) lemma matrix_eq: fixes A B :: "'a::semiring_1 ^ 'n ^ 'm" shows "A = B \ (\x. A *v x = B *v x)" (is "?lhs \ ?rhs") apply auto apply (subst vec_eq_iff) apply clarify apply (clarsimp simp add: matrix_vector_mult_def if_distrib if_distribR vec_eq_iff cong del: if_weak_cong) apply (erule_tac x="axis ia 1" in allE) apply (erule_tac x="i" in allE) apply (auto simp add: if_distrib if_distribR axis_def sum.delta[OF finite] cong del: if_weak_cong) done lemma matrix_vector_mul_component: "(A *v x)$k = inner (A$k) x" by (simp add: matrix_vector_mult_def inner_vec_def) lemma dot_lmul_matrix: "inner ((x::real ^_) v* A) y = inner x (A *v y)" apply (simp add: inner_vec_def matrix_vector_mult_def vector_matrix_mult_def sum_distrib_right sum_distrib_left ac_simps) apply (subst sum.swap) apply simp done lemma transpose_mat [simp]: "transpose (mat n) = mat n" by (vector transpose_def mat_def) lemma transpose_transpose [simp]: "transpose(transpose A) = A" by (vector transpose_def) lemma row_transpose [simp]: "row i (transpose A) = column i A" by (simp add: row_def column_def transpose_def vec_eq_iff) lemma column_transpose [simp]: "column i (transpose A) = row i A" by (simp add: row_def column_def transpose_def vec_eq_iff) lemma rows_transpose [simp]: "rows(transpose A) = columns A" by (auto simp add: rows_def columns_def intro: set_eqI) lemma columns_transpose [simp]: "columns(transpose A) = rows A" by (metis transpose_transpose rows_transpose) lemma transpose_scalar: "transpose (k *\<^sub>R A) = k *\<^sub>R transpose A" unfolding transpose_def by (simp add: vec_eq_iff) lemma transpose_iff [iff]: "transpose A = transpose B \ A = B" by (metis transpose_transpose) lemma matrix_mult_sum: "(A::'a::comm_semiring_1^'n^'m) *v x = sum (\i. (x$i) *s column i A) (UNIV:: 'n set)" by (simp add: matrix_vector_mult_def vec_eq_iff column_def mult.commute) lemma vector_componentwise: "(x::'a::ring_1^'n) = (\ j. \i\UNIV. (x$i) * (axis i 1 :: 'a^'n) $ j)" by (simp add: axis_def if_distrib sum.If_cases vec_eq_iff) lemma basis_expansion: "sum (\i. (x$i) *s axis i 1) UNIV = (x::('a::ring_1) ^'n)" by (auto simp add: axis_def vec_eq_iff if_distrib sum.If_cases cong del: if_weak_cong) text\Correspondence between matrices and linear operators.\ definition\<^marker>\tag important\ matrix :: "('a::{plus,times, one, zero}^'m \ 'a ^ 'n) \ 'a^'m^'n" where "matrix f = (\ i j. (f(axis j 1))$i)" lemma matrix_id_mat_1: "matrix id = mat 1" by (simp add: mat_def matrix_def axis_def) lemma matrix_scaleR: "(matrix ((*\<^sub>R) r)) = mat r" by (simp add: mat_def matrix_def axis_def if_distrib cong: if_cong) lemma matrix_vector_mul_linear[intro, simp]: "linear (\x. A *v (x::'a::real_algebra_1 ^ _))" by (simp add: linear_iff matrix_vector_mult_def vec_eq_iff field_simps sum_distrib_left sum.distrib scaleR_right.sum) lemma vector_matrix_left_distrib [algebra_simps]: shows "(x + y) v* A = x v* A + y v* A" unfolding vector_matrix_mult_def by (simp add: algebra_simps sum.distrib vec_eq_iff) lemma matrix_vector_right_distrib [algebra_simps]: "A *v (x + y) = A *v x + A *v y" by (vector matrix_vector_mult_def sum.distrib distrib_left) lemma matrix_vector_mult_diff_distrib [algebra_simps]: fixes A :: "'a::ring_1^'n^'m" shows "A *v (x - y) = A *v x - A *v y" by (vector matrix_vector_mult_def sum_subtractf right_diff_distrib) lemma matrix_vector_mult_scaleR[algebra_simps]: fixes A :: "real^'n^'m" shows "A *v (c *\<^sub>R x) = c *\<^sub>R (A *v x)" using linear_iff matrix_vector_mul_linear by blast lemma matrix_vector_mult_0_right [simp]: "A *v 0 = 0" by (simp add: matrix_vector_mult_def vec_eq_iff) lemma matrix_vector_mult_0 [simp]: "0 *v w = 0" by (simp add: matrix_vector_mult_def vec_eq_iff) lemma matrix_vector_mult_add_rdistrib [algebra_simps]: "(A + B) *v x = (A *v x) + (B *v x)" by (vector matrix_vector_mult_def sum.distrib distrib_right) lemma matrix_vector_mult_diff_rdistrib [algebra_simps]: fixes A :: "'a :: ring_1^'n^'m" shows "(A - B) *v x = (A *v x) - (B *v x)" by (vector matrix_vector_mult_def sum_subtractf left_diff_distrib) lemma matrix_vector_column: "(A::'a::comm_semiring_1^'n^_) *v x = sum (\i. (x$i) *s ((transpose A)$i)) (UNIV:: 'n set)" by (simp add: matrix_vector_mult_def transpose_def vec_eq_iff mult.commute) subsection\Inverse matrices (not necessarily square)\ definition\<^marker>\tag important\ "invertible(A::'a::semiring_1^'n^'m) \ (\A'::'a^'m^'n. A ** A' = mat 1 \ A' ** A = mat 1)" definition\<^marker>\tag important\ "matrix_inv(A:: 'a::semiring_1^'n^'m) = (SOME A'::'a^'m^'n. A ** A' = mat 1 \ A' ** A = mat 1)" lemma inj_matrix_vector_mult: fixes A::"'a::field^'n^'m" assumes "invertible A" shows "inj ((*v) A)" by (metis assms inj_on_inverseI invertible_def matrix_vector_mul_assoc matrix_vector_mul_lid) lemma scalar_invertible: fixes A :: "('a::real_algebra_1)^'m^'n" assumes "k \ 0" and "invertible A" shows "invertible (k *\<^sub>R A)" proof - obtain A' where "A ** A' = mat 1" and "A' ** A = mat 1" using assms unfolding invertible_def by auto with \k \ 0\ have "(k *\<^sub>R A) ** ((1/k) *\<^sub>R A') = mat 1" "((1/k) *\<^sub>R A') ** (k *\<^sub>R A) = mat 1" by (simp_all add: assms matrix_scalar_ac) thus "invertible (k *\<^sub>R A)" unfolding invertible_def by auto qed proposition scalar_invertible_iff: fixes A :: "('a::real_algebra_1)^'m^'n" assumes "k \ 0" and "invertible A" shows "invertible (k *\<^sub>R A) \ k \ 0 \ invertible A" by (simp add: assms scalar_invertible) lemma vector_transpose_matrix [simp]: "x v* transpose A = A *v x" unfolding transpose_def vector_matrix_mult_def matrix_vector_mult_def by simp lemma transpose_matrix_vector [simp]: "transpose A *v x = x v* A" unfolding transpose_def vector_matrix_mult_def matrix_vector_mult_def by simp lemma vector_scalar_commute: fixes A :: "'a::{field}^'m^'n" shows "A *v (c *s x) = c *s (A *v x)" by (simp add: vector_scalar_mult_def matrix_vector_mult_def mult_ac sum_distrib_left) lemma scalar_vector_matrix_assoc: fixes k :: "'a::{field}" and x :: "'a::{field}^'n" and A :: "'a^'m^'n" shows "(k *s x) v* A = k *s (x v* A)" by (metis transpose_matrix_vector vector_scalar_commute) lemma vector_matrix_mult_0 [simp]: "0 v* A = 0" unfolding vector_matrix_mult_def by (simp add: zero_vec_def) lemma vector_matrix_mult_0_right [simp]: "x v* 0 = 0" unfolding vector_matrix_mult_def by (simp add: zero_vec_def) lemma vector_matrix_mul_rid [simp]: fixes v :: "('a::semiring_1)^'n" shows "v v* mat 1 = v" by (metis matrix_vector_mul_lid transpose_mat vector_transpose_matrix) lemma scaleR_vector_matrix_assoc: fixes k :: real and x :: "real^'n" and A :: "real^'m^'n" shows "(k *\<^sub>R x) v* A = k *\<^sub>R (x v* A)" by (metis matrix_vector_mult_scaleR transpose_matrix_vector) proposition vector_scaleR_matrix_ac: fixes k :: real and x :: "real^'n" and A :: "real^'m^'n" shows "x v* (k *\<^sub>R A) = k *\<^sub>R (x v* A)" proof - have "x v* (k *\<^sub>R A) = (k *\<^sub>R x) v* A" unfolding vector_matrix_mult_def by (simp add: algebra_simps) with scaleR_vector_matrix_assoc show "x v* (k *\<^sub>R A) = k *\<^sub>R (x v* A)" by auto qed end diff --git a/src/HOL/Analysis/Further_Topology.thy b/src/HOL/Analysis/Further_Topology.thy --- a/src/HOL/Analysis/Further_Topology.thy +++ b/src/HOL/Analysis/Further_Topology.thy @@ -1,5706 +1,5706 @@ section \Extending Continous Maps, Invariance of Domain, etc\ (*FIX rename? *) text\Ported from HOL Light (moretop.ml) by L C Paulson\ theory Further_Topology imports Weierstrass_Theorems Polytope Complex_Transcendental Equivalence_Lebesgue_Henstock_Integration Retracts begin subsection\A map from a sphere to a higher dimensional sphere is nullhomotopic\ lemma spheremap_lemma1: fixes f :: "'a::euclidean_space \ 'a::euclidean_space" assumes "subspace S" "subspace T" and dimST: "dim S < dim T" and "S \ T" and diff_f: "f differentiable_on sphere 0 1 \ S" shows "f ` (sphere 0 1 \ S) \ sphere 0 1 \ T" proof assume fim: "f ` (sphere 0 1 \ S) = sphere 0 1 \ T" have inS: "\x. \x \ S; x \ 0\ \ (x /\<^sub>R norm x) \ S" using subspace_mul \subspace S\ by blast have subS01: "(\x. x /\<^sub>R norm x) ` (S - {0}) \ sphere 0 1 \ S" using \subspace S\ subspace_mul by fastforce then have diff_f': "f differentiable_on (\x. x /\<^sub>R norm x) ` (S - {0})" by (rule differentiable_on_subset [OF diff_f]) define g where "g \ \x. norm x *\<^sub>R f(inverse(norm x) *\<^sub>R x)" have gdiff: "g differentiable_on S - {0}" unfolding g_def by (rule diff_f' derivative_intros differentiable_on_compose [where f=f] | force)+ have geq: "g ` (S - {0}) = T - {0}" proof have "\u. \u \ S; norm u *\<^sub>R f (u /\<^sub>R norm u) \ T\ \ u = 0" by (metis (mono_tags, lifting) DiffI subS01 subspace_mul [OF \subspace T\] fim image_subset_iff inf_le2 singletonD) then have "g ` (S - {0}) \ T" using g_def by blast moreover have "g ` (S - {0}) \ UNIV - {0}" proof (clarsimp simp: g_def) fix y assume "y \ S" and f0: "f (y /\<^sub>R norm y) = 0" then have "y \ 0 \ y /\<^sub>R norm y \ sphere 0 1 \ S" by (auto simp: subspace_mul [OF \subspace S\]) then show "y = 0" by (metis fim f0 Int_iff image_iff mem_sphere_0 norm_eq_zero zero_neq_one) qed ultimately show "g ` (S - {0}) \ T - {0}" by auto next have *: "sphere 0 1 \ T \ f ` (sphere 0 1 \ S)" using fim by (simp add: image_subset_iff) have "x \ (\x. norm x *\<^sub>R f (x /\<^sub>R norm x)) ` (S - {0})" if "x \ T" "x \ 0" for x proof - have "x /\<^sub>R norm x \ T" using \subspace T\ subspace_mul that by blast then obtain u where u: "f u \ T" "x /\<^sub>R norm x = f u" "norm u = 1" "u \ S" using * [THEN subsetD, of "x /\<^sub>R norm x"] \x \ 0\ by auto with that have [simp]: "norm x *\<^sub>R f u = x" by (metis divideR_right norm_eq_zero) moreover have "norm x *\<^sub>R u \ S - {0}" using \subspace S\ subspace_scale that(2) u by auto with u show ?thesis by (simp add: image_eqI [where x="norm x *\<^sub>R u"]) qed then have "T - {0} \ (\x. norm x *\<^sub>R f (x /\<^sub>R norm x)) ` (S - {0})" by force then show "T - {0} \ g ` (S - {0})" by (simp add: g_def) qed define T' where "T' \ {y. \x \ T. orthogonal x y}" have "subspace T'" by (simp add: subspace_orthogonal_to_vectors T'_def) have dim_eq: "dim T' + dim T = DIM('a)" using dim_subspace_orthogonal_to_vectors [of T UNIV] \subspace T\ by (simp add: T'_def) have "\v1 v2. v1 \ span T \ (\w \ span T. orthogonal v2 w) \ x = v1 + v2" for x by (force intro: orthogonal_subspace_decomp_exists [of T x]) then obtain p1 p2 where p1span: "p1 x \ span T" and "\w. w \ span T \ orthogonal (p2 x) w" and eq: "p1 x + p2 x = x" for x by metis then have p1: "\z. p1 z \ T" and ortho: "\w. w \ T \ orthogonal (p2 x) w" for x using span_eq_iff \subspace T\ by blast+ then have p2: "\z. p2 z \ T'" by (simp add: T'_def orthogonal_commute) have p12_eq: "\x y. \x \ T; y \ T'\ \ p1(x + y) = x \ p2(x + y) = y" proof (rule orthogonal_subspace_decomp_unique [OF eq p1span, where T=T']) show "\x y. \x \ T; y \ T'\ \ p2 (x + y) \ span T'" using span_eq_iff p2 \subspace T'\ by blast show "\a b. \a \ T; b \ T'\ \ orthogonal a b" using T'_def by blast qed (auto simp: span_base) then have "\c x. p1 (c *\<^sub>R x) = c *\<^sub>R p1 x \ p2 (c *\<^sub>R x) = c *\<^sub>R p2 x" proof - fix c :: real and x :: 'a have f1: "c *\<^sub>R x = c *\<^sub>R p1 x + c *\<^sub>R p2 x" by (metis eq pth_6) have f2: "c *\<^sub>R p2 x \ T'" by (simp add: \subspace T'\ p2 subspace_scale) have "c *\<^sub>R p1 x \ T" by (metis (full_types) assms(2) p1span span_eq_iff subspace_scale) then show "p1 (c *\<^sub>R x) = c *\<^sub>R p1 x \ p2 (c *\<^sub>R x) = c *\<^sub>R p2 x" using f2 f1 p12_eq by presburger qed moreover have lin_add: "\x y. p1 (x + y) = p1 x + p1 y \ p2 (x + y) = p2 x + p2 y" proof (rule orthogonal_subspace_decomp_unique [OF _ p1span, where T=T']) show "\x y. p1 (x + y) + p2 (x + y) = p1 x + p1 y + (p2 x + p2 y)" by (simp add: add.assoc add.left_commute eq) show "\a b. \a \ T; b \ T'\ \ orthogonal a b" using T'_def by blast qed (auto simp: p1span p2 span_base span_add) ultimately have "linear p1" "linear p2" by unfold_locales auto have "g differentiable_on p1 ` {x + y |x y. x \ S - {0} \ y \ T'}" using p12_eq \S \ T\ by (force intro: differentiable_on_subset [OF gdiff]) then have "(\z. g (p1 z)) differentiable_on {x + y |x y. x \ S - {0} \ y \ T'}" by (rule differentiable_on_compose [OF linear_imp_differentiable_on [OF \linear p1\]]) then have diff: "(\x. g (p1 x) + p2 x) differentiable_on {x + y |x y. x \ S - {0} \ y \ T'}" by (intro derivative_intros linear_imp_differentiable_on [OF \linear p2\]) have "dim {x + y |x y. x \ S - {0} \ y \ T'} \ dim {x + y |x y. x \ S \ y \ T'}" by (blast intro: dim_subset) also have "... = dim S + dim T' - dim (S \ T')" using dim_sums_Int [OF \subspace S\ \subspace T'\] by (simp add: algebra_simps) also have "... < DIM('a)" using dimST dim_eq by auto finally have neg: "negligible {x + y |x y. x \ S - {0} \ y \ T'}" by (rule negligible_lowdim) have "negligible ((\x. g (p1 x) + p2 x) ` {x + y |x y. x \ S - {0} \ y \ T'})" by (rule negligible_differentiable_image_negligible [OF order_refl neg diff]) then have "negligible {x + y |x y. x \ g ` (S - {0}) \ y \ T'}" proof (rule negligible_subset) have "\t' \ T'; s \ S; s \ 0\ \ g s + t' \ (\x. g (p1 x) + p2 x) ` {x + t' |x t'. x \ S \ x \ 0 \ t' \ T'}" for t' s apply (rule_tac x="s + t'" in image_eqI) using \S \ T\ p12_eq by auto then show "{x + y |x y. x \ g ` (S - {0}) \ y \ T'} \ (\x. g (p1 x) + p2 x) ` {x + y |x y. x \ S - {0} \ y \ T'}" by auto qed moreover have "- T' \ {x + y |x y. x \ g ` (S - {0}) \ y \ T'}" proof clarsimp fix z assume "z \ T'" show "\x y. z = x + y \ x \ g ` (S - {0}) \ y \ T'" by (metis Diff_iff \z \ T'\ add.left_neutral eq geq p1 p2 singletonD) qed ultimately have "negligible (-T')" using negligible_subset by blast moreover have "negligible T'" using negligible_lowdim by (metis add.commute assms(3) diff_add_inverse2 diff_self_eq_0 dim_eq le_add1 le_antisym linordered_semidom_class.add_diff_inverse not_less0) ultimately have "negligible (-T' \ T')" by (metis negligible_Un_eq) then show False using negligible_Un_eq non_negligible_UNIV by simp qed lemma spheremap_lemma2: fixes f :: "'a::euclidean_space \ 'a::euclidean_space" assumes ST: "subspace S" "subspace T" "dim S < dim T" and "S \ T" and contf: "continuous_on (sphere 0 1 \ S) f" and fim: "f ` (sphere 0 1 \ S) \ sphere 0 1 \ T" shows "\c. homotopic_with_canon (\x. True) (sphere 0 1 \ S) (sphere 0 1 \ T) f (\x. c)" proof - have [simp]: "\x. \norm x = 1; x \ S\ \ norm (f x) = 1" using fim by (simp add: image_subset_iff) have "compact (sphere 0 1 \ S)" by (simp add: \subspace S\ closed_subspace compact_Int_closed) then obtain g where pfg: "polynomial_function g" and gim: "g ` (sphere 0 1 \ S) \ T" and g12: "\x. x \ sphere 0 1 \ S \ norm(f x - g x) < 1/2" apply (rule Stone_Weierstrass_polynomial_function_subspace [OF _ contf _ \subspace T\, of "1/2"]) using fim by auto have gnz: "g x \ 0" if "x \ sphere 0 1 \ S" for x proof - have "norm (f x) = 1" using fim that by (simp add: image_subset_iff) then show ?thesis using g12 [OF that] by auto qed have diffg: "g differentiable_on sphere 0 1 \ S" by (metis pfg differentiable_on_polynomial_function) define h where "h \ \x. inverse(norm(g x)) *\<^sub>R g x" have h: "x \ sphere 0 1 \ S \ h x \ sphere 0 1 \ T" for x unfolding h_def using gnz [of x] by (auto simp: subspace_mul [OF \subspace T\] subsetD [OF gim]) have diffh: "h differentiable_on sphere 0 1 \ S" unfolding h_def using gnz by (fastforce intro: derivative_intros diffg differentiable_on_compose [OF diffg]) have homfg: "homotopic_with_canon (\z. True) (sphere 0 1 \ S) (T - {0}) f g" proof (rule homotopic_with_linear [OF contf]) show "continuous_on (sphere 0 1 \ S) g" using pfg by (simp add: differentiable_imp_continuous_on diffg) next have non0fg: "0 \ closed_segment (f x) (g x)" if "norm x = 1" "x \ S" for x proof - have "f x \ sphere 0 1" using fim that by (simp add: image_subset_iff) moreover have "norm(f x - g x) < 1/2" using g12 that by auto ultimately show ?thesis by (auto simp: norm_minus_commute dest: segment_bound) qed show "closed_segment (f x) (g x) \ T - {0}" if "x \ sphere 0 1 \ S" for x proof - have "convex T" by (simp add: \subspace T\ subspace_imp_convex) then have "convex hull {f x, g x} \ T" by (metis IntD2 closed_segment_subset fim gim image_subset_iff segment_convex_hull that) then show ?thesis using that non0fg segment_convex_hull by fastforce qed qed obtain d where d: "d \ (sphere 0 1 \ T) - h ` (sphere 0 1 \ S)" using h spheremap_lemma1 [OF ST \S \ T\ diffh] by force then have non0hd: "0 \ closed_segment (h x) (- d)" if "norm x = 1" "x \ S" for x using midpoint_between [of 0 "h x" "-d"] that h [of x] by (auto simp: between_mem_segment midpoint_def) have conth: "continuous_on (sphere 0 1 \ S) h" using differentiable_imp_continuous_on diffh by blast have hom_hd: "homotopic_with_canon (\z. True) (sphere 0 1 \ S) (T - {0}) h (\x. -d)" proof (rule homotopic_with_linear [OF conth continuous_on_const]) fix x assume x: "x \ sphere 0 1 \ S" have "convex hull {h x, - d} \ T" proof (rule hull_minimal) show "{h x, - d} \ T" using h d x by (force simp: subspace_neg [OF \subspace T\]) qed (simp add: subspace_imp_convex [OF \subspace T\]) with x segment_convex_hull show "closed_segment (h x) (- d) \ T - {0}" by (auto simp add: subset_Diff_insert non0hd) qed have conT0: "continuous_on (T - {0}) (\y. inverse(norm y) *\<^sub>R y)" by (intro continuous_intros) auto have sub0T: "(\y. y /\<^sub>R norm y) ` (T - {0}) \ sphere 0 1 \ T" by (fastforce simp: assms(2) subspace_mul) obtain c where homhc: "homotopic_with_canon (\z. True) (sphere 0 1 \ S) (sphere 0 1 \ T) h (\x. c)" apply (rule_tac c="-d" in that) apply (rule homotopic_with_eq) apply (rule homotopic_with_compose_continuous_left [OF hom_hd conT0 sub0T]) using d apply (auto simp: h_def) done have "homotopic_with_canon (\x. True) (sphere 0 1 \ S) (sphere 0 1 \ T) f h" apply (rule homotopic_with_eq [OF homotopic_with_compose_continuous_left [OF homfg conT0 sub0T]]) by (auto simp: h_def) then show ?thesis by (metis homotopic_with_trans [OF _ homhc]) qed lemma spheremap_lemma3: assumes "bounded S" "convex S" "subspace U" and affSU: "aff_dim S \ dim U" obtains T where "subspace T" "T \ U" "S \ {} \ aff_dim T = aff_dim S" "(rel_frontier S) homeomorphic (sphere 0 1 \ T)" proof (cases "S = {}") case True with \subspace U\ subspace_0 show ?thesis by (rule_tac T = "{0}" in that) auto next case False then obtain a where "a \ S" by auto then have affS: "aff_dim S = int (dim ((\x. -a+x) ` S))" by (metis hull_inc aff_dim_eq_dim) with affSU have "dim ((\x. -a+x) ` S) \ dim U" by linarith with choose_subspace_of_subspace obtain T where "subspace T" "T \ span U" and dimT: "dim T = dim ((\x. -a+x) ` S)" . show ?thesis proof (rule that [OF \subspace T\]) show "T \ U" using span_eq_iff \subspace U\ \T \ span U\ by blast show "aff_dim T = aff_dim S" using dimT \subspace T\ affS aff_dim_subspace by fastforce show "rel_frontier S homeomorphic sphere 0 1 \ T" proof - have "aff_dim (ball 0 1 \ T) = aff_dim (T)" by (metis IntI interior_ball \subspace T\ aff_dim_convex_Int_nonempty_interior centre_in_ball empty_iff inf_commute subspace_0 subspace_imp_convex zero_less_one) then have affS_eq: "aff_dim S = aff_dim (ball 0 1 \ T)" using \aff_dim T = aff_dim S\ by simp have "rel_frontier S homeomorphic rel_frontier(ball 0 1 \ T)" proof (rule homeomorphic_rel_frontiers_convex_bounded_sets [OF \convex S\ \bounded S\]) show "convex (ball 0 1 \ T)" by (simp add: \subspace T\ convex_Int subspace_imp_convex) show "bounded (ball 0 1 \ T)" by (simp add: bounded_Int) show "aff_dim S = aff_dim (ball 0 1 \ T)" by (rule affS_eq) qed also have "... = frontier (ball 0 1) \ T" proof (rule convex_affine_rel_frontier_Int [OF convex_ball]) show "affine T" by (simp add: \subspace T\ subspace_imp_affine) show "interior (ball 0 1) \ T \ {}" using \subspace T\ subspace_0 by force qed also have "... = sphere 0 1 \ T" by auto finally show ?thesis . qed qed qed proposition inessential_spheremap_lowdim_gen: fixes f :: "'M::euclidean_space \ 'a::euclidean_space" assumes "convex S" "bounded S" "convex T" "bounded T" and affST: "aff_dim S < aff_dim T" and contf: "continuous_on (rel_frontier S) f" and fim: "f ` (rel_frontier S) \ rel_frontier T" obtains c where "homotopic_with_canon (\z. True) (rel_frontier S) (rel_frontier T) f (\x. c)" proof (cases "S = {}") case True then show ?thesis by (simp add: that) next case False then show ?thesis proof (cases "T = {}") case True then show ?thesis using fim that by auto next case False obtain T':: "'a set" where "subspace T'" and affT': "aff_dim T' = aff_dim T" and homT: "rel_frontier T homeomorphic sphere 0 1 \ T'" apply (rule spheremap_lemma3 [OF \bounded T\ \convex T\ subspace_UNIV, where 'b='a]) apply (simp add: aff_dim_le_DIM) using \T \ {}\ by blast with homeomorphic_imp_homotopy_eqv have relT: "sphere 0 1 \ T' homotopy_eqv rel_frontier T" using homotopy_equivalent_space_sym by blast have "aff_dim S \ int (dim T')" using affT' \subspace T'\ affST aff_dim_subspace by force with spheremap_lemma3 [OF \bounded S\ \convex S\ \subspace T'\] \S \ {}\ obtain S':: "'a set" where "subspace S'" "S' \ T'" and affS': "aff_dim S' = aff_dim S" and homT: "rel_frontier S homeomorphic sphere 0 1 \ S'" by metis with homeomorphic_imp_homotopy_eqv have relS: "sphere 0 1 \ S' homotopy_eqv rel_frontier S" using homotopy_equivalent_space_sym by blast have dimST': "dim S' < dim T'" by (metis \S' \ T'\ \subspace S'\ \subspace T'\ affS' affST affT' less_irrefl not_le subspace_dim_equal) have "\c. homotopic_with_canon (\z. True) (rel_frontier S) (rel_frontier T) f (\x. c)" apply (rule homotopy_eqv_homotopic_triviality_null_imp [OF relT contf fim]) apply (rule homotopy_eqv_cohomotopic_triviality_null[OF relS, THEN iffD1, rule_format]) apply (metis dimST' \subspace S'\ \subspace T'\ \S' \ T'\ spheremap_lemma2, blast) done with that show ?thesis by blast qed qed lemma inessential_spheremap_lowdim: fixes f :: "'M::euclidean_space \ 'a::euclidean_space" assumes "DIM('M) < DIM('a)" and f: "continuous_on (sphere a r) f" "f ` (sphere a r) \ (sphere b s)" obtains c where "homotopic_with_canon (\z. True) (sphere a r) (sphere b s) f (\x. c)" proof (cases "s \ 0") case True then show ?thesis by (meson nullhomotopic_into_contractible f contractible_sphere that) next case False show ?thesis proof (cases "r \ 0") case True then show ?thesis by (meson f nullhomotopic_from_contractible contractible_sphere that) next case False with \\ s \ 0\ have "r > 0" "s > 0" by auto show thesis apply (rule inessential_spheremap_lowdim_gen [of "cball a r" "cball b s" f]) using \0 < r\ \0 < s\ assms(1) that by (simp_all add: f aff_dim_cball inessential_spheremap_lowdim_gen [of "cball a r" "cball b s" f]) qed qed subsection\ Some technical lemmas about extending maps from cell complexes\ lemma extending_maps_Union_aux: assumes fin: "finite \" and "\S. S \ \ \ closed S" and "\S T. \S \ \; T \ \; S \ T\ \ S \ T \ K" and "\S. S \ \ \ \g. continuous_on S g \ g ` S \ T \ (\x \ S \ K. g x = h x)" shows "\g. continuous_on (\\) g \ g ` (\\) \ T \ (\x \ \\ \ K. g x = h x)" using assms proof (induction \) case empty show ?case by simp next case (insert S \) then obtain f where contf: "continuous_on (S) f" and fim: "f ` S \ T" and feq: "\x \ S \ K. f x = h x" by (meson insertI1) obtain g where contg: "continuous_on (\\) g" and gim: "g ` \\ \ T" and geq: "\x \ \\ \ K. g x = h x" using insert by auto have fg: "f x = g x" if "x \ T" "T \ \" "x \ S" for x T proof - have "T \ S \ K \ S = T" using that by (metis (no_types) insert.prems(2) insertCI) then show ?thesis using UnionI feq geq \S \ \\ subsetD that by fastforce qed show ?case apply (rule_tac x="\x. if x \ S then f x else g x" in exI, simp) apply (intro conjI continuous_on_cases) using fim gim feq geq apply (force simp: insert closed_Union contf contg inf_commute intro: fg)+ done qed lemma extending_maps_Union: assumes fin: "finite \" and "\S. S \ \ \ \g. continuous_on S g \ g ` S \ T \ (\x \ S \ K. g x = h x)" and "\S. S \ \ \ closed S" and K: "\X Y. \X \ \; Y \ \; \ X \ Y; \ Y \ X\ \ X \ Y \ K" shows "\g. continuous_on (\\) g \ g ` (\\) \ T \ (\x \ \\ \ K. g x = h x)" apply (simp add: Union_maximal_sets [OF fin, symmetric]) apply (rule extending_maps_Union_aux) apply (simp_all add: Union_maximal_sets [OF fin] assms) by (metis K psubsetI) lemma extend_map_lemma: assumes "finite \" "\ \ \" "convex T" "bounded T" and poly: "\X. X \ \ \ polytope X" and aff: "\X. X \ \ - \ \ aff_dim X < aff_dim T" and face: "\S T. \S \ \; T \ \\ \ (S \ T) face_of S" and contf: "continuous_on (\\) f" and fim: "f ` (\\) \ rel_frontier T" obtains g where "continuous_on (\\) g" "g ` (\\) \ rel_frontier T" "\x. x \ \\ \ g x = f x" proof (cases "\ - \ = {}") case True show ?thesis proof show "continuous_on (\ \) f" using True \\ \ \\ contf by auto show "f ` \ \ \ rel_frontier T" using True fim by auto qed auto next case False then have "0 \ aff_dim T" by (metis aff aff_dim_empty aff_dim_geq aff_dim_negative_iff all_not_in_conv not_less) then obtain i::nat where i: "int i = aff_dim T" by (metis nonneg_eq_int) have Union_empty_eq: "\{D. D = {} \ P D} = {}" for P :: "'a set \ bool" by auto have face': "\S T. \S \ \; T \ \\ \ (S \ T) face_of S \ (S \ T) face_of T" by (metis face inf_commute) have extendf: "\g. continuous_on (\(\ \ {D. \C \ \. D face_of C \ aff_dim D < i})) g \ g ` (\ (\ \ {D. \C \ \. D face_of C \ aff_dim D < i})) \ rel_frontier T \ (\x \ \\. g x = f x)" if "i \ aff_dim T" for i::nat using that proof (induction i) case 0 show ?case using 0 contf fim by (auto simp add: Union_empty_eq) next case (Suc p) with \bounded T\ have "rel_frontier T \ {}" by (auto simp: rel_frontier_eq_empty affine_bounded_eq_lowdim [of T]) then obtain t where t: "t \ rel_frontier T" by auto have ple: "int p \ aff_dim T" using Suc.prems by force obtain h where conth: "continuous_on (\(\ \ {D. \C \ \. D face_of C \ aff_dim D < p})) h" and him: "h ` (\ (\ \ {D. \C \ \. D face_of C \ aff_dim D < p})) \ rel_frontier T" and heq: "\x. x \ \\ \ h x = f x" using Suc.IH [OF ple] by auto let ?Faces = "{D. \C \ \. D face_of C \ aff_dim D \ p}" have extendh: "\g. continuous_on D g \ g ` D \ rel_frontier T \ (\x \ D \ \(\ \ {D. \C \ \. D face_of C \ aff_dim D < p}). g x = h x)" if D: "D \ \ \ ?Faces" for D proof (cases "D \ \(\ \ {D. \C \ \. D face_of C \ aff_dim D < p})") case True have "continuous_on D h" using True conth continuous_on_subset by blast moreover have "h ` D \ rel_frontier T" using True him by blast ultimately show ?thesis by blast next case False note notDsub = False show ?thesis proof (cases "\a. D = {a}") case True then obtain a where "D = {a}" by auto with notDsub t show ?thesis by (rule_tac x="\x. t" in exI) simp next case False have "D \ {}" using notDsub by auto have Dnotin: "D \ \ \ {D. \C \ \. D face_of C \ aff_dim D < p}" using notDsub by auto then have "D \ \" by simp have "D \ ?Faces - {D. \C \ \. D face_of C \ aff_dim D < p}" using Dnotin that by auto then obtain C where "C \ \" "D face_of C" and affD: "aff_dim D = int p" by auto then have "bounded D" using face_of_polytope_polytope poly polytope_imp_bounded by blast then have [simp]: "\ affine D" using affine_bounded_eq_trivial False \D \ {}\ \bounded D\ by blast have "{F. F facet_of D} \ {E. E face_of C \ aff_dim E < int p}" by clarify (metis \D face_of C\ affD eq_iff face_of_trans facet_of_def zle_diff1_eq) moreover have "polyhedron D" using \C \ \\ \D face_of C\ face_of_polytope_polytope poly polytope_imp_polyhedron by auto ultimately have relf_sub: "rel_frontier D \ \ {E. E face_of C \ aff_dim E < p}" by (simp add: rel_frontier_of_polyhedron Union_mono) then have him_relf: "h ` rel_frontier D \ rel_frontier T" using \C \ \\ him by blast have "convex D" by (simp add: \polyhedron D\ polyhedron_imp_convex) have affD_lessT: "aff_dim D < aff_dim T" using Suc.prems affD by linarith have contDh: "continuous_on (rel_frontier D) h" using \C \ \\ relf_sub by (blast intro: continuous_on_subset [OF conth]) then have *: "(\c. homotopic_with_canon (\x. True) (rel_frontier D) (rel_frontier T) h (\x. c)) = (\g. continuous_on UNIV g \ range g \ rel_frontier T \ (\x\rel_frontier D. g x = h x))" apply (rule nullhomotopic_into_rel_frontier_extension [OF closed_rel_frontier]) apply (simp_all add: assms rel_frontier_eq_empty him_relf) done have "(\c. homotopic_with_canon (\x. True) (rel_frontier D) (rel_frontier T) h (\x. c))" by (metis inessential_spheremap_lowdim_gen [OF \convex D\ \bounded D\ \convex T\ \bounded T\ affD_lessT contDh him_relf]) then obtain g where contg: "continuous_on UNIV g" and gim: "range g \ rel_frontier T" and gh: "\x. x \ rel_frontier D \ g x = h x" by (metis *) have "D \ E \ rel_frontier D" if "E \ \ \ {D. Bex \ ((face_of) D) \ aff_dim D < int p}" for E proof (rule face_of_subset_rel_frontier) show "D \ E face_of D" using that proof safe assume "E \ \" then show "D \ E face_of D" by (meson \C \ \\ \D face_of C\ assms(2) face' face_of_Int_subface face_of_refl_eq poly polytope_imp_convex subsetD) next fix x assume "aff_dim E < int p" "x \ \" "E face_of x" then show "D \ E face_of D" by (meson \C \ \\ \D face_of C\ face' face_of_Int_subface that) qed show "D \ E \ D" using that notDsub by auto qed moreover have "continuous_on D g" using contg continuous_on_subset by blast ultimately show ?thesis by (rule_tac x=g in exI) (use gh gim in fastforce) qed qed have intle: "i < 1 + int j \ i \ int j" for i j by auto have "finite \" using \finite \\ \\ \ \\ rev_finite_subset by blast moreover have "finite (?Faces)" proof - have \
: "finite (\ {{D. D face_of C} |C. C \ \})" by (auto simp: \finite \\ finite_polytope_faces poly) show ?thesis by (auto intro: finite_subset [OF _ \
]) qed ultimately have fin: "finite (\ \ ?Faces)" by simp have clo: "closed S" if "S \ \ \ ?Faces" for S using that \\ \ \\ face_of_polytope_polytope poly polytope_imp_closed by blast have K: "X \ Y \ \(\ \ {D. \C\\. D face_of C \ aff_dim D < int p})" if "X \ \ \ ?Faces" "Y \ \ \ ?Faces" "\ Y \ X" for X Y proof - have ff: "X \ Y face_of X \ X \ Y face_of Y" if XY: "X face_of D" "Y face_of E" and DE: "D \ \" "E \ \" for D E by (rule face_of_Int_subface [OF _ _ XY]) (auto simp: face' DE) show ?thesis using that apply auto apply (drule_tac x="X \ Y" in spec, safe) using ff face_of_imp_convex [of X] face_of_imp_convex [of Y] apply (fastforce dest: face_of_aff_dim_lt) by (meson face_of_trans ff) qed obtain g where "continuous_on (\(\ \ ?Faces)) g" "g ` \(\ \ ?Faces) \ rel_frontier T" "(\x \ \(\ \ ?Faces) \ \(\ \ {D. \C\\. D face_of C \ aff_dim D < p}). g x = h x)" by (rule exE [OF extending_maps_Union [OF fin extendh clo K]], blast+) then show ?case by (simp add: intle local.heq [symmetric], blast) qed have eq: "\(\ \ {D. \C \ \. D face_of C \ aff_dim D < i}) = \\" proof show "\(\ \ {D. \C\\. D face_of C \ aff_dim D < int i}) \ \\" using \\ \ \\ face_of_imp_subset by fastforce show "\\ \ \(\ \ {D. \C\\. D face_of C \ aff_dim D < i})" proof (rule Union_mono) show "\ \ \ \ {D. \C\\. D face_of C \ aff_dim D < int i}" using face by (fastforce simp: aff i) qed qed have "int i \ aff_dim T" by (simp add: i) then show ?thesis using extendf [of i] unfolding eq by (metis that) qed lemma extend_map_lemma_cofinite0: assumes "finite \" and "pairwise (\S T. S \ T \ K) \" and "\S. S \ \ \ \a g. a \ U \ continuous_on (S - {a}) g \ g ` (S - {a}) \ T \ (\x \ S \ K. g x = h x)" and "\S. S \ \ \ closed S" shows "\C g. finite C \ disjnt C U \ card C \ card \ \ continuous_on (\\ - C) g \ g ` (\\ - C) \ T \ (\x \ (\\ - C) \ K. g x = h x)" using assms proof induction case empty then show ?case by force next case (insert X \) then have "closed X" and clo: "\X. X \ \ \ closed X" and \: "\S. S \ \ \ \a g. a \ U \ continuous_on (S - {a}) g \ g ` (S - {a}) \ T \ (\x \ S \ K. g x = h x)" and pwX: "\Y. Y \ \ \ Y \ X \ X \ Y \ K \ Y \ X \ K" and pwF: "pairwise (\ S T. S \ T \ K) \" by (simp_all add: pairwise_insert) obtain C g where C: "finite C" "disjnt C U" "card C \ card \" and contg: "continuous_on (\\ - C) g" and gim: "g ` (\\ - C) \ T" and gh: "\x. x \ (\\ - C) \ K \ g x = h x" using insert.IH [OF pwF \ clo] by auto obtain a f where "a \ U" and contf: "continuous_on (X - {a}) f" and fim: "f ` (X - {a}) \ T" and fh: "(\x \ X \ K. f x = h x)" using insert.prems by (meson insertI1) show ?case proof (intro exI conjI) show "finite (insert a C)" by (simp add: C) show "disjnt (insert a C) U" using C \a \ U\ by simp show "card (insert a C) \ card (insert X \)" by (simp add: C card_insert_if insert.hyps le_SucI) have "closed (\\)" using clo insert.hyps by blast have "continuous_on (X - insert a C) f" using contf by (force simp: elim: continuous_on_subset) moreover have "continuous_on (\ \ - insert a C) g" using contg by (force simp: elim: continuous_on_subset) ultimately have "continuous_on (X - insert a C \ (\\ - insert a C)) (\x. if x \ X then f x else g x)" apply (intro continuous_on_cases_local; simp add: closedin_closed) using \closed X\ apply blast using \closed (\\)\ apply blast using fh gh insert.hyps pwX by fastforce then show "continuous_on (\(insert X \) - insert a C) (\a. if a \ X then f a else g a)" by (blast intro: continuous_on_subset) show "\x\(\(insert X \) - insert a C) \ K. (if x \ X then f x else g x) = h x" using gh by (auto simp: fh) show "(\a. if a \ X then f a else g a) ` (\(insert X \) - insert a C) \ T" using fim gim by auto force qed qed lemma extend_map_lemma_cofinite1: assumes "finite \" and \: "\X. X \ \ \ \a g. a \ U \ continuous_on (X - {a}) g \ g ` (X - {a}) \ T \ (\x \ X \ K. g x = h x)" and clo: "\X. X \ \ \ closed X" and K: "\X Y. \X \ \; Y \ \; \ X \ Y; \ Y \ X\ \ X \ Y \ K" obtains C g where "finite C" "disjnt C U" "card C \ card \" "continuous_on (\\ - C) g" "g ` (\\ - C) \ T" "\x. x \ (\\ - C) \ K \ g x = h x" proof - let ?\ = "{X \ \. \Y\\. \ X \ Y}" have [simp]: "\?\ = \\" by (simp add: Union_maximal_sets assms) have fin: "finite ?\" by (force intro: finite_subset [OF _ \finite \\]) have pw: "pairwise (\ S T. S \ T \ K) ?\" by (simp add: pairwise_def) (metis K psubsetI) have "card {X \ \. \Y\\. \ X \ Y} \ card \" by (simp add: \finite \\ card_mono) moreover obtain C g where "finite C \ disjnt C U \ card C \ card ?\ \ continuous_on (\?\ - C) g \ g ` (\?\ - C) \ T \ (\x \ (\?\ - C) \ K. g x = h x)" using extend_map_lemma_cofinite0 [OF fin pw, of U T h] by (fastforce intro!: clo \) ultimately show ?thesis by (rule_tac C=C and g=g in that) auto qed lemma extend_map_lemma_cofinite: assumes "finite \" "\ \ \" and T: "convex T" "bounded T" and poly: "\X. X \ \ \ polytope X" and contf: "continuous_on (\\) f" and fim: "f ` (\\) \ rel_frontier T" and face: "\X Y. \X \ \; Y \ \\ \ (X \ Y) face_of X" and aff: "\X. X \ \ - \ \ aff_dim X \ aff_dim T" obtains C g where "finite C" "disjnt C (\\)" "card C \ card \" "continuous_on (\\ - C) g" "g ` (\ \ - C) \ rel_frontier T" "\x. x \ \\ \ g x = f x" proof - define \ where "\ \ \ \ {D. \C \ \ - \. D face_of C \ aff_dim D < aff_dim T}" have "finite \" using assms finite_subset by blast have *: "finite (\{{D. D face_of C} |C. C \ \})" using finite_polytope_faces poly \finite \\ by force then have "finite \" by (auto simp: \_def \finite \\ intro: finite_subset [OF _ *]) have face': "\S T. \S \ \; T \ \\ \ (S \ T) face_of S \ (S \ T) face_of T" by (metis face inf_commute) have *: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X" unfolding \_def using subsetD [OF \\ \ \\] apply (auto simp add: face) apply (meson face' face_of_Int_subface face_of_refl_eq poly polytope_imp_convex)+ done obtain h where conth: "continuous_on (\\) h" and him: "h ` (\\) \ rel_frontier T" and hf: "\x. x \ \\ \ h x = f x" proof (rule extend_map_lemma [OF \finite \\ [unfolded \_def] Un_upper1 T]) show "\X. \X \ \ \ {D. \C\\ - \. D face_of C \ aff_dim D < aff_dim T}\ \ polytope X" using \\ \ \\ face_of_polytope_polytope poly by fastforce qed (use * \_def contf fim in auto) have "bounded (\\)" using \finite \\ \\ \ \\ poly polytope_imp_bounded by blast then have "\\ \ UNIV" by auto then obtain a where a: "a \ \\" by blast have \: "\a g. a \ \\ \ continuous_on (D - {a}) g \ g ` (D - {a}) \ rel_frontier T \ (\x \ D \ \\. g x = h x)" if "D \ \" for D proof (cases "D \ \\") case True then show ?thesis apply (rule_tac x=a in exI) apply (rule_tac x=h in exI) using him apply (blast intro!: \a \ \\\ continuous_on_subset [OF conth]) + done next case False note D_not_subset = False show ?thesis proof (cases "D \ \") case True with D_not_subset show ?thesis by (auto simp: \_def) next case False then have affD: "aff_dim D \ aff_dim T" by (simp add: \D \ \\ aff) show ?thesis proof (cases "rel_interior D = {}") case True with \D \ \\ poly a show ?thesis by (force simp: rel_interior_eq_empty polytope_imp_convex) next case False then obtain b where brelD: "b \ rel_interior D" by blast have "polyhedron D" by (simp add: poly polytope_imp_polyhedron that) have "rel_frontier D retract_of affine hull D - {b}" by (simp add: rel_frontier_retract_of_punctured_affine_hull poly polytope_imp_bounded polytope_imp_convex that brelD) then obtain r where relfD: "rel_frontier D \ affine hull D - {b}" and contr: "continuous_on (affine hull D - {b}) r" and rim: "r ` (affine hull D - {b}) \ rel_frontier D" and rid: "\x. x \ rel_frontier D \ r x = x" by (auto simp: retract_of_def retraction_def) show ?thesis proof (intro exI conjI ballI) show "b \ \\" proof clarify fix E assume "b \ E" "E \ \" then have "E \ D face_of E \ E \ D face_of D" using \\ \ \\ face' that by auto with face_of_subset_rel_frontier \E \ \\ \b \ E\ brelD rel_interior_subset [of D] D_not_subset rel_frontier_def \_def show False by blast qed have "r ` (D - {b}) \ r ` (affine hull D - {b})" by (simp add: Diff_mono hull_subset image_mono) also have "... \ rel_frontier D" by (rule rim) also have "... \ \{E. E face_of D \ aff_dim E < aff_dim T}" using affD by (force simp: rel_frontier_of_polyhedron [OF \polyhedron D\] facet_of_def) also have "... \ \(\)" using D_not_subset \_def that by fastforce finally have rsub: "r ` (D - {b}) \ \(\)" . show "continuous_on (D - {b}) (h \ r)" apply (intro conjI \b \ \\\ continuous_on_compose) apply (rule continuous_on_subset [OF contr]) apply (simp add: Diff_mono hull_subset) apply (rule continuous_on_subset [OF conth rsub]) done show "(h \ r) ` (D - {b}) \ rel_frontier T" using brelD him rsub by fastforce show "(h \ r) x = h x" if x: "x \ D \ \\" for x proof - consider A where "x \ D" "A \ \" "x \ A" | A B where "x \ D" "A face_of B" "B \ \" "B \ \" "aff_dim A < aff_dim T" "x \ A" using x by (auto simp: \_def) then have xrel: "x \ rel_frontier D" proof cases case 1 show ?thesis proof (rule face_of_subset_rel_frontier [THEN subsetD]) show "D \ A face_of D" using \A \ \\ \\ \ \\ face \D \ \\ by blast show "D \ A \ D" using \A \ \\ D_not_subset \_def by blast qed (auto simp: 1) next case 2 show ?thesis proof (rule face_of_subset_rel_frontier [THEN subsetD]) have "D face_of D" by (simp add: \polyhedron D\ polyhedron_imp_convex face_of_refl) then show "D \ A face_of D" by (meson "2"(2) "2"(3) \D \ \\ face' face_of_Int_Int face_of_face) show "D \ A \ D" using "2" D_not_subset \_def by blast qed (auto simp: 2) qed show ?thesis by (simp add: rid xrel) qed qed qed qed qed have clo: "\S. S \ \ \ closed S" by (simp add: poly polytope_imp_closed) obtain C g where "finite C" "disjnt C (\\)" "card C \ card \" "continuous_on (\\ - C) g" "g ` (\\ - C) \ rel_frontier T" and gh: "\x. x \ (\\ - C) \ \\ \ g x = h x" proof (rule extend_map_lemma_cofinite1 [OF \finite \\ \ clo]) show "X \ Y \ \\" if XY: "X \ \" "Y \ \" and "\ X \ Y" "\ Y \ X" for X Y proof (cases "X \ \") case True then show ?thesis by (auto simp: \_def) next case False have "X \ Y \ X" using \\ X \ Y\ by blast with XY show ?thesis by (clarsimp simp: \_def) (metis Diff_iff Int_iff aff antisym_conv face face_of_aff_dim_lt face_of_refl not_le poly polytope_imp_convex) qed qed (blast)+ with \\ \ \\ show ?thesis apply (rule_tac C=C and g=g in that) apply (auto simp: disjnt_def hf [symmetric] \_def intro!: gh) done qed text\The next two proofs are similar\ theorem extend_map_cell_complex_to_sphere: assumes "finite \" and S: "S \ \\" "closed S" and T: "convex T" "bounded T" and poly: "\X. X \ \ \ polytope X" and aff: "\X. X \ \ \ aff_dim X < aff_dim T" and face: "\X Y. \X \ \; Y \ \\ \ (X \ Y) face_of X" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier T" obtains g where "continuous_on (\\) g" "g ` (\\) \ rel_frontier T" "\x. x \ S \ g x = f x" proof - obtain V g where "S \ V" "open V" "continuous_on V g" and gim: "g ` V \ rel_frontier T" and gf: "\x. x \ S \ g x = f x" using neighbourhood_extension_into_ANR [OF contf fim _ \closed S\] ANR_rel_frontier_convex T by blast have "compact S" by (meson assms compact_Union poly polytope_imp_compact seq_compact_closed_subset seq_compact_eq_compact) then obtain d where "d > 0" and d: "\x y. \x \ S; y \ - V\ \ d \ dist x y" using separate_compact_closed [of S "-V"] \open V\ \S \ V\ by force obtain \ where "finite \" "\\ = \\" and diaG: "\X. X \ \ \ diameter X < d" and polyG: "\X. X \ \ \ polytope X" and affG: "\X. X \ \ \ aff_dim X \ aff_dim T - 1" and faceG: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X" proof (rule cell_complex_subdivision_exists [OF \d>0\ \finite \\ poly _ face]) show "\X. X \ \ \ aff_dim X \ aff_dim T - 1" by (simp add: aff) qed auto obtain h where conth: "continuous_on (\\) h" and him: "h ` \\ \ rel_frontier T" and hg: "\x. x \ \(\ \ Pow V) \ h x = g x" proof (rule extend_map_lemma [of \ "\ \ Pow V" T g]) show "continuous_on (\(\ \ Pow V)) g" by (metis Union_Int_subset Union_Pow_eq \continuous_on V g\ continuous_on_subset le_inf_iff) qed (use \finite \\ T polyG affG faceG gim in fastforce)+ show ?thesis proof show "continuous_on (\\) h" using \\\ = \\\ conth by auto show "h ` \\ \ rel_frontier T" using \\\ = \\\ him by auto show "h x = f x" if "x \ S" for x proof - have "x \ \\" using \\\ = \\\ \S \ \\\ that by auto then obtain X where "x \ X" "X \ \" by blast then have "diameter X < d" "bounded X" by (auto simp: diaG \X \ \\ polyG polytope_imp_bounded) then have "X \ V" using d [OF \x \ S\] diameter_bounded_bound [OF \bounded X\ \x \ X\] by fastforce have "h x = g x" apply (rule hg) using \X \ \\ \X \ V\ \x \ X\ by blast also have "... = f x" by (simp add: gf that) finally show "h x = f x" . qed qed qed theorem extend_map_cell_complex_to_sphere_cofinite: assumes "finite \" and S: "S \ \\" "closed S" and T: "convex T" "bounded T" and poly: "\X. X \ \ \ polytope X" and aff: "\X. X \ \ \ aff_dim X \ aff_dim T" and face: "\X Y. \X \ \; Y \ \\ \ (X \ Y) face_of X" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier T" obtains C g where "finite C" "disjnt C S" "continuous_on (\\ - C) g" "g ` (\\ - C) \ rel_frontier T" "\x. x \ S \ g x = f x" proof - obtain V g where "S \ V" "open V" "continuous_on V g" and gim: "g ` V \ rel_frontier T" and gf: "\x. x \ S \ g x = f x" using neighbourhood_extension_into_ANR [OF contf fim _ \closed S\] ANR_rel_frontier_convex T by blast have "compact S" by (meson assms compact_Union poly polytope_imp_compact seq_compact_closed_subset seq_compact_eq_compact) then obtain d where "d > 0" and d: "\x y. \x \ S; y \ - V\ \ d \ dist x y" using separate_compact_closed [of S "-V"] \open V\ \S \ V\ by force obtain \ where "finite \" "\\ = \\" and diaG: "\X. X \ \ \ diameter X < d" and polyG: "\X. X \ \ \ polytope X" and affG: "\X. X \ \ \ aff_dim X \ aff_dim T" and faceG: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X" by (rule cell_complex_subdivision_exists [OF \d>0\ \finite \\ poly aff face]) auto obtain C h where "finite C" and dis: "disjnt C (\(\ \ Pow V))" and card: "card C \ card \" and conth: "continuous_on (\\ - C) h" and him: "h ` (\\ - C) \ rel_frontier T" and hg: "\x. x \ \(\ \ Pow V) \ h x = g x" proof (rule extend_map_lemma_cofinite [of \ "\ \ Pow V" T g]) show "continuous_on (\(\ \ Pow V)) g" by (metis Union_Int_subset Union_Pow_eq \continuous_on V g\ continuous_on_subset le_inf_iff) show "g ` \(\ \ Pow V) \ rel_frontier T" using gim by force qed (auto intro: \finite \\ T polyG affG dest: faceG) have Ssub: "S \ \(\ \ Pow V)" proof fix x assume "x \ S" then have "x \ \\" using \\\ = \\\ \S \ \\\ by auto then obtain X where "x \ X" "X \ \" by blast then have "diameter X < d" "bounded X" by (auto simp: diaG \X \ \\ polyG polytope_imp_bounded) then have "X \ V" using d [OF \x \ S\] diameter_bounded_bound [OF \bounded X\ \x \ X\] by fastforce then show "x \ \(\ \ Pow V)" using \X \ \\ \x \ X\ by blast qed show ?thesis proof show "continuous_on (\\-C) h" using \\\ = \\\ conth by auto show "h ` (\\ - C) \ rel_frontier T" using \\\ = \\\ him by auto show "h x = f x" if "x \ S" for x proof - have "h x = g x" using Ssub hg that by blast also have "... = f x" by (simp add: gf that) finally show "h x = f x" . qed show "disjnt C S" using dis Ssub by (meson disjnt_iff subset_eq) qed (intro \finite C\) qed subsection\ Special cases and corollaries involving spheres\ lemma disjnt_Diff1: "X \ Y' \ disjnt (X - Y) (X' - Y')" by (auto simp: disjnt_def) proposition extend_map_affine_to_sphere_cofinite_simple: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact S" "convex U" "bounded U" and aff: "aff_dim T \ aff_dim U" and "S \ T" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier U" obtains K g where "finite K" "K \ T" "disjnt K S" "continuous_on (T - K) g" "g ` (T - K) \ rel_frontier U" "\x. x \ S \ g x = f x" proof - have "\K g. finite K \ disjnt K S \ continuous_on (T - K) g \ g ` (T - K) \ rel_frontier U \ (\x \ S. g x = f x)" if "affine T" "S \ T" and aff: "aff_dim T \ aff_dim U" for T proof (cases "S = {}") case True show ?thesis proof (cases "rel_frontier U = {}") case True with \bounded U\ have "aff_dim U \ 0" using affine_bounded_eq_lowdim rel_frontier_eq_empty by auto with aff have "aff_dim T \ 0" by auto then obtain a where "T \ {a}" using \affine T\ affine_bounded_eq_lowdim affine_bounded_eq_trivial by auto then show ?thesis using \S = {}\ fim by (metis Diff_cancel contf disjnt_empty2 finite.emptyI finite_insert finite_subset) next case False then obtain a where "a \ rel_frontier U" by auto then show ?thesis using continuous_on_const [of _ a] \S = {}\ by force qed next case False have "bounded S" by (simp add: \compact S\ compact_imp_bounded) then obtain b where b: "S \ cbox (-b) b" using bounded_subset_cbox_symmetric by blast define bbox where "bbox \ cbox (-(b+One)) (b+One)" have "cbox (-b) b \ bbox" by (auto simp: bbox_def algebra_simps intro!: subset_box_imp) with b \S \ T\ have "S \ bbox \ T" by auto then have Ssub: "S \ \{bbox \ T}" by auto then have "aff_dim (bbox \ T) \ aff_dim U" by (metis aff aff_dim_subset inf_commute inf_le1 order_trans) obtain K g where K: "finite K" "disjnt K S" and contg: "continuous_on (\{bbox \ T} - K) g" and gim: "g ` (\{bbox \ T} - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" proof (rule extend_map_cell_complex_to_sphere_cofinite [OF _ Ssub _ \convex U\ \bounded U\ _ _ _ contf fim]) show "closed S" using \compact S\ compact_eq_bounded_closed by auto show poly: "\X. X \ {bbox \ T} \ polytope X" by (simp add: polytope_Int_polyhedron bbox_def polytope_interval affine_imp_polyhedron \affine T\) show "\X Y. \X \ {bbox \ T}; Y \ {bbox \ T}\ \ X \ Y face_of X" by (simp add:poly face_of_refl polytope_imp_convex) show "\X. X \ {bbox \ T} \ aff_dim X \ aff_dim U" by (simp add: \aff_dim (bbox \ T) \ aff_dim U\) qed auto define fro where "fro \ \d. frontier(cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One))" obtain d where d12: "1/2 \ d" "d \ 1" and dd: "disjnt K (fro d)" proof (rule disjoint_family_elem_disjnt [OF _ \finite K\]) show "infinite {1/2..1::real}" by (simp add: infinite_Icc) have dis1: "disjnt (fro x) (fro y)" if "x b + d *\<^sub>R One" have cbsub: "cbox (-b) b \ box (-c) c" "cbox (-b) b \ cbox (-c) c" "cbox (-c) c \ bbox" using d12 by (auto simp: algebra_simps subset_box_imp c_def bbox_def) have clo_cbT: "closed (cbox (- c) c \ T)" by (simp add: affine_closed closed_Int closed_cbox \affine T\) have cpT_ne: "cbox (- c) c \ T \ {}" using \S \ {}\ b cbsub(2) \S \ T\ by fastforce have "closest_point (cbox (- c) c \ T) x \ K" if "x \ T" "x \ K" for x proof (cases "x \ cbox (-c) c") case True with that show ?thesis by (simp add: closest_point_self) next case False have int_ne: "interior (cbox (-c) c) \ T \ {}" using \S \ {}\ \S \ T\ b \cbox (- b) b \ box (- c) c\ by force have "convex T" by (meson \affine T\ affine_imp_convex) then have "x \ affine hull (cbox (- c) c \ T)" by (metis Int_commute Int_iff \S \ {}\ \S \ T\ cbsub(1) \x \ T\ affine_hull_convex_Int_nonempty_interior all_not_in_conv b hull_inc inf.orderE interior_cbox) then have "x \ affine hull (cbox (- c) c \ T) - rel_interior (cbox (- c) c \ T)" by (meson DiffI False Int_iff rel_interior_subset subsetCE) then have "closest_point (cbox (- c) c \ T) x \ rel_frontier (cbox (- c) c \ T)" by (rule closest_point_in_rel_frontier [OF clo_cbT cpT_ne]) moreover have "(rel_frontier (cbox (- c) c \ T)) \ fro d" apply (subst convex_affine_rel_frontier_Int [OF _ \affine T\ int_ne]) apply (auto simp: fro_def c_def) done ultimately show ?thesis using dd by (force simp: disjnt_def) qed then have cpt_subset: "closest_point (cbox (- c) c \ T) ` (T - K) \ \{bbox \ T} - K" using closest_point_in_set [OF clo_cbT cpT_ne] cbsub(3) by force show ?thesis proof (intro conjI ballI exI) have "continuous_on (T - K) (closest_point (cbox (- c) c \ T))" apply (rule continuous_on_closest_point) using \S \ {}\ cbsub(2) b that by (auto simp: affine_imp_convex convex_Int affine_closed closed_Int closed_cbox \affine T\) then show "continuous_on (T - K) (g \ closest_point (cbox (- c) c \ T))" by (metis continuous_on_compose continuous_on_subset [OF contg cpt_subset]) have "(g \ closest_point (cbox (- c) c \ T)) ` (T - K) \ g ` (\{bbox \ T} - K)" by (metis image_comp image_mono cpt_subset) also have "... \ rel_frontier U" by (rule gim) finally show "(g \ closest_point (cbox (- c) c \ T)) ` (T - K) \ rel_frontier U" . show "(g \ closest_point (cbox (- c) c \ T)) x = f x" if "x \ S" for x proof - have "(g \ closest_point (cbox (- c) c \ T)) x = g x" unfolding o_def by (metis IntI \S \ T\ b cbsub(2) closest_point_self subset_eq that) also have "... = f x" by (simp add: that gf) finally show ?thesis . qed qed (auto simp: K) qed then obtain K g where "finite K" "disjnt K S" and contg: "continuous_on (affine hull T - K) g" and gim: "g ` (affine hull T - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" by (metis aff affine_affine_hull aff_dim_affine_hull order_trans [OF \S \ T\ hull_subset [of T affine]]) then obtain K g where "finite K" "disjnt K S" and contg: "continuous_on (T - K) g" and gim: "g ` (T - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" by (rule_tac K=K and g=g in that) (auto simp: hull_inc elim: continuous_on_subset) then show ?thesis by (rule_tac K="K \ T" and g=g in that) (auto simp: disjnt_iff Diff_Int contg) qed subsection\Extending maps to spheres\ (*Up to extend_map_affine_to_sphere_cofinite_gen*) lemma extend_map_affine_to_sphere1: fixes f :: "'a::euclidean_space \ 'b::topological_space" assumes "finite K" "affine U" and contf: "continuous_on (U - K) f" and fim: "f ` (U - K) \ T" and comps: "\C. \C \ components(U - S); C \ K \ {}\ \ C \ L \ {}" and clo: "closedin (top_of_set U) S" and K: "disjnt K S" "K \ U" obtains g where "continuous_on (U - L) g" "g ` (U - L) \ T" "\x. x \ S \ g x = f x" proof (cases "K = {}") case True then show ?thesis by (metis Diff_empty Diff_subset contf fim continuous_on_subset image_subsetI rev_image_eqI subset_iff that) next case False have "S \ U" using clo closedin_limpt by blast then have "(U - S) \ K \ {}" by (metis Diff_triv False Int_Diff K disjnt_def inf.absorb_iff2 inf_commute) then have "\(components (U - S)) \ K \ {}" using Union_components by simp then obtain C0 where C0: "C0 \ components (U - S)" "C0 \ K \ {}" by blast have "convex U" by (simp add: affine_imp_convex \affine U\) then have "locally connected U" by (rule convex_imp_locally_connected) have "\a g. a \ C \ a \ L \ continuous_on (S \ (C - {a})) g \ g ` (S \ (C - {a})) \ T \ (\x \ S. g x = f x)" if C: "C \ components (U - S)" and CK: "C \ K \ {}" for C proof - have "C \ U-S" "C \ L \ {}" by (simp_all add: in_components_subset comps that) then obtain a where a: "a \ C" "a \ L" by auto have opeUC: "openin (top_of_set U) C" proof (rule openin_trans) show "openin (top_of_set (U-S)) C" by (simp add: \locally connected U\ clo locally_diff_closed openin_components_locally_connected [OF _ C]) show "openin (top_of_set U) (U - S)" by (simp add: clo openin_diff) qed then obtain d where "C \ U" "0 < d" and d: "cball a d \ U \ C" using openin_contains_cball by (metis \a \ C\) then have "ball a d \ U \ C" by auto obtain h k where homhk: "homeomorphism (S \ C) (S \ C) h k" and subC: "{x. (\ (h x = x \ k x = x))} \ C" and bou: "bounded {x. (\ (h x = x \ k x = x))}" and hin: "\x. x \ C \ K \ h x \ ball a d \ U" proof (rule homeomorphism_grouping_points_exists_gen [of C "ball a d \ U" "C \ K" "S \ C"]) show "openin (top_of_set C) (ball a d \ U)" by (metis open_ball \C \ U\ \ball a d \ U \ C\ inf.absorb_iff2 inf.orderE inf_assoc open_openin openin_subtopology) show "openin (top_of_set (affine hull C)) C" by (metis \a \ C\ \openin (top_of_set U) C\ affine_hull_eq affine_hull_openin all_not_in_conv \affine U\) show "ball a d \ U \ {}" using \0 < d\ \C \ U\ \a \ C\ by force show "finite (C \ K)" by (simp add: \finite K\) show "S \ C \ affine hull C" by (metis \C \ U\ \S \ U\ \a \ C\ opeUC affine_hull_eq affine_hull_openin all_not_in_conv assms(2) sup.bounded_iff) show "connected C" by (metis C in_components_connected) qed auto have a_BU: "a \ ball a d \ U" using \0 < d\ \C \ U\ \a \ C\ by auto have "rel_frontier (cball a d \ U) retract_of (affine hull (cball a d \ U) - {a})" apply (rule rel_frontier_retract_of_punctured_affine_hull) apply (auto simp: \convex U\ convex_Int) by (metis \affine U\ convex_cball empty_iff interior_cball a_BU rel_interior_convex_Int_affine) moreover have "rel_frontier (cball a d \ U) = frontier (cball a d) \ U" apply (rule convex_affine_rel_frontier_Int) using a_BU by (force simp: \affine U\)+ moreover have "affine hull (cball a d \ U) = U" by (metis \convex U\ a_BU affine_hull_convex_Int_nonempty_interior affine_hull_eq \affine U\ equals0D inf.commute interior_cball) ultimately have "frontier (cball a d) \ U retract_of (U - {a})" by metis then obtain r where contr: "continuous_on (U - {a}) r" and rim: "r ` (U - {a}) \ sphere a d" "r ` (U - {a}) \ U" and req: "\x. x \ sphere a d \ U \ r x = x" using \affine U\ by (auto simp: retract_of_def retraction_def hull_same) define j where "j \ \x. if x \ ball a d then r x else x" have kj: "\x. x \ S \ k (j x) = x" using \C \ U - S\ \S \ U\ \ball a d \ U \ C\ j_def subC by auto have Uaeq: "U - {a} = (cball a d - {a}) \ U \ (U - ball a d)" using \0 < d\ by auto have jim: "j ` (S \ (C - {a})) \ (S \ C) - ball a d" proof clarify fix y assume "y \ S \ (C - {a})" then have "y \ U - {a}" using \C \ U - S\ \S \ U\ \a \ C\ by auto then have "r y \ sphere a d" using rim by auto then show "j y \ S \ C - ball a d" apply (simp add: j_def) using \r y \ sphere a d\ \y \ U - {a}\ \y \ S \ (C - {a})\ d rim by fastforce qed have contj: "continuous_on (U - {a}) j" unfolding j_def Uaeq proof (intro continuous_on_cases_local continuous_on_id, simp_all add: req closedin_closed Uaeq [symmetric]) show "\T. closed T \ (cball a d - {a}) \ U = (U - {a}) \ T" apply (rule_tac x="(cball a d) \ U" in exI) using affine_closed \affine U\ by blast show "\T. closed T \ U - ball a d = (U - {a}) \ T" apply (rule_tac x="U - ball a d" in exI) using \0 < d\ by (force simp: affine_closed \affine U\ closed_Diff) show "continuous_on ((cball a d - {a}) \ U) r" by (force intro: continuous_on_subset [OF contr]) qed have fT: "x \ U - K \ f x \ T" for x using fim by blast show ?thesis proof (intro conjI exI) show "continuous_on (S \ (C - {a})) (f \ k \ j)" proof (intro continuous_on_compose) show "continuous_on (S \ (C - {a})) j" apply (rule continuous_on_subset [OF contj]) using \C \ U - S\ \S \ U\ \a \ C\ by force show "continuous_on (j ` (S \ (C - {a}))) k" apply (rule continuous_on_subset [OF homeomorphism_cont2 [OF homhk]]) - using jim \C \ U - S\ \S \ U\ \ball a d \ U \ C\ j_def by fastforce + using jim \C \ U - S\ \S \ U\ \ball a d \ U \ C\ j_def by blast show "continuous_on (k ` j ` (S \ (C - {a}))) f" proof (clarify intro!: continuous_on_subset [OF contf]) fix y assume "y \ S \ (C - {a})" have ky: "k y \ S \ C" using homeomorphism_image2 [OF homhk] \y \ S \ (C - {a})\ by blast have jy: "j y \ S \ C - ball a d" using Un_iff \y \ S \ (C - {a})\ jim by auto show "k (j y) \ U - K" apply safe using \C \ U\ \S \ U\ homeomorphism_image2 [OF homhk] jy apply blast by (metis DiffD1 DiffD2 Int_iff Un_iff \disjnt K S\ disjnt_def empty_iff hin homeomorphism_apply2 homeomorphism_image2 homhk imageI jy) qed qed have ST: "\x. x \ S \ (f \ k \ j) x \ T" apply (simp add: kj) apply (metis DiffI \S \ U\ \disjnt K S\ subsetD disjnt_iff fim image_subset_iff) done moreover have "(f \ k \ j) x \ T" if "x \ C" "x \ a" "x \ S" for x proof - have rx: "r x \ sphere a d" using \C \ U\ rim that by fastforce have jj: "j x \ S \ C - ball a d" using jim that by blast have "k (j x) = j x \ k (j x) \ C \ j x \ C" by (metis Diff_iff Int_iff Un_iff \S \ U\ subsetD d j_def jj rx sphere_cball that(1)) then have "k (j x) \ C" using homeomorphism_apply2 [OF homhk, of "j x"] \C \ U\ \S \ U\ a rx by (metis (mono_tags, lifting) Diff_iff subsetD jj mem_Collect_eq subC) with jj \C \ U\ show ?thesis apply safe using ST j_def apply fastforce apply (auto simp: not_less intro!: fT) by (metis DiffD1 DiffD2 Int_iff hin homeomorphism_apply2 [OF homhk] jj) qed ultimately show "(f \ k \ j) ` (S \ (C - {a})) \ T" by force show "\x\S. (f \ k \ j) x = f x" using kj by simp qed (auto simp: a) qed then obtain a h where ah: "\C. \C \ components (U - S); C \ K \ {}\ \ a C \ C \ a C \ L \ continuous_on (S \ (C - {a C})) (h C) \ h C ` (S \ (C - {a C})) \ T \ (\x \ S. h C x = f x)" using that by metis define F where "F \ {C \ components (U - S). C \ K \ {}}" define G where "G \ {C \ components (U - S). C \ K = {}}" define UF where "UF \ (\C\F. C - {a C})" have "C0 \ F" by (auto simp: F_def C0) have "finite F" proof (subst finite_image_iff [of "\C. C \ K" F, symmetric]) show "inj_on (\C. C \ K) F" unfolding F_def inj_on_def using components_nonoverlap by blast show "finite ((\C. C \ K) ` F)" unfolding F_def by (rule finite_subset [of _ "Pow K"]) (auto simp: \finite K\) qed obtain g where contg: "continuous_on (S \ UF) g" and gh: "\x i. \i \ F; x \ (S \ UF) \ (S \ (i - {a i}))\ \ g x = h i x" proof (rule pasting_lemma_exists_closed [OF \finite F\]) let ?X = "top_of_set (S \ UF)" show "topspace ?X \ (\C\F. S \ (C - {a C}))" using \C0 \ F\ by (force simp: UF_def) show "closedin (top_of_set (S \ UF)) (S \ (C - {a C}))" if "C \ F" for C proof (rule closedin_closed_subset [of U "S \ C"]) show "closedin (top_of_set U) (S \ C)" apply (rule closedin_Un_complement_component [OF \locally connected U\ clo]) using F_def that by blast next have "x = a C'" if "C' \ F" "x \ C'" "x \ U" for x C' proof - have "\A. x \ \A \ C' \ A" using \x \ C'\ by blast with that show "x = a C'" by (metis (lifting) DiffD1 F_def Union_components mem_Collect_eq) qed then show "S \ UF \ U" using \S \ U\ by (force simp: UF_def) next show "S \ (C - {a C}) = (S \ C) \ (S \ UF)" using F_def UF_def components_nonoverlap that by auto qed show "continuous_map (subtopology ?X (S \ (C' - {a C'}))) euclidean (h C')" if "C' \ F" for C' proof - have C': "C' \ components (U - S)" "C' \ K \ {}" using F_def that by blast+ show ?thesis using ah [OF C'] by (auto simp: F_def subtopology_subtopology intro: continuous_on_subset) qed show "\i j x. \i \ F; j \ F; x \ topspace ?X \ (S \ (i - {a i})) \ (S \ (j - {a j}))\ \ h i x = h j x" using components_eq by (fastforce simp: components_eq F_def ah) qed auto have SU': "S \ \G \ (S \ UF) \ U" using \S \ U\ in_components_subset by (auto simp: F_def G_def UF_def) have clo1: "closedin (top_of_set (S \ \G \ (S \ UF))) (S \ \G)" proof (rule closedin_closed_subset [OF _ SU']) have *: "\C. C \ F \ openin (top_of_set U) C" unfolding F_def by clarify (metis (no_types, lifting) \locally connected U\ clo closedin_def locally_diff_closed openin_components_locally_connected openin_trans topspace_euclidean_subtopology) show "closedin (top_of_set U) (U - UF)" unfolding UF_def by (force intro: openin_delete *) show "S \ \G = (U - UF) \ (S \ \G \ (S \ UF))" using \S \ U\ apply (auto simp: F_def G_def UF_def) apply (metis Diff_iff UnionI Union_components) apply (metis DiffD1 UnionI Union_components) by (metis (no_types, lifting) IntI components_nonoverlap empty_iff) qed have clo2: "closedin (top_of_set (S \ \G \ (S \ UF))) (S \ UF)" proof (rule closedin_closed_subset [OF _ SU']) show "closedin (top_of_set U) (\C\F. S \ C)" apply (rule closedin_Union) apply (simp add: \finite F\) using F_def \locally connected U\ clo closedin_Un_complement_component by blast show "S \ UF = (\C\F. S \ C) \ (S \ \G \ (S \ UF))" using \S \ U\ apply (auto simp: F_def G_def UF_def) using C0 apply blast by (metis components_nonoverlap disjnt_def disjnt_iff) qed have SUG: "S \ \G \ U - K" using \S \ U\ K apply (auto simp: G_def disjnt_iff) by (meson Diff_iff subsetD in_components_subset) then have contf': "continuous_on (S \ \G) f" by (rule continuous_on_subset [OF contf]) have contg': "continuous_on (S \ UF) g" apply (rule continuous_on_subset [OF contg]) using \S \ U\ by (auto simp: F_def G_def) have "\x. \S \ U; x \ S\ \ f x = g x" by (subst gh) (auto simp: ah C0 intro: \C0 \ F\) then have f_eq_g: "\x. x \ S \ UF \ x \ S \ \G \ f x = g x" using \S \ U\ apply (auto simp: F_def G_def UF_def dest: in_components_subset) using components_eq by blast have cont: "continuous_on (S \ \G \ (S \ UF)) (\x. if x \ S \ \G then f x else g x)" by (blast intro: continuous_on_cases_local [OF clo1 clo2 contf' contg' f_eq_g, of "\x. x \ S \ \G"]) show ?thesis proof have UF: "\F - L \ UF" unfolding F_def UF_def using ah by blast have "U - S - L = \(components (U - S)) - L" by simp also have "... = \F \ \G - L" unfolding F_def G_def by blast also have "... \ UF \ \G" using UF by blast finally have "U - L \ S \ \G \ (S \ UF)" by blast then show "continuous_on (U - L) (\x. if x \ S \ \G then f x else g x)" by (rule continuous_on_subset [OF cont]) have "((U - L) \ {x. x \ S \ (\xa\G. x \ xa)}) \ ((U - L) \ (-S \ UF))" using \U - L \ S \ \G \ (S \ UF)\ by auto moreover have "g ` ((U - L) \ (-S \ UF)) \ T" proof - have "g x \ T" if "x \ U" "x \ L" "x \ S" "C \ F" "x \ C" "x \ a C" for x C proof (subst gh) show "x \ (S \ UF) \ (S \ (C - {a C}))" using that by (auto simp: UF_def) show "h C x \ T" using ah that by (fastforce simp add: F_def) qed (rule that) then show ?thesis by (force simp: UF_def) qed ultimately have "g ` ((U - L) \ {x. x \ S \ (\xa\G. x \ xa)}) \ T" using image_mono order_trans by blast moreover have "f ` ((U - L) \ (S \ \G)) \ T" using fim SUG by blast ultimately show "(\x. if x \ S \ \G then f x else g x) ` (U - L) \ T" by force show "\x. x \ S \ (if x \ S \ \G then f x else g x) = f x" by (simp add: F_def G_def) qed qed lemma extend_map_affine_to_sphere2: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact S" "convex U" "bounded U" "affine T" "S \ T" and affTU: "aff_dim T \ aff_dim U" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier U" and ovlap: "\C. C \ components(T - S) \ C \ L \ {}" obtains K g where "finite K" "K \ L" "K \ T" "disjnt K S" "continuous_on (T - K) g" "g ` (T - K) \ rel_frontier U" "\x. x \ S \ g x = f x" proof - obtain K g where K: "finite K" "K \ T" "disjnt K S" and contg: "continuous_on (T - K) g" and gim: "g ` (T - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" using assms extend_map_affine_to_sphere_cofinite_simple by metis have "(\y C. C \ components (T - S) \ x \ C \ y \ C \ y \ L)" if "x \ K" for x proof - have "x \ T-S" using \K \ T\ \disjnt K S\ disjnt_def that by fastforce then obtain C where "C \ components(T - S)" "x \ C" by (metis UnionE Union_components) with ovlap [of C] show ?thesis by blast qed then obtain \ where \: "\x. x \ K \ \C. C \ components (T - S) \ x \ C \ \ x \ C \ \ x \ L" by metis obtain h where conth: "continuous_on (T - \ ` K) h" and him: "h ` (T - \ ` K) \ rel_frontier U" and hg: "\x. x \ S \ h x = g x" proof (rule extend_map_affine_to_sphere1 [OF \finite K\ \affine T\ contg gim, of S "\ ` K"]) show cloTS: "closedin (top_of_set T) S" by (simp add: \compact S\ \S \ T\ closed_subset compact_imp_closed) show "\C. \C \ components (T - S); C \ K \ {}\ \ C \ \ ` K \ {}" using \ components_eq by blast qed (use K in auto) show ?thesis proof show *: "\ ` K \ L" using \ by blast show "finite (\ ` K)" by (simp add: K) show "\ ` K \ T" by clarify (meson \ Diff_iff contra_subsetD in_components_subset) show "continuous_on (T - \ ` K) h" by (rule conth) show "disjnt (\ ` K) S" using K apply (auto simp: disjnt_def) by (metis \ DiffD2 UnionI Union_components) qed (simp_all add: him hg gf) qed proposition extend_map_affine_to_sphere_cofinite_gen: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes SUT: "compact S" "convex U" "bounded U" "affine T" "S \ T" and aff: "aff_dim T \ aff_dim U" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier U" and dis: "\C. \C \ components(T - S); bounded C\ \ C \ L \ {}" obtains K g where "finite K" "K \ L" "K \ T" "disjnt K S" "continuous_on (T - K) g" "g ` (T - K) \ rel_frontier U" "\x. x \ S \ g x = f x" proof (cases "S = {}") case True show ?thesis proof (cases "rel_frontier U = {}") case True with aff have "aff_dim T \ 0" apply (simp add: rel_frontier_eq_empty) using affine_bounded_eq_lowdim \bounded U\ order_trans by auto with aff_dim_geq [of T] consider "aff_dim T = -1" | "aff_dim T = 0" by linarith then show ?thesis proof cases assume "aff_dim T = -1" then have "T = {}" by (simp add: aff_dim_empty) then show ?thesis by (rule_tac K="{}" in that) auto next assume "aff_dim T = 0" then obtain a where "T = {a}" using aff_dim_eq_0 by blast then have "a \ L" using dis [of "{a}"] \S = {}\ by (auto simp: in_components_self) with \S = {}\ \T = {a}\ show ?thesis by (rule_tac K="{a}" and g=f in that) auto qed next case False then obtain y where "y \ rel_frontier U" by auto with \S = {}\ show ?thesis by (rule_tac K="{}" and g="\x. y" in that) (auto) qed next case False have "bounded S" by (simp add: assms compact_imp_bounded) then obtain b where b: "S \ cbox (-b) b" using bounded_subset_cbox_symmetric by blast define LU where "LU \ L \ (\ {C \ components (T - S). \bounded C} - cbox (-(b+One)) (b+One))" obtain K g where "finite K" "K \ LU" "K \ T" "disjnt K S" and contg: "continuous_on (T - K) g" and gim: "g ` (T - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" proof (rule extend_map_affine_to_sphere2 [OF SUT aff contf fim]) show "C \ LU \ {}" if "C \ components (T - S)" for C proof (cases "bounded C") case True with dis that show ?thesis unfolding LU_def by fastforce next case False then have "\ bounded (\{C \ components (T - S). \ bounded C})" by (metis (no_types, lifting) Sup_upper bounded_subset mem_Collect_eq that) then show ?thesis apply (clarsimp simp: LU_def Int_Un_distrib Diff_Int_distrib Int_UN_distrib) by (metis (no_types, lifting) False Sup_upper bounded_cbox bounded_subset inf.orderE mem_Collect_eq that) qed qed blast have *: False if "x \ cbox (- b - m *\<^sub>R One) (b + m *\<^sub>R One)" "x \ box (- b - n *\<^sub>R One) (b + n *\<^sub>R One)" "0 \ m" "m < n" "n \ 1" for m n x using that by (auto simp: mem_box algebra_simps) have "disjoint_family_on (\d. frontier (cbox (- b - d *\<^sub>R One) (b + d *\<^sub>R One))) {1 / 2..1}" by (auto simp: disjoint_family_on_def neq_iff frontier_def dest: *) then obtain d where d12: "1/2 \ d" "d \ 1" and ddis: "disjnt K (frontier (cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One)))" using disjoint_family_elem_disjnt [of "{1/2..1::real}" K "\d. frontier (cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One))"] by (auto simp: \finite K\) define c where "c \ b + d *\<^sub>R One" have cbsub: "cbox (-b) b \ box (-c) c" "cbox (-b) b \ cbox (-c) c" "cbox (-c) c \ cbox (-(b+One)) (b+One)" using d12 by (simp_all add: subset_box c_def inner_diff_left inner_left_distrib) have clo_cT: "closed (cbox (- c) c \ T)" using affine_closed \affine T\ by blast have cT_ne: "cbox (- c) c \ T \ {}" using \S \ {}\ \S \ T\ b cbsub by fastforce have S_sub_cc: "S \ cbox (- c) c" using \cbox (- b) b \ cbox (- c) c\ b by auto show ?thesis proof show "finite (K \ cbox (-(b+One)) (b+One))" using \finite K\ by blast show "K \ cbox (- (b + One)) (b + One) \ L" using \K \ LU\ by (auto simp: LU_def) show "K \ cbox (- (b + One)) (b + One) \ T" using \K \ T\ by auto show "disjnt (K \ cbox (- (b + One)) (b + One)) S" using \disjnt K S\ by (simp add: disjnt_def disjoint_eq_subset_Compl inf.coboundedI1) have cloTK: "closest_point (cbox (- c) c \ T) x \ T - K" if "x \ T" and Knot: "x \ K \ x \ cbox (- b - One) (b + One)" for x proof (cases "x \ cbox (- c) c") case True with \x \ T\ show ?thesis using cbsub(3) Knot by (force simp: closest_point_self) next case False have clo_in_rf: "closest_point (cbox (- c) c \ T) x \ rel_frontier (cbox (- c) c \ T)" proof (intro closest_point_in_rel_frontier [OF clo_cT cT_ne] DiffI notI) have "T \ interior (cbox (- c) c) \ {}" using \S \ {}\ \S \ T\ b cbsub(1) by fastforce then show "x \ affine hull (cbox (- c) c \ T)" by (simp add: Int_commute affine_hull_affine_Int_nonempty_interior \affine T\ hull_inc that(1)) next show "False" if "x \ rel_interior (cbox (- c) c \ T)" proof - have "interior (cbox (- c) c) \ T \ {}" using \S \ {}\ \S \ T\ b cbsub(1) by fastforce then have "affine hull (T \ cbox (- c) c) = T" using affine_hull_convex_Int_nonempty_interior [of T "cbox (- c) c"] by (simp add: affine_imp_convex \affine T\ inf_commute) then show ?thesis by (meson subsetD le_inf_iff rel_interior_subset that False) qed qed have "closest_point (cbox (- c) c \ T) x \ K" proof assume inK: "closest_point (cbox (- c) c \ T) x \ K" have "\x. x \ K \ x \ frontier (cbox (- (b + d *\<^sub>R One)) (b + d *\<^sub>R One))" by (metis ddis disjnt_iff) then show False by (metis DiffI Int_iff \affine T\ cT_ne c_def clo_cT clo_in_rf closest_point_in_set convex_affine_rel_frontier_Int convex_box(1) empty_iff frontier_cbox inK interior_cbox) qed then show ?thesis using cT_ne clo_cT closest_point_in_set by blast qed show "continuous_on (T - K \ cbox (- (b + One)) (b + One)) (g \ closest_point (cbox (-c) c \ T))" apply (intro continuous_on_compose continuous_on_closest_point continuous_on_subset [OF contg]) apply (simp_all add: clo_cT affine_imp_convex \affine T\ convex_Int cT_ne) using cloTK by blast have "g (closest_point (cbox (- c) c \ T) x) \ rel_frontier U" if "x \ T" "x \ K \ x \ cbox (- b - One) (b + One)" for x apply (rule gim [THEN subsetD]) using that cloTK by blast then show "(g \ closest_point (cbox (- c) c \ T)) ` (T - K \ cbox (- (b + One)) (b + One)) \ rel_frontier U" by force show "\x. x \ S \ (g \ closest_point (cbox (- c) c \ T)) x = f x" by simp (metis (mono_tags, lifting) IntI \S \ T\ cT_ne clo_cT closest_point_refl gf subsetD S_sub_cc) qed qed corollary extend_map_affine_to_sphere_cofinite: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes SUT: "compact S" "affine T" "S \ T" and aff: "aff_dim T \ DIM('b)" and "0 \ r" and contf: "continuous_on S f" and fim: "f ` S \ sphere a r" and dis: "\C. \C \ components(T - S); bounded C\ \ C \ L \ {}" obtains K g where "finite K" "K \ L" "K \ T" "disjnt K S" "continuous_on (T - K) g" "g ` (T - K) \ sphere a r" "\x. x \ S \ g x = f x" proof (cases "r = 0") case True with fim show ?thesis by (rule_tac K="{}" and g = "\x. a" in that) (auto) next case False with assms have "0 < r" by auto then have "aff_dim T \ aff_dim (cball a r)" by (simp add: aff aff_dim_cball) then show ?thesis apply (rule extend_map_affine_to_sphere_cofinite_gen [OF \compact S\ convex_cball bounded_cball \affine T\ \S \ T\ _ contf]) using fim apply (auto simp: assms False that dest: dis) done qed corollary extend_map_UNIV_to_sphere_cofinite: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes aff: "DIM('a) \ DIM('b)" and "0 \ r" and SUT: "compact S" and contf: "continuous_on S f" and fim: "f ` S \ sphere a r" and dis: "\C. \C \ components(- S); bounded C\ \ C \ L \ {}" obtains K g where "finite K" "K \ L" "disjnt K S" "continuous_on (- K) g" "g ` (- K) \ sphere a r" "\x. x \ S \ g x = f x" apply (rule extend_map_affine_to_sphere_cofinite [OF \compact S\ affine_UNIV subset_UNIV _ \0 \ r\ contf fim dis]) apply (auto simp: assms that Compl_eq_Diff_UNIV [symmetric]) done corollary extend_map_UNIV_to_sphere_no_bounded_component: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes aff: "DIM('a) \ DIM('b)" and "0 \ r" and SUT: "compact S" and contf: "continuous_on S f" and fim: "f ` S \ sphere a r" and dis: "\C. C \ components(- S) \ \ bounded C" obtains g where "continuous_on UNIV g" "g ` UNIV \ sphere a r" "\x. x \ S \ g x = f x" apply (rule extend_map_UNIV_to_sphere_cofinite [OF aff \0 \ r\ \compact S\ contf fim, of "{}"]) apply (auto simp: that dest: dis) done theorem Borsuk_separation_theorem_gen: fixes S :: "'a::euclidean_space set" assumes "compact S" shows "(\c \ components(- S). \bounded c) \ (\f. continuous_on S f \ f ` S \ sphere (0::'a) 1 \ (\c. homotopic_with_canon (\x. True) S (sphere 0 1) f (\x. c)))" (is "?lhs = ?rhs") proof assume L [rule_format]: ?lhs show ?rhs proof clarify fix f :: "'a \ 'a" assume contf: "continuous_on S f" and fim: "f ` S \ sphere 0 1" obtain g where contg: "continuous_on UNIV g" and gim: "range g \ sphere 0 1" and gf: "\x. x \ S \ g x = f x" by (rule extend_map_UNIV_to_sphere_no_bounded_component [OF _ _ \compact S\ contf fim L]) auto then obtain c where c: "homotopic_with_canon (\h. True) UNIV (sphere 0 1) g (\x. c)" using contractible_UNIV nullhomotopic_from_contractible by blast then show "\c. homotopic_with_canon (\x. True) S (sphere 0 1) f (\x. c)" by (metis assms compact_imp_closed contf contg contractible_empty fim gf gim nullhomotopic_from_contractible nullhomotopic_into_sphere_extension) qed next assume R [rule_format]: ?rhs show ?lhs unfolding components_def proof clarify fix a assume "a \ S" and a: "bounded (connected_component_set (- S) a)" have cont: "continuous_on S (\x. inverse(norm(x - a)) *\<^sub>R (x - a))" apply (intro continuous_intros) using \a \ S\ by auto have im: "(\x. inverse(norm(x - a)) *\<^sub>R (x - a)) ` S \ sphere 0 1" by clarsimp (metis \a \ S\ eq_iff_diff_eq_0 left_inverse norm_eq_zero) show False using R cont im Borsuk_map_essential_bounded_component [OF \compact S\ \a \ S\] a by blast qed qed corollary Borsuk_separation_theorem: fixes S :: "'a::euclidean_space set" assumes "compact S" and 2: "2 \ DIM('a)" shows "connected(- S) \ (\f. continuous_on S f \ f ` S \ sphere (0::'a) 1 \ (\c. homotopic_with_canon (\x. True) S (sphere 0 1) f (\x. c)))" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof (cases "S = {}") case True then show ?thesis by auto next case False then have "(\c\components (- S). \ bounded c)" by (metis L assms(1) bounded_empty cobounded_imp_unbounded compact_imp_bounded in_components_maximal order_refl) then show ?thesis by (simp add: Borsuk_separation_theorem_gen [OF \compact S\]) qed next assume R: ?rhs then show ?lhs apply (simp add: Borsuk_separation_theorem_gen [OF \compact S\, symmetric]) apply (auto simp: components_def connected_iff_eq_connected_component_set) using connected_component_in apply fastforce using cobounded_unique_unbounded_component [OF _ 2, of "-S"] \compact S\ compact_eq_bounded_closed by fastforce qed lemma homotopy_eqv_separation: fixes S :: "'a::euclidean_space set" and T :: "'a set" assumes "S homotopy_eqv T" and "compact S" and "compact T" shows "connected(- S) \ connected(- T)" proof - consider "DIM('a) = 1" | "2 \ DIM('a)" by (metis DIM_ge_Suc0 One_nat_def Suc_1 dual_order.antisym not_less_eq_eq) then show ?thesis proof cases case 1 then show ?thesis using bounded_connected_Compl_1 compact_imp_bounded homotopy_eqv_empty1 homotopy_eqv_empty2 assms by metis next case 2 with assms show ?thesis by (simp add: Borsuk_separation_theorem homotopy_eqv_cohomotopic_triviality_null) qed qed proposition Jordan_Brouwer_separation: fixes S :: "'a::euclidean_space set" and a::'a assumes hom: "S homeomorphic sphere a r" and "0 < r" shows "\ connected(- S)" proof - have "- sphere a r \ ball a r \ {}" using \0 < r\ by (simp add: Int_absorb1 subset_eq) moreover have eq: "- sphere a r - ball a r = - cball a r" by auto have "- cball a r \ {}" proof - have "frontier (cball a r) \ {}" using \0 < r\ by auto then show ?thesis by (metis frontier_complement frontier_empty) qed with eq have "- sphere a r - ball a r \ {}" by auto moreover have "connected (- S) = connected (- sphere a r)" proof (rule homotopy_eqv_separation) show "S homotopy_eqv sphere a r" using hom homeomorphic_imp_homotopy_eqv by blast show "compact (sphere a r)" by simp then show " compact S" using hom homeomorphic_compactness by blast qed ultimately show ?thesis using connected_Int_frontier [of "- sphere a r" "ball a r"] by (auto simp: \0 < r\) qed proposition Jordan_Brouwer_frontier: fixes S :: "'a::euclidean_space set" and a::'a assumes S: "S homeomorphic sphere a r" and T: "T \ components(- S)" and 2: "2 \ DIM('a)" shows "frontier T = S" proof (cases r rule: linorder_cases) assume "r < 0" with S T show ?thesis by auto next assume "r = 0" with S T card_eq_SucD obtain b where "S = {b}" by (auto simp: homeomorphic_finite [of "{a}" S]) have "components (- {b}) = { -{b}}" using T \S = {b}\ by (auto simp: components_eq_sing_iff connected_punctured_universe 2) with T show ?thesis by (metis \S = {b}\ cball_trivial frontier_cball frontier_complement singletonD sphere_trivial) next assume "r > 0" have "compact S" using homeomorphic_compactness compact_sphere S by blast show ?thesis proof (rule frontier_minimal_separating_closed) show "closed S" using \compact S\ compact_eq_bounded_closed by blast show "\ connected (- S)" using Jordan_Brouwer_separation S \0 < r\ by blast obtain f g where hom: "homeomorphism S (sphere a r) f g" using S by (auto simp: homeomorphic_def) show "connected (- T)" if "closed T" "T \ S" for T proof - have "f ` T \ sphere a r" using \T \ S\ hom homeomorphism_image1 by blast moreover have "f ` T \ sphere a r" using \T \ S\ hom by (metis homeomorphism_image2 homeomorphism_of_subsets order_refl psubsetE) ultimately have "f ` T \ sphere a r" by blast then have "connected (- f ` T)" by (rule psubset_sphere_Compl_connected [OF _ \0 < r\ 2]) moreover have "compact T" using \compact S\ bounded_subset compact_eq_bounded_closed that by blast moreover then have "compact (f ` T)" by (meson compact_continuous_image continuous_on_subset hom homeomorphism_def psubsetE \T \ S\) moreover have "T homotopy_eqv f ` T" by (meson \f ` T \ sphere a r\ dual_order.strict_implies_order hom homeomorphic_def homeomorphic_imp_homotopy_eqv homeomorphism_of_subsets \T \ S\) ultimately show ?thesis using homotopy_eqv_separation [of T "f`T"] by blast qed qed (rule T) qed proposition Jordan_Brouwer_nonseparation: fixes S :: "'a::euclidean_space set" and a::'a assumes S: "S homeomorphic sphere a r" and "T \ S" and 2: "2 \ DIM('a)" shows "connected(- T)" proof - have *: "connected(C \ (S - T))" if "C \ components(- S)" for C proof (rule connected_intermediate_closure) show "connected C" using in_components_connected that by auto have "S = frontier C" using "2" Jordan_Brouwer_frontier S that by blast with closure_subset show "C \ (S - T) \ closure C" by (auto simp: frontier_def) qed auto have "components(- S) \ {}" by (metis S bounded_empty cobounded_imp_unbounded compact_eq_bounded_closed compact_sphere components_eq_empty homeomorphic_compactness) then have "- T = (\C \ components(- S). C \ (S - T))" using Union_components [of "-S"] \T \ S\ by auto then show ?thesis apply (rule ssubst) apply (rule connected_Union) using \T \ S\ apply (auto simp: *) done qed subsection\ Invariance of domain and corollaries\ lemma invariance_of_domain_ball: fixes f :: "'a \ 'a::euclidean_space" assumes contf: "continuous_on (cball a r) f" and "0 < r" and inj: "inj_on f (cball a r)" shows "open(f ` ball a r)" proof (cases "DIM('a) = 1") case True obtain h::"'a\real" and k where "linear h" "linear k" "h ` UNIV = UNIV" "k ` UNIV = UNIV" "\x. norm(h x) = norm x" "\x. norm(k x) = norm x" "\x. k(h x) = x" "\x. h(k x) = x" apply (rule isomorphisms_UNIV_UNIV [where 'M='a and 'N=real]) using True apply force by (metis UNIV_I UNIV_eq_I imageI) have cont: "continuous_on S h" "continuous_on T k" for S T by (simp_all add: \linear h\ \linear k\ linear_continuous_on linear_linear) have "continuous_on (h ` cball a r) (h \ f \ k)" apply (intro continuous_on_compose cont continuous_on_subset [OF contf]) apply (auto simp: \\x. k (h x) = x\) done moreover have "is_interval (h ` cball a r)" by (simp add: is_interval_connected_1 \linear h\ linear_continuous_on linear_linear connected_continuous_image) moreover have "inj_on (h \ f \ k) (h ` cball a r)" using inj by (simp add: inj_on_def) (metis \\x. k (h x) = x\) ultimately have *: "\T. \open T; T \ h ` cball a r\ \ open ((h \ f \ k) ` T)" using injective_eq_1d_open_map_UNIV by blast have "open ((h \ f \ k) ` (h ` ball a r))" by (rule *) (auto simp: \linear h\ \range h = UNIV\ open_surjective_linear_image) then have "open ((h \ f) ` ball a r)" by (simp add: image_comp \\x. k (h x) = x\ cong: image_cong) then show ?thesis apply (simp only: image_comp [symmetric]) apply (metis open_bijective_linear_image_eq \linear h\ \\x. k (h x) = x\ \range h = UNIV\ bijI inj_on_def) done next case False then have 2: "DIM('a) \ 2" by (metis DIM_ge_Suc0 One_nat_def Suc_1 antisym not_less_eq_eq) have fimsub: "f ` ball a r \ - f ` sphere a r" using inj by clarsimp (metis inj_onD less_eq_real_def mem_cball order_less_irrefl) have hom: "f ` sphere a r homeomorphic sphere a r" by (meson compact_sphere contf continuous_on_subset homeomorphic_compact homeomorphic_sym inj inj_on_subset sphere_cball) then have nconn: "\ connected (- f ` sphere a r)" by (rule Jordan_Brouwer_separation) (auto simp: \0 < r\) obtain C where C: "C \ components (- f ` sphere a r)" and "bounded C" apply (rule cobounded_has_bounded_component [OF _ nconn]) apply (simp_all add: 2) by (meson compact_imp_bounded compact_continuous_image_eq compact_sphere contf inj sphere_cball) moreover have "f ` (ball a r) = C" proof have "C \ {}" by (rule in_components_nonempty [OF C]) show "C \ f ` ball a r" proof (rule ccontr) assume nonsub: "\ C \ f ` ball a r" have "- f ` cball a r \ C" proof (rule components_maximal [OF C]) have "f ` cball a r homeomorphic cball a r" using compact_cball contf homeomorphic_compact homeomorphic_sym inj by blast then show "connected (- f ` cball a r)" by (auto intro: connected_complement_homeomorphic_convex_compact 2) show "- f ` cball a r \ - f ` sphere a r" by auto then show "C \ - f ` cball a r \ {}" using \C \ {}\ in_components_subset [OF C] nonsub using image_iff by fastforce qed then have "bounded (- f ` cball a r)" using bounded_subset \bounded C\ by auto then have "\ bounded (f ` cball a r)" using cobounded_imp_unbounded by blast then show "False" using compact_continuous_image [OF contf] compact_cball compact_imp_bounded by blast qed with \C \ {}\ have "C \ f ` ball a r \ {}" by (simp add: inf.absorb_iff1) then show "f ` ball a r \ C" by (metis components_maximal [OF C _ fimsub] connected_continuous_image ball_subset_cball connected_ball contf continuous_on_subset) qed moreover have "open (- f ` sphere a r)" using hom compact_eq_bounded_closed compact_sphere homeomorphic_compactness by blast ultimately show ?thesis using open_components by blast qed text\Proved by L. E. J. Brouwer (1912)\ theorem invariance_of_domain: fixes f :: "'a \ 'a::euclidean_space" assumes "continuous_on S f" "open S" "inj_on f S" shows "open(f ` S)" unfolding open_subopen [of "f`S"] proof clarify fix a assume "a \ S" obtain \ where "\ > 0" and \: "cball a \ \ S" using \open S\ \a \ S\ open_contains_cball_eq by blast show "\T. open T \ f a \ T \ T \ f ` S" proof (intro exI conjI) show "open (f ` (ball a \))" by (meson \ \0 < \\ assms continuous_on_subset inj_on_subset invariance_of_domain_ball) show "f a \ f ` ball a \" by (simp add: \0 < \\) show "f ` ball a \ \ f ` S" using \ ball_subset_cball by blast qed qed lemma inv_of_domain_ss0: fixes f :: "'a \ 'a::euclidean_space" assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f ` U \ S" and "subspace S" and dimS: "dim S = DIM('b::euclidean_space)" and ope: "openin (top_of_set S) U" shows "openin (top_of_set S) (f ` U)" proof - have "U \ S" using ope openin_imp_subset by blast have "(UNIV::'b set) homeomorphic S" by (simp add: \subspace S\ dimS homeomorphic_subspaces) then obtain h k where homhk: "homeomorphism (UNIV::'b set) S h k" using homeomorphic_def by blast have homkh: "homeomorphism S (k ` S) k h" using homhk homeomorphism_image2 homeomorphism_sym by fastforce have "open ((k \ f \ h) ` k ` U)" proof (rule invariance_of_domain) show "continuous_on (k ` U) (k \ f \ h)" proof (intro continuous_intros) show "continuous_on (k ` U) h" by (meson continuous_on_subset [OF homeomorphism_cont1 [OF homhk]] top_greatest) show "continuous_on (h ` k ` U) f" apply (rule continuous_on_subset [OF contf], clarify) apply (metis homhk homeomorphism_def ope openin_imp_subset rev_subsetD) done show "continuous_on (f ` h ` k ` U) k" apply (rule continuous_on_subset [OF homeomorphism_cont2 [OF homhk]]) using fim homhk homeomorphism_apply2 ope openin_subset by fastforce qed have ope_iff: "\T. open T \ openin (top_of_set (k ` S)) T" using homhk homeomorphism_image2 open_openin by fastforce show "open (k ` U)" by (simp add: ope_iff homeomorphism_imp_open_map [OF homkh ope]) show "inj_on (k \ f \ h) (k ` U)" apply (clarsimp simp: inj_on_def) by (metis subsetD fim homeomorphism_apply2 [OF homhk] image_subset_iff inj_on_eq_iff injf \U \ S\) qed moreover have eq: "f ` U = h ` (k \ f \ h \ k) ` U" unfolding image_comp [symmetric] using \U \ S\ fim by (metis homeomorphism_image2 homeomorphism_of_subsets homkh subset_image_iff) ultimately show ?thesis by (metis (no_types, hide_lams) homeomorphism_imp_open_map homhk image_comp open_openin subtopology_UNIV) qed lemma inv_of_domain_ss1: fixes f :: "'a \ 'a::euclidean_space" assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f ` U \ S" and "subspace S" and ope: "openin (top_of_set S) U" shows "openin (top_of_set S) (f ` U)" proof - define S' where "S' \ {y. \x \ S. orthogonal x y}" have "subspace S'" by (simp add: S'_def subspace_orthogonal_to_vectors) define g where "g \ \z::'a*'a. ((f \ fst)z, snd z)" have "openin (top_of_set (S \ S')) (g ` (U \ S'))" proof (rule inv_of_domain_ss0) show "continuous_on (U \ S') g" apply (simp add: g_def) apply (intro continuous_intros continuous_on_compose2 [OF contf continuous_on_fst], auto) done show "g ` (U \ S') \ S \ S'" using fim by (auto simp: g_def) show "inj_on g (U \ S')" using injf by (auto simp: g_def inj_on_def) show "subspace (S \ S')" by (simp add: \subspace S'\ \subspace S\ subspace_Times) show "openin (top_of_set (S \ S')) (U \ S')" by (simp add: openin_Times [OF ope]) have "dim (S \ S') = dim S + dim S'" by (simp add: \subspace S'\ \subspace S\ dim_Times) also have "... = DIM('a)" using dim_subspace_orthogonal_to_vectors [OF \subspace S\ subspace_UNIV] by (simp add: add.commute S'_def) finally show "dim (S \ S') = DIM('a)" . qed moreover have "g ` (U \ S') = f ` U \ S'" by (auto simp: g_def image_iff) moreover have "0 \ S'" using \subspace S'\ subspace_affine by blast ultimately show ?thesis by (auto simp: openin_Times_eq) qed corollary invariance_of_domain_subspaces: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and "subspace U" "subspace V" and VU: "dim V \ dim U" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" shows "openin (top_of_set V) (f ` S)" proof - obtain V' where "subspace V'" "V' \ U" "dim V' = dim V" using choose_subspace_of_subspace [OF VU] by (metis span_eq_iff \subspace U\) then have "V homeomorphic V'" by (simp add: \subspace V\ homeomorphic_subspaces) then obtain h k where homhk: "homeomorphism V V' h k" using homeomorphic_def by blast have eq: "f ` S = k ` (h \ f) ` S" proof - have "k ` h ` f ` S = f ` S" by (meson fim homeomorphism_def homeomorphism_of_subsets homhk subset_refl) then show ?thesis by (simp add: image_comp) qed show ?thesis unfolding eq proof (rule homeomorphism_imp_open_map) show homkh: "homeomorphism V' V k h" by (simp add: homeomorphism_symD homhk) have hfV': "(h \ f) ` S \ V'" using fim homeomorphism_image1 homhk by fastforce moreover have "openin (top_of_set U) ((h \ f) ` S)" proof (rule inv_of_domain_ss1) show "continuous_on S (h \ f)" by (meson contf continuous_on_compose continuous_on_subset fim homeomorphism_cont1 homhk) show "inj_on (h \ f) S" apply (clarsimp simp: inj_on_def) by (metis fim homeomorphism_apply2 [OF homkh] image_subset_iff inj_onD injf) show "(h \ f) ` S \ U" using \V' \ U\ hfV' by auto qed (auto simp: assms) ultimately show "openin (top_of_set V') ((h \ f) ` S)" using openin_subset_trans \V' \ U\ by force qed qed corollary invariance_of_dimension_subspaces: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and "subspace U" "subspace V" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" and "S \ {}" shows "dim U \ dim V" proof - have "False" if "dim V < dim U" proof - obtain T where "subspace T" "T \ U" "dim T = dim V" using choose_subspace_of_subspace [of "dim V" U] by (metis \dim V < dim U\ assms(2) order.strict_implies_order span_eq_iff) then have "V homeomorphic T" by (simp add: \subspace V\ homeomorphic_subspaces) then obtain h k where homhk: "homeomorphism V T h k" using homeomorphic_def by blast have "continuous_on S (h \ f)" by (meson contf continuous_on_compose continuous_on_subset fim homeomorphism_cont1 homhk) moreover have "(h \ f) ` S \ U" using \T \ U\ fim homeomorphism_image1 homhk by fastforce moreover have "inj_on (h \ f) S" apply (clarsimp simp: inj_on_def) by (metis fim homeomorphism_apply1 homhk image_subset_iff inj_onD injf) ultimately have ope_hf: "openin (top_of_set U) ((h \ f) ` S)" using invariance_of_domain_subspaces [OF ope \subspace U\ \subspace U\] by blast have "(h \ f) ` S \ T" using fim homeomorphism_image1 homhk by fastforce then have "dim ((h \ f) ` S) \ dim T" by (rule dim_subset) also have "dim ((h \ f) ` S) = dim U" using \S \ {}\ \subspace U\ by (blast intro: dim_openin ope_hf) finally show False using \dim V < dim U\ \dim T = dim V\ by simp qed then show ?thesis using not_less by blast qed corollary invariance_of_domain_affine_sets: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and aff: "affine U" "affine V" "aff_dim V \ aff_dim U" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" shows "openin (top_of_set V) (f ` S)" proof (cases "S = {}") case True then show ?thesis by auto next case False obtain a b where "a \ S" "a \ U" "b \ V" using False fim ope openin_contains_cball by fastforce have "openin (top_of_set ((+) (- b) ` V)) (((+) (- b) \ f \ (+) a) ` (+) (- a) ` S)" proof (rule invariance_of_domain_subspaces) show "openin (top_of_set ((+) (- a) ` U)) ((+) (- a) ` S)" by (metis ope homeomorphism_imp_open_map homeomorphism_translation translation_galois) show "subspace ((+) (- a) ` U)" by (simp add: \a \ U\ affine_diffs_subspace_subtract \affine U\ cong: image_cong_simp) show "subspace ((+) (- b) ` V)" by (simp add: \b \ V\ affine_diffs_subspace_subtract \affine V\ cong: image_cong_simp) show "dim ((+) (- b) ` V) \ dim ((+) (- a) ` U)" by (metis \a \ U\ \b \ V\ aff_dim_eq_dim affine_hull_eq aff of_nat_le_iff) show "continuous_on ((+) (- a) ` S) ((+) (- b) \ f \ (+) a)" by (metis contf continuous_on_compose homeomorphism_cont2 homeomorphism_translation translation_galois) show "((+) (- b) \ f \ (+) a) ` (+) (- a) ` S \ (+) (- b) ` V" using fim by auto show "inj_on ((+) (- b) \ f \ (+) a) ((+) (- a) ` S)" by (auto simp: inj_on_def) (meson inj_onD injf) qed then show ?thesis by (metis (no_types, lifting) homeomorphism_imp_open_map homeomorphism_translation image_comp translation_galois) qed corollary invariance_of_dimension_affine_sets: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and aff: "affine U" "affine V" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" and "S \ {}" shows "aff_dim U \ aff_dim V" proof - obtain a b where "a \ S" "a \ U" "b \ V" using \S \ {}\ fim ope openin_contains_cball by fastforce have "dim ((+) (- a) ` U) \ dim ((+) (- b) ` V)" proof (rule invariance_of_dimension_subspaces) show "openin (top_of_set ((+) (- a) ` U)) ((+) (- a) ` S)" by (metis ope homeomorphism_imp_open_map homeomorphism_translation translation_galois) show "subspace ((+) (- a) ` U)" by (simp add: \a \ U\ affine_diffs_subspace_subtract \affine U\ cong: image_cong_simp) show "subspace ((+) (- b) ` V)" by (simp add: \b \ V\ affine_diffs_subspace_subtract \affine V\ cong: image_cong_simp) show "continuous_on ((+) (- a) ` S) ((+) (- b) \ f \ (+) a)" by (metis contf continuous_on_compose homeomorphism_cont2 homeomorphism_translation translation_galois) show "((+) (- b) \ f \ (+) a) ` (+) (- a) ` S \ (+) (- b) ` V" using fim by auto show "inj_on ((+) (- b) \ f \ (+) a) ((+) (- a) ` S)" by (auto simp: inj_on_def) (meson inj_onD injf) qed (use \S \ {}\ in auto) then show ?thesis by (metis \a \ U\ \b \ V\ aff_dim_eq_dim affine_hull_eq aff of_nat_le_iff) qed corollary invariance_of_dimension: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and "open S" and injf: "inj_on f S" and "S \ {}" shows "DIM('a) \ DIM('b)" using invariance_of_dimension_subspaces [of UNIV S UNIV f] assms by auto corollary continuous_injective_image_subspace_dim_le: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "subspace S" "subspace T" and contf: "continuous_on S f" and fim: "f ` S \ T" and injf: "inj_on f S" shows "dim S \ dim T" apply (rule invariance_of_dimension_subspaces [of S S _ f]) using assms by (auto simp: subspace_affine) lemma invariance_of_dimension_convex_domain: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "convex S" and contf: "continuous_on S f" and fim: "f ` S \ affine hull T" and injf: "inj_on f S" shows "aff_dim S \ aff_dim T" proof (cases "S = {}") case True then show ?thesis by (simp add: aff_dim_geq) next case False have "aff_dim (affine hull S) \ aff_dim (affine hull T)" proof (rule invariance_of_dimension_affine_sets) show "openin (top_of_set (affine hull S)) (rel_interior S)" by (simp add: openin_rel_interior) show "continuous_on (rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast show "f ` rel_interior S \ affine hull T" using fim rel_interior_subset by blast show "inj_on f (rel_interior S)" using inj_on_subset injf rel_interior_subset by blast show "rel_interior S \ {}" by (simp add: False \convex S\ rel_interior_eq_empty) qed auto then show ?thesis by simp qed lemma homeomorphic_convex_sets_le: assumes "convex S" "S homeomorphic T" shows "aff_dim S \ aff_dim T" proof - obtain h k where homhk: "homeomorphism S T h k" using homeomorphic_def assms by blast show ?thesis proof (rule invariance_of_dimension_convex_domain [OF \convex S\]) show "continuous_on S h" using homeomorphism_def homhk by blast show "h ` S \ affine hull T" by (metis homeomorphism_def homhk hull_subset) show "inj_on h S" by (meson homeomorphism_apply1 homhk inj_on_inverseI) qed qed lemma homeomorphic_convex_sets: assumes "convex S" "convex T" "S homeomorphic T" shows "aff_dim S = aff_dim T" by (meson assms dual_order.antisym homeomorphic_convex_sets_le homeomorphic_sym) lemma homeomorphic_convex_compact_sets_eq: assumes "convex S" "compact S" "convex T" "compact T" shows "S homeomorphic T \ aff_dim S = aff_dim T" by (meson assms homeomorphic_convex_compact_sets homeomorphic_convex_sets) lemma invariance_of_domain_gen: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "inj_on f S" "DIM('b) \ DIM('a)" shows "open(f ` S)" using invariance_of_domain_subspaces [of UNIV S UNIV f] assms by auto lemma injective_into_1d_imp_open_map_UNIV: fixes f :: "'a::euclidean_space \ real" assumes "open T" "continuous_on S f" "inj_on f S" "T \ S" shows "open (f ` T)" apply (rule invariance_of_domain_gen [OF \open T\]) using assms apply (auto simp: elim: continuous_on_subset subset_inj_on) done lemma continuous_on_inverse_open: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "DIM('b) \ DIM('a)" and gf: "\x. x \ S \ g(f x) = x" shows "continuous_on (f ` S) g" proof (clarsimp simp add: continuous_openin_preimage_eq) fix T :: "'a set" assume "open T" have eq: "f ` S \ g -` T = f ` (S \ T)" by (auto simp: gf) show "openin (top_of_set (f ` S)) (f ` S \ g -` T)" apply (subst eq) apply (rule open_openin_trans) apply (rule invariance_of_domain_gen) using assms apply auto using inj_on_inverseI apply auto[1] by (metis \open T\ continuous_on_subset inj_onI inj_on_subset invariance_of_domain_gen openin_open openin_open_eq) qed lemma invariance_of_domain_homeomorphism: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "DIM('b) \ DIM('a)" "inj_on f S" obtains g where "homeomorphism S (f ` S) f g" proof show "homeomorphism S (f ` S) f (inv_into S f)" by (simp add: assms continuous_on_inverse_open homeomorphism_def) qed corollary invariance_of_domain_homeomorphic: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "DIM('b) \ DIM('a)" "inj_on f S" shows "S homeomorphic (f ` S)" using invariance_of_domain_homeomorphism [OF assms] by (meson homeomorphic_def) lemma continuous_image_subset_interior: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "continuous_on S f" "inj_on f S" "DIM('b) \ DIM('a)" shows "f ` (interior S) \ interior(f ` S)" apply (rule interior_maximal) apply (simp add: image_mono interior_subset) apply (rule invariance_of_domain_gen) using assms apply (auto simp: subset_inj_on interior_subset continuous_on_subset) done lemma homeomorphic_interiors_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" and dimeq: "DIM('a) = DIM('b)" shows "(interior S) homeomorphic (interior T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have fim: "f ` interior S \ interior T" using continuous_image_subset_interior [OF contf \inj_on f S\] dimeq fST by simp have gim: "g ` interior T \ interior S" using continuous_image_subset_interior [OF contg \inj_on g T\] dimeq gTS by simp show "homeomorphism (interior S) (interior T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show "\x. x \ interior S \ g (f x) = x" by (meson \\x\S. f x \ T \ g (f x) = x\ subsetD interior_subset) have "interior T \ f ` interior S" proof fix x assume "x \ interior T" then have "g x \ interior S" using gim by blast then show "x \ f ` interior S" by (metis T \x \ interior T\ image_iff interior_subset subsetCE) qed then show "f ` interior S = interior T" using fim by blast show "continuous_on (interior S) f" by (metis interior_subset continuous_on_subset contf) show "\y. y \ interior T \ f (g y) = y" by (meson T subsetD interior_subset) have "interior S \ g ` interior T" proof fix x assume "x \ interior S" then have "f x \ interior T" using fim by blast then show "x \ g ` interior T" by (metis S \x \ interior S\ image_iff interior_subset subsetCE) qed then show "g ` interior T = interior S" using gim by blast show "continuous_on (interior T) g" by (metis interior_subset continuous_on_subset contg) qed qed lemma homeomorphic_open_imp_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "open S" "S \ {}" "open T" "T \ {}" shows "DIM('a) = DIM('b)" using assms apply (simp add: homeomorphic_minimal) apply (rule order_antisym; metis inj_onI invariance_of_dimension) done proposition homeomorphic_interiors: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "interior S = {} \ interior T = {}" shows "(interior S) homeomorphic (interior T)" proof (cases "interior T = {}") case True with assms show ?thesis by auto next case False then have "DIM('a) = DIM('b)" using assms apply (simp add: homeomorphic_minimal) apply (rule order_antisym; metis continuous_on_subset inj_onI inj_on_subset interior_subset invariance_of_dimension open_interior) done then show ?thesis by (rule homeomorphic_interiors_same_dimension [OF \S homeomorphic T\]) qed lemma homeomorphic_frontiers_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "closed S" "closed T" and dimeq: "DIM('a) = DIM('b)" shows "(frontier S) homeomorphic (frontier T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have "g ` interior T \ interior S" using continuous_image_subset_interior [OF contg \inj_on g T\] dimeq gTS by simp then have fim: "f ` frontier S \ frontier T" apply (simp add: frontier_def) using continuous_image_subset_interior assms(2) assms(3) S by auto have "f ` interior S \ interior T" using continuous_image_subset_interior [OF contf \inj_on f S\] dimeq fST by simp then have gim: "g ` frontier T \ frontier S" apply (simp add: frontier_def) using continuous_image_subset_interior T assms(2) assms(3) by auto show "homeomorphism (frontier S) (frontier T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show gf: "\x. x \ frontier S \ g (f x) = x" by (simp add: S assms(2) frontier_def) show fg: "\y. y \ frontier T \ f (g y) = y" by (simp add: T assms(3) frontier_def) have "frontier T \ f ` frontier S" proof fix x assume "x \ frontier T" then have "g x \ frontier S" using gim by blast then show "x \ f ` frontier S" by (metis fg \x \ frontier T\ imageI) qed then show "f ` frontier S = frontier T" using fim by blast show "continuous_on (frontier S) f" by (metis Diff_subset assms(2) closure_eq contf continuous_on_subset frontier_def) have "frontier S \ g ` frontier T" proof fix x assume "x \ frontier S" then have "f x \ frontier T" using fim by blast then show "x \ g ` frontier T" by (metis gf \x \ frontier S\ imageI) qed then show "g ` frontier T = frontier S" using gim by blast show "continuous_on (frontier T) g" by (metis Diff_subset assms(3) closure_closed contg continuous_on_subset frontier_def) qed qed lemma homeomorphic_frontiers: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "closed S" "closed T" "interior S = {} \ interior T = {}" shows "(frontier S) homeomorphic (frontier T)" proof (cases "interior T = {}") case True then show ?thesis by (metis Diff_empty assms closure_eq frontier_def) next case False show ?thesis apply (rule homeomorphic_frontiers_same_dimension) apply (simp_all add: assms) using False assms homeomorphic_interiors homeomorphic_open_imp_same_dimension by blast qed lemma continuous_image_subset_rel_interior: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \ T" and TS: "aff_dim T \ aff_dim S" shows "f ` (rel_interior S) \ rel_interior(f ` S)" proof (rule rel_interior_maximal) show "f ` rel_interior S \ f ` S" by(simp add: image_mono rel_interior_subset) show "openin (top_of_set (affine hull f ` S)) (f ` rel_interior S)" proof (rule invariance_of_domain_affine_sets) show "openin (top_of_set (affine hull S)) (rel_interior S)" by (simp add: openin_rel_interior) show "aff_dim (affine hull f ` S) \ aff_dim (affine hull S)" by (metis aff_dim_affine_hull aff_dim_subset fim TS order_trans) show "f ` rel_interior S \ affine hull f ` S" by (meson \f ` rel_interior S \ f ` S\ hull_subset order_trans) show "continuous_on (rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast show "inj_on f (rel_interior S)" using inj_on_subset injf rel_interior_subset by blast qed auto qed lemma homeomorphic_rel_interiors_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" and aff: "aff_dim S = aff_dim T" shows "(rel_interior S) homeomorphic (rel_interior T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have fim: "f ` rel_interior S \ rel_interior T" by (metis \inj_on f S\ aff contf continuous_image_subset_rel_interior fST order_refl) have gim: "g ` rel_interior T \ rel_interior S" by (metis \inj_on g T\ aff contg continuous_image_subset_rel_interior gTS order_refl) show "homeomorphism (rel_interior S) (rel_interior T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show gf: "\x. x \ rel_interior S \ g (f x) = x" using S rel_interior_subset by blast show fg: "\y. y \ rel_interior T \ f (g y) = y" using T mem_rel_interior_ball by blast have "rel_interior T \ f ` rel_interior S" proof fix x assume "x \ rel_interior T" then have "g x \ rel_interior S" using gim by blast then show "x \ f ` rel_interior S" by (metis fg \x \ rel_interior T\ imageI) qed moreover have "f ` rel_interior S \ rel_interior T" by (metis \inj_on f S\ aff contf continuous_image_subset_rel_interior fST order_refl) ultimately show "f ` rel_interior S = rel_interior T" by blast show "continuous_on (rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast have "rel_interior S \ g ` rel_interior T" proof fix x assume "x \ rel_interior S" then have "f x \ rel_interior T" using fim by blast then show "x \ g ` rel_interior T" by (metis gf \x \ rel_interior S\ imageI) qed then show "g ` rel_interior T = rel_interior S" using gim by blast show "continuous_on (rel_interior T) g" using contg continuous_on_subset rel_interior_subset by blast qed qed lemma homeomorphic_rel_interiors: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "rel_interior S = {} \ rel_interior T = {}" shows "(rel_interior S) homeomorphic (rel_interior T)" proof (cases "rel_interior T = {}") case True with assms show ?thesis by auto next case False obtain f g where S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" using assms [unfolded homeomorphic_minimal] by auto have "aff_dim (affine hull S) \ aff_dim (affine hull T)" apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior S" _ f]) apply (simp_all add: openin_rel_interior False assms) using contf continuous_on_subset rel_interior_subset apply blast apply (meson S hull_subset image_subsetI rel_interior_subset rev_subsetD) apply (metis S inj_on_inverseI inj_on_subset rel_interior_subset) done moreover have "aff_dim (affine hull T) \ aff_dim (affine hull S)" apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior T" _ g]) apply (simp_all add: openin_rel_interior False assms) using contg continuous_on_subset rel_interior_subset apply blast apply (meson T hull_subset image_subsetI rel_interior_subset rev_subsetD) apply (metis T inj_on_inverseI inj_on_subset rel_interior_subset) done ultimately have "aff_dim S = aff_dim T" by force then show ?thesis by (rule homeomorphic_rel_interiors_same_dimension [OF \S homeomorphic T\]) qed lemma homeomorphic_rel_boundaries_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" and aff: "aff_dim S = aff_dim T" shows "(S - rel_interior S) homeomorphic (T - rel_interior T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have fim: "f ` rel_interior S \ rel_interior T" by (metis \inj_on f S\ aff contf continuous_image_subset_rel_interior fST order_refl) have gim: "g ` rel_interior T \ rel_interior S" by (metis \inj_on g T\ aff contg continuous_image_subset_rel_interior gTS order_refl) show "homeomorphism (S - rel_interior S) (T - rel_interior T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show gf: "\x. x \ S - rel_interior S \ g (f x) = x" using S rel_interior_subset by blast show fg: "\y. y \ T - rel_interior T \ f (g y) = y" using T mem_rel_interior_ball by blast show "f ` (S - rel_interior S) = T - rel_interior T" using S fST fim gim by auto show "continuous_on (S - rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast show "g ` (T - rel_interior T) = S - rel_interior S" using T gTS gim fim by auto show "continuous_on (T - rel_interior T) g" using contg continuous_on_subset rel_interior_subset by blast qed qed lemma homeomorphic_rel_boundaries: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "rel_interior S = {} \ rel_interior T = {}" shows "(S - rel_interior S) homeomorphic (T - rel_interior T)" proof (cases "rel_interior T = {}") case True with assms show ?thesis by auto next case False obtain f g where S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" using assms [unfolded homeomorphic_minimal] by auto have "aff_dim (affine hull S) \ aff_dim (affine hull T)" apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior S" _ f]) apply (simp_all add: openin_rel_interior False assms) using contf continuous_on_subset rel_interior_subset apply blast apply (meson S hull_subset image_subsetI rel_interior_subset rev_subsetD) apply (metis S inj_on_inverseI inj_on_subset rel_interior_subset) done moreover have "aff_dim (affine hull T) \ aff_dim (affine hull S)" apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior T" _ g]) apply (simp_all add: openin_rel_interior False assms) using contg continuous_on_subset rel_interior_subset apply blast apply (meson T hull_subset image_subsetI rel_interior_subset rev_subsetD) apply (metis T inj_on_inverseI inj_on_subset rel_interior_subset) done ultimately have "aff_dim S = aff_dim T" by force then show ?thesis by (rule homeomorphic_rel_boundaries_same_dimension [OF \S homeomorphic T\]) qed proposition uniformly_continuous_homeomorphism_UNIV_trivial: fixes f :: "'a::euclidean_space \ 'a" assumes contf: "uniformly_continuous_on S f" and hom: "homeomorphism S UNIV f g" shows "S = UNIV" proof (cases "S = {}") case True then show ?thesis by (metis UNIV_I hom empty_iff homeomorphism_def image_eqI) next case False have "inj g" by (metis UNIV_I hom homeomorphism_apply2 injI) then have "open (g ` UNIV)" by (blast intro: invariance_of_domain hom homeomorphism_cont2) then have "open S" using hom homeomorphism_image2 by blast moreover have "complete S" unfolding complete_def proof clarify fix \ assume \: "\n. \ n \ S" and "Cauchy \" have "Cauchy (f o \)" using uniformly_continuous_imp_Cauchy_continuous \Cauchy \\ \ contf by blast then obtain l where "(f \ \) \ l" by (auto simp: convergent_eq_Cauchy [symmetric]) show "\l\S. \ \ l" proof show "g l \ S" using hom homeomorphism_image2 by blast have "(g \ (f \ \)) \ g l" by (meson UNIV_I \(f \ \) \ l\ continuous_on_sequentially hom homeomorphism_cont2) then show "\ \ g l" proof - have "\n. \ n = (g \ (f \ \)) n" by (metis (no_types) \ comp_eq_dest_lhs hom homeomorphism_apply1) then show ?thesis by (metis (no_types) LIMSEQ_iff \(g \ (f \ \)) \ g l\) qed qed qed then have "closed S" by (simp add: complete_eq_closed) ultimately show ?thesis using clopen [of S] False by simp qed subsection\Formulation of loop homotopy in terms of maps out of type complex\ lemma homotopic_circlemaps_imp_homotopic_loops: assumes "homotopic_with_canon (\h. True) (sphere 0 1) S f g" shows "homotopic_loops S (f \ exp \ (\t. 2 * of_real pi * of_real t * \)) (g \ exp \ (\t. 2 * of_real pi * of_real t * \))" proof - have "homotopic_with_canon (\f. True) {z. cmod z = 1} S f g" using assms by (auto simp: sphere_def) moreover have "continuous_on {0..1} (exp \ (\t. 2 * of_real pi * of_real t * \))" by (intro continuous_intros) moreover have "(exp \ (\t. 2 * of_real pi * of_real t * \)) ` {0..1} \ {z. cmod z = 1}" by (auto simp: norm_mult) ultimately show ?thesis apply (simp add: homotopic_loops_def comp_assoc) apply (rule homotopic_with_compose_continuous_right) apply (auto simp: pathstart_def pathfinish_def) done qed lemma homotopic_loops_imp_homotopic_circlemaps: assumes "homotopic_loops S p q" shows "homotopic_with_canon (\h. True) (sphere 0 1) S (p \ (\z. (Arg2pi z / (2 * pi)))) (q \ (\z. (Arg2pi z / (2 * pi))))" proof - obtain h where conth: "continuous_on ({0..1::real} \ {0..1}) h" and him: "h ` ({0..1} \ {0..1}) \ S" and h0: "(\x. h (0, x) = p x)" and h1: "(\x. h (1, x) = q x)" and h01: "(\t\{0..1}. h (t, 1) = h (t, 0)) " using assms by (auto simp: homotopic_loops_def sphere_def homotopic_with_def pathstart_def pathfinish_def) define j where "j \ \z. if 0 \ Im (snd z) then h (fst z, Arg2pi (snd z) / (2 * pi)) else h (fst z, 1 - Arg2pi (cnj (snd z)) / (2 * pi))" have Arg2pi_eq: "1 - Arg2pi (cnj y) / (2 * pi) = Arg2pi y / (2 * pi) \ Arg2pi y = 0 \ Arg2pi (cnj y) = 0" if "cmod y = 1" for y using that Arg2pi_eq_0_pi Arg2pi_eq_pi by (force simp: Arg2pi_cnj field_split_simps) show ?thesis proof (simp add: homotopic_with; intro conjI ballI exI) show "continuous_on ({0..1} \ sphere 0 1) (\w. h (fst w, Arg2pi (snd w) / (2 * pi)))" proof (rule continuous_on_eq) show j: "j x = h (fst x, Arg2pi (snd x) / (2 * pi))" if "x \ {0..1} \ sphere 0 1" for x using Arg2pi_eq that h01 by (force simp: j_def) have eq: "S = S \ (UNIV \ {z. 0 \ Im z}) \ S \ (UNIV \ {z. Im z \ 0})" for S :: "(real*complex)set" by auto have c1: "continuous_on ({0..1} \ sphere 0 1 \ UNIV \ {z. 0 \ Im z}) (\x. h (fst x, Arg2pi (snd x) / (2 * pi)))" apply (intro continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF continuous_on_upperhalf_Arg2pi]) apply (auto simp: Arg2pi) apply (meson Arg2pi_lt_2pi linear not_le) done have c2: "continuous_on ({0..1} \ sphere 0 1 \ UNIV \ {z. Im z \ 0}) (\x. h (fst x, 1 - Arg2pi (cnj (snd x)) / (2 * pi)))" apply (intro continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF continuous_on_upperhalf_Arg2pi]) apply (auto simp: Arg2pi) apply (meson Arg2pi_lt_2pi linear not_le) done show "continuous_on ({0..1} \ sphere 0 1) j" apply (simp add: j_def) apply (subst eq) apply (rule continuous_on_cases_local) apply (simp_all add: eq [symmetric] closedin_closed_Int closed_Times closed_halfspace_Im_le closed_halfspace_Im_ge c1 c2) using Arg2pi_eq h01 by force qed have "(\w. h (fst w, Arg2pi (snd w) / (2 * pi))) ` ({0..1} \ sphere 0 1) \ h ` ({0..1} \ {0..1})" by (auto simp: Arg2pi_ge_0 Arg2pi_lt_2pi less_imp_le) also have "... \ S" using him by blast finally show "(\w. h (fst w, Arg2pi (snd w) / (2 * pi))) ` ({0..1} \ sphere 0 1) \ S" . qed (auto simp: h0 h1) qed lemma simply_connected_homotopic_loops: "simply_connected S \ (\p q. homotopic_loops S p p \ homotopic_loops S q q \ homotopic_loops S p q)" unfolding simply_connected_def using homotopic_loops_refl by metis lemma simply_connected_eq_homotopic_circlemaps1: fixes f :: "complex \ 'a::topological_space" and g :: "complex \ 'a" assumes S: "simply_connected S" and contf: "continuous_on (sphere 0 1) f" and fim: "f ` (sphere 0 1) \ S" and contg: "continuous_on (sphere 0 1) g" and gim: "g ` (sphere 0 1) \ S" shows "homotopic_with_canon (\h. True) (sphere 0 1) S f g" proof - have "homotopic_loops S (f \ exp \ (\t. of_real(2 * pi * t) * \)) (g \ exp \ (\t. of_real(2 * pi * t) * \))" apply (rule S [unfolded simply_connected_homotopic_loops, rule_format]) apply (simp add: homotopic_circlemaps_imp_homotopic_loops contf fim contg gim) done then show ?thesis apply (rule homotopic_with_eq [OF homotopic_loops_imp_homotopic_circlemaps]) apply (auto simp: o_def complex_norm_eq_1_exp mult.commute) done qed lemma simply_connected_eq_homotopic_circlemaps2a: fixes h :: "complex \ 'a::topological_space" assumes conth: "continuous_on (sphere 0 1) h" and him: "h ` (sphere 0 1) \ S" and hom: "\f g::complex \ 'a. \continuous_on (sphere 0 1) f; f ` (sphere 0 1) \ S; continuous_on (sphere 0 1) g; g ` (sphere 0 1) \ S\ \ homotopic_with_canon (\h. True) (sphere 0 1) S f g" shows "\a. homotopic_with_canon (\h. True) (sphere 0 1) S h (\x. a)" apply (rule_tac x="h 1" in exI) apply (rule hom) using assms by (auto) lemma simply_connected_eq_homotopic_circlemaps2b: fixes S :: "'a::real_normed_vector set" assumes "\f g::complex \ 'a. \continuous_on (sphere 0 1) f; f ` (sphere 0 1) \ S; continuous_on (sphere 0 1) g; g ` (sphere 0 1) \ S\ \ homotopic_with_canon (\h. True) (sphere 0 1) S f g" shows "path_connected S" proof (clarsimp simp add: path_connected_eq_homotopic_points) fix a b assume "a \ S" "b \ S" then show "homotopic_loops S (linepath a a) (linepath b b)" using homotopic_circlemaps_imp_homotopic_loops [OF assms [of "\x. a" "\x. b"]] by (auto simp: o_def linepath_def) qed lemma simply_connected_eq_homotopic_circlemaps3: fixes h :: "complex \ 'a::real_normed_vector" assumes "path_connected S" and hom: "\f::complex \ 'a. \continuous_on (sphere 0 1) f; f `(sphere 0 1) \ S\ \ \a. homotopic_with_canon (\h. True) (sphere 0 1) S f (\x. a)" shows "simply_connected S" proof (clarsimp simp add: simply_connected_eq_contractible_loop_some assms) fix p assume p: "path p" "path_image p \ S" "pathfinish p = pathstart p" then have "homotopic_loops S p p" by (simp add: homotopic_loops_refl) then obtain a where homp: "homotopic_with_canon (\h. True) (sphere 0 1) S (p \ (\z. Arg2pi z / (2 * pi))) (\x. a)" by (metis homotopic_with_imp_subset2 homotopic_loops_imp_homotopic_circlemaps homotopic_with_imp_continuous hom) show "\a. a \ S \ homotopic_loops S p (linepath a a)" proof (intro exI conjI) show "a \ S" using homotopic_with_imp_subset2 [OF homp] by (metis dist_0_norm image_subset_iff mem_sphere norm_one) have teq: "\t. \0 \ t; t \ 1\ \ t = Arg2pi (exp (2 * of_real pi * of_real t * \)) / (2 * pi) \ t=1 \ Arg2pi (exp (2 * of_real pi * of_real t * \)) = 0" apply (rule disjCI) using Arg2pi_of_real [of 1] apply (auto simp: Arg2pi_exp) done have "homotopic_loops S p (p \ (\z. Arg2pi z / (2 * pi)) \ exp \ (\t. 2 * complex_of_real pi * complex_of_real t * \))" apply (rule homotopic_loops_eq [OF p]) using p teq apply (fastforce simp: pathfinish_def pathstart_def) done then show "homotopic_loops S p (linepath a a)" by (simp add: linepath_refl homotopic_loops_trans [OF _ homotopic_circlemaps_imp_homotopic_loops [OF homp, simplified K_record_comp]]) qed qed proposition simply_connected_eq_homotopic_circlemaps: fixes S :: "'a::real_normed_vector set" shows "simply_connected S \ (\f g::complex \ 'a. continuous_on (sphere 0 1) f \ f ` (sphere 0 1) \ S \ continuous_on (sphere 0 1) g \ g ` (sphere 0 1) \ S \ homotopic_with_canon (\h. True) (sphere 0 1) S f g)" apply (rule iffI) apply (blast elim: dest: simply_connected_eq_homotopic_circlemaps1) by (simp add: simply_connected_eq_homotopic_circlemaps2a simply_connected_eq_homotopic_circlemaps2b simply_connected_eq_homotopic_circlemaps3) proposition simply_connected_eq_contractible_circlemap: fixes S :: "'a::real_normed_vector set" shows "simply_connected S \ path_connected S \ (\f::complex \ 'a. continuous_on (sphere 0 1) f \ f `(sphere 0 1) \ S \ (\a. homotopic_with_canon (\h. True) (sphere 0 1) S f (\x. a)))" apply (rule iffI) apply (simp add: simply_connected_eq_homotopic_circlemaps1 simply_connected_eq_homotopic_circlemaps2a simply_connected_eq_homotopic_circlemaps2b) using simply_connected_eq_homotopic_circlemaps3 by blast corollary homotopy_eqv_simple_connectedness: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" shows "S homotopy_eqv T \ simply_connected S \ simply_connected T" by (simp add: simply_connected_eq_homotopic_circlemaps homotopy_eqv_homotopic_triviality) subsection\Homeomorphism of simple closed curves to circles\ proposition homeomorphic_simple_path_image_circle: fixes a :: complex and \ :: "real \ 'a::t2_space" assumes "simple_path \" and loop: "pathfinish \ = pathstart \" and "0 < r" shows "(path_image \) homeomorphic sphere a r" proof - have "homotopic_loops (path_image \) \ \" by (simp add: assms homotopic_loops_refl simple_path_imp_path) then have hom: "homotopic_with_canon (\h. True) (sphere 0 1) (path_image \) (\ \ (\z. Arg2pi z / (2*pi))) (\ \ (\z. Arg2pi z / (2*pi)))" by (rule homotopic_loops_imp_homotopic_circlemaps) have "\g. homeomorphism (sphere 0 1) (path_image \) (\ \ (\z. Arg2pi z / (2*pi))) g" proof (rule homeomorphism_compact) show "continuous_on (sphere 0 1) (\ \ (\z. Arg2pi z / (2*pi)))" using hom homotopic_with_imp_continuous by blast show "inj_on (\ \ (\z. Arg2pi z / (2*pi))) (sphere 0 1)" proof fix x y assume xy: "x \ sphere 0 1" "y \ sphere 0 1" and eq: "(\ \ (\z. Arg2pi z / (2*pi))) x = (\ \ (\z. Arg2pi z / (2*pi))) y" then have "(Arg2pi x / (2*pi)) = (Arg2pi y / (2*pi))" proof - have "(Arg2pi x / (2*pi)) \ {0..1}" "(Arg2pi y / (2*pi)) \ {0..1}" using Arg2pi_ge_0 Arg2pi_lt_2pi dual_order.strict_iff_order by fastforce+ with eq show ?thesis using \simple_path \\ Arg2pi_lt_2pi unfolding simple_path_def o_def by (metis eq_divide_eq_1 not_less_iff_gr_or_eq) qed with xy show "x = y" by (metis is_Arg_def Arg2pi Arg2pi_0 dist_0_norm divide_cancel_right dual_order.strict_iff_order mem_sphere) qed have "\z. cmod z = 1 \ \x\{0..1}. \ (Arg2pi z / (2*pi)) = \ x" by (metis Arg2pi_ge_0 Arg2pi_lt_2pi atLeastAtMost_iff divide_less_eq_1 less_eq_real_def zero_less_mult_iff pi_gt_zero zero_le_divide_iff zero_less_numeral) moreover have "\z\sphere 0 1. \ x = \ (Arg2pi z / (2*pi))" if "0 \ x" "x \ 1" for x proof (cases "x=1") case True with Arg2pi_of_real [of 1] loop show ?thesis by (rule_tac x=1 in bexI) (auto simp: pathfinish_def pathstart_def \0 \ x\) next case False then have *: "(Arg2pi (exp (\*(2* of_real pi* of_real x))) / (2*pi)) = x" using that by (auto simp: Arg2pi_exp field_split_simps) show ?thesis by (rule_tac x="exp(\ * of_real(2*pi*x))" in bexI) (auto simp: *) qed ultimately show "(\ \ (\z. Arg2pi z / (2*pi))) ` sphere 0 1 = path_image \" by (auto simp: path_image_def image_iff) qed auto then have "path_image \ homeomorphic sphere (0::complex) 1" using homeomorphic_def homeomorphic_sym by blast also have "... homeomorphic sphere a r" by (simp add: assms homeomorphic_spheres) finally show ?thesis . qed lemma homeomorphic_simple_path_images: fixes \1 :: "real \ 'a::t2_space" and \2 :: "real \ 'b::t2_space" assumes "simple_path \1" and loop: "pathfinish \1 = pathstart \1" assumes "simple_path \2" and loop: "pathfinish \2 = pathstart \2" shows "(path_image \1) homeomorphic (path_image \2)" by (meson assms homeomorphic_simple_path_image_circle homeomorphic_sym homeomorphic_trans loop pi_gt_zero) subsection\Dimension-based conditions for various homeomorphisms\ lemma homeomorphic_subspaces_eq: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "subspace S" "subspace T" shows "S homeomorphic T \ dim S = dim T" proof assume "S homeomorphic T" then obtain f g where hom: "homeomorphism S T f g" using homeomorphic_def by blast show "dim S = dim T" proof (rule order_antisym) show "dim S \ dim T" by (metis assms dual_order.refl inj_onI homeomorphism_cont1 [OF hom] homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] continuous_injective_image_subspace_dim_le) show "dim T \ dim S" by (metis assms dual_order.refl inj_onI homeomorphism_cont2 [OF hom] homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] continuous_injective_image_subspace_dim_le) qed next assume "dim S = dim T" then show "S homeomorphic T" by (simp add: assms homeomorphic_subspaces) qed lemma homeomorphic_affine_sets_eq: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "affine S" "affine T" shows "S homeomorphic T \ aff_dim S = aff_dim T" proof (cases "S = {} \ T = {}") case True then show ?thesis using assms homeomorphic_affine_sets by force next case False then obtain a b where "a \ S" "b \ T" by blast then have "subspace ((+) (- a) ` S)" "subspace ((+) (- b) ` T)" using affine_diffs_subspace assms by blast+ then show ?thesis by (metis affine_imp_convex assms homeomorphic_affine_sets homeomorphic_convex_sets) qed lemma homeomorphic_hyperplanes_eq: fixes a :: "'a::euclidean_space" and c :: "'b::euclidean_space" assumes "a \ 0" "c \ 0" shows "({x. a \ x = b} homeomorphic {x. c \ x = d} \ DIM('a) = DIM('b))" apply (auto simp: homeomorphic_affine_sets_eq affine_hyperplane assms) by (metis DIM_positive Suc_pred) lemma homeomorphic_UNIV_UNIV: shows "(UNIV::'a set) homeomorphic (UNIV::'b set) \ DIM('a::euclidean_space) = DIM('b::euclidean_space)" by (simp add: homeomorphic_subspaces_eq) lemma simply_connected_sphere_gen: assumes "convex S" "bounded S" and 3: "3 \ aff_dim S" shows "simply_connected(rel_frontier S)" proof - have pa: "path_connected (rel_frontier S)" using assms by (simp add: path_connected_sphere_gen) show ?thesis proof (clarsimp simp add: simply_connected_eq_contractible_circlemap pa) fix f assume f: "continuous_on (sphere (0::complex) 1) f" "f ` sphere 0 1 \ rel_frontier S" have eq: "sphere (0::complex) 1 = rel_frontier(cball 0 1)" by simp have "convex (cball (0::complex) 1)" by (rule convex_cball) then obtain c where "homotopic_with_canon (\z. True) (sphere (0::complex) 1) (rel_frontier S) f (\x. c)" apply (rule inessential_spheremap_lowdim_gen [OF _ bounded_cball \convex S\ \bounded S\, where f=f]) using f 3 apply (auto simp: aff_dim_cball) done then show "\a. homotopic_with_canon (\h. True) (sphere 0 1) (rel_frontier S) f (\x. a)" by blast qed qed subsection\more invariance of domain\(*FIX ME title? *) proposition invariance_of_domain_sphere_affine_set_gen: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \ T" and U: "bounded U" "convex U" and "affine T" and affTU: "aff_dim T < aff_dim U" and ope: "openin (top_of_set (rel_frontier U)) S" shows "openin (top_of_set T) (f ` S)" proof (cases "rel_frontier U = {}") case True then show ?thesis using ope openin_subset by force next case False obtain b c where b: "b \ rel_frontier U" and c: "c \ rel_frontier U" and "b \ c" using \bounded U\ rel_frontier_not_sing [of U] subset_singletonD False by fastforce obtain V :: "'a set" where "affine V" and affV: "aff_dim V = aff_dim U - 1" proof (rule choose_affine_subset [OF affine_UNIV]) show "- 1 \ aff_dim U - 1" by (metis aff_dim_empty aff_dim_geq aff_dim_negative_iff affTU diff_0 diff_right_mono not_le) show "aff_dim U - 1 \ aff_dim (UNIV::'a set)" by (metis aff_dim_UNIV aff_dim_le_DIM le_cases not_le zle_diff1_eq) qed auto have SU: "S \ rel_frontier U" using ope openin_imp_subset by auto have homb: "rel_frontier U - {b} homeomorphic V" and homc: "rel_frontier U - {c} homeomorphic V" using homeomorphic_punctured_sphere_affine_gen [of U _ V] by (simp_all add: \affine V\ affV U b c) then obtain g h j k where gh: "homeomorphism (rel_frontier U - {b}) V g h" and jk: "homeomorphism (rel_frontier U - {c}) V j k" by (auto simp: homeomorphic_def) with SU have hgsub: "(h ` g ` (S - {b})) \ S" and kjsub: "(k ` j ` (S - {c})) \ S" by (simp_all add: homeomorphism_def subset_eq) have [simp]: "aff_dim T \ aff_dim V" by (simp add: affTU affV) have "openin (top_of_set T) ((f \ h) ` g ` (S - {b}))" proof (rule invariance_of_domain_affine_sets [OF _ \affine V\]) show "openin (top_of_set V) (g ` (S - {b}))" apply (rule homeomorphism_imp_open_map [OF gh]) by (meson Diff_mono Diff_subset SU ope openin_delete openin_subset_trans order_refl) show "continuous_on (g ` (S - {b})) (f \ h)" apply (rule continuous_on_compose) apply (meson Diff_mono SU homeomorphism_def homeomorphism_of_subsets gh set_eq_subset) using contf continuous_on_subset hgsub by blast show "inj_on (f \ h) (g ` (S - {b}))" using kjsub apply (clarsimp simp add: inj_on_def) by (metis SU b homeomorphism_def inj_onD injf insert_Diff insert_iff gh rev_subsetD) show "(f \ h) ` g ` (S - {b}) \ T" by (metis fim image_comp image_mono hgsub subset_trans) qed (auto simp: assms) moreover have "openin (top_of_set T) ((f \ k) ` j ` (S - {c}))" proof (rule invariance_of_domain_affine_sets [OF _ \affine V\]) show "openin (top_of_set V) (j ` (S - {c}))" apply (rule homeomorphism_imp_open_map [OF jk]) by (meson Diff_mono Diff_subset SU ope openin_delete openin_subset_trans order_refl) show "continuous_on (j ` (S - {c})) (f \ k)" apply (rule continuous_on_compose) apply (meson Diff_mono SU homeomorphism_def homeomorphism_of_subsets jk set_eq_subset) using contf continuous_on_subset kjsub by blast show "inj_on (f \ k) (j ` (S - {c}))" using kjsub apply (clarsimp simp add: inj_on_def) by (metis SU c homeomorphism_def inj_onD injf insert_Diff insert_iff jk rev_subsetD) show "(f \ k) ` j ` (S - {c}) \ T" by (metis fim image_comp image_mono kjsub subset_trans) qed (auto simp: assms) ultimately have "openin (top_of_set T) ((f \ h) ` g ` (S - {b}) \ ((f \ k) ` j ` (S - {c})))" by (rule openin_Un) moreover have "(f \ h) ` g ` (S - {b}) = f ` (S - {b})" proof - have "h ` g ` (S - {b}) = (S - {b})" proof show "h ` g ` (S - {b}) \ S - {b}" using homeomorphism_apply1 [OF gh] SU by (fastforce simp add: image_iff image_subset_iff) show "S - {b} \ h ` g ` (S - {b})" apply clarify by (metis SU subsetD homeomorphism_apply1 [OF gh] image_iff member_remove remove_def) qed then show ?thesis by (metis image_comp) qed moreover have "(f \ k) ` j ` (S - {c}) = f ` (S - {c})" proof - have "k ` j ` (S - {c}) = (S - {c})" proof show "k ` j ` (S - {c}) \ S - {c}" using homeomorphism_apply1 [OF jk] SU by (fastforce simp add: image_iff image_subset_iff) show "S - {c} \ k ` j ` (S - {c})" apply clarify by (metis SU subsetD homeomorphism_apply1 [OF jk] image_iff member_remove remove_def) qed then show ?thesis by (metis image_comp) qed moreover have "f ` (S - {b}) \ f ` (S - {c}) = f ` (S)" using \b \ c\ by blast ultimately show ?thesis by simp qed lemma invariance_of_domain_sphere_affine_set: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \ T" and "r \ 0" "affine T" and affTU: "aff_dim T < DIM('a)" and ope: "openin (top_of_set (sphere a r)) S" shows "openin (top_of_set T) (f ` S)" proof (cases "sphere a r = {}") case True then show ?thesis using ope openin_subset by force next case False show ?thesis proof (rule invariance_of_domain_sphere_affine_set_gen [OF contf injf fim bounded_cball convex_cball \affine T\]) show "aff_dim T < aff_dim (cball a r)" by (metis False affTU aff_dim_cball assms(4) linorder_cases sphere_empty) show "openin (top_of_set (rel_frontier (cball a r))) S" by (simp add: \r \ 0\ ope) qed qed lemma no_embedding_sphere_lowdim: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on (sphere a r) f" and injf: "inj_on f (sphere a r)" and "r > 0" shows "DIM('a) \ DIM('b)" proof - have "False" if "DIM('a) > DIM('b)" proof - have "compact (f ` sphere a r)" using compact_continuous_image by (simp add: compact_continuous_image contf) then have "\ open (f ` sphere a r)" using compact_open by (metis assms(3) image_is_empty not_less_iff_gr_or_eq sphere_eq_empty) then show False using invariance_of_domain_sphere_affine_set [OF contf injf subset_UNIV] \r > 0\ by (metis aff_dim_UNIV affine_UNIV less_irrefl of_nat_less_iff open_openin openin_subtopology_self subtopology_UNIV that) qed then show ?thesis using not_less by blast qed lemma simply_connected_sphere: fixes a :: "'a::euclidean_space" assumes "3 \ DIM('a)" shows "simply_connected(sphere a r)" proof (cases rule: linorder_cases [of r 0]) case less then show ?thesis by simp next case equal then show ?thesis by (auto simp: convex_imp_simply_connected) next case greater then show ?thesis using simply_connected_sphere_gen [of "cball a r"] assms by (simp add: aff_dim_cball) qed lemma simply_connected_sphere_eq: fixes a :: "'a::euclidean_space" shows "simply_connected(sphere a r) \ 3 \ DIM('a) \ r \ 0" (is "?lhs = ?rhs") proof (cases "r \ 0") case True have "simply_connected (sphere a r)" apply (rule convex_imp_simply_connected) using True less_eq_real_def by auto with True show ?thesis by auto next case False show ?thesis proof assume L: ?lhs have "False" if "DIM('a) = 1 \ DIM('a) = 2" using that proof assume "DIM('a) = 1" with L show False using connected_sphere_eq simply_connected_imp_connected by (metis False Suc_1 not_less_eq_eq order_refl) next assume "DIM('a) = 2" then have "sphere a r homeomorphic sphere (0::complex) 1" by (metis DIM_complex False homeomorphic_spheres_gen not_less zero_less_one) then have "simply_connected(sphere (0::complex) 1)" using L homeomorphic_simply_connected_eq by blast then obtain a::complex where "homotopic_with_canon (\h. True) (sphere 0 1) (sphere 0 1) id (\x. a)" apply (simp add: simply_connected_eq_contractible_circlemap) by (metis continuous_on_id' id_apply image_id subset_refl) then show False using contractible_sphere contractible_def not_one_le_zero by blast qed with False show ?rhs apply simp by (metis DIM_ge_Suc0 le_antisym not_less_eq_eq numeral_2_eq_2 numeral_3_eq_3) next assume ?rhs with False show ?lhs by (simp add: simply_connected_sphere) qed qed lemma simply_connected_punctured_universe_eq: fixes a :: "'a::euclidean_space" shows "simply_connected(- {a}) \ 3 \ DIM('a)" proof - have [simp]: "a \ rel_interior (cball a 1)" by (simp add: rel_interior_nonempty_interior) have [simp]: "affine hull cball a 1 - {a} = -{a}" by (metis Compl_eq_Diff_UNIV aff_dim_cball aff_dim_lt_full not_less_iff_gr_or_eq zero_less_one) have "simply_connected(- {a}) \ simply_connected(sphere a 1)" apply (rule sym) apply (rule homotopy_eqv_simple_connectedness) using homotopy_eqv_rel_frontier_punctured_affine_hull [of "cball a 1" a] apply auto done also have "... \ 3 \ DIM('a)" by (simp add: simply_connected_sphere_eq) finally show ?thesis . qed lemma not_simply_connected_circle: fixes a :: complex shows "0 < r \ \ simply_connected(sphere a r)" by (simp add: simply_connected_sphere_eq) proposition simply_connected_punctured_convex: fixes a :: "'a::euclidean_space" assumes "convex S" and 3: "3 \ aff_dim S" shows "simply_connected(S - {a})" proof (cases "a \ rel_interior S") case True then obtain e where "a \ S" "0 < e" and e: "cball a e \ affine hull S \ S" by (auto simp: rel_interior_cball) have con: "convex (cball a e \ affine hull S)" by (simp add: convex_Int) have bo: "bounded (cball a e \ affine hull S)" by (simp add: bounded_Int) have "affine hull S \ interior (cball a e) \ {}" using \0 < e\ \a \ S\ hull_subset by fastforce then have "3 \ aff_dim (affine hull S \ cball a e)" by (simp add: 3 aff_dim_convex_Int_nonempty_interior [OF convex_affine_hull]) also have "... = aff_dim (cball a e \ affine hull S)" by (simp add: Int_commute) finally have "3 \ aff_dim (cball a e \ affine hull S)" . moreover have "rel_frontier (cball a e \ affine hull S) homotopy_eqv S - {a}" proof (rule homotopy_eqv_rel_frontier_punctured_convex) show "a \ rel_interior (cball a e \ affine hull S)" by (meson IntI Int_mono \a \ S\ \0 < e\ e \cball a e \ affine hull S \ S\ ball_subset_cball centre_in_cball dual_order.strict_implies_order hull_inc hull_mono mem_rel_interior_ball) have "closed (cball a e \ affine hull S)" by blast then show "rel_frontier (cball a e \ affine hull S) \ S" apply (simp add: rel_frontier_def) using e by blast show "S \ affine hull (cball a e \ affine hull S)" by (metis (no_types, lifting) IntI \a \ S\ \0 < e\ affine_hull_convex_Int_nonempty_interior centre_in_ball convex_affine_hull empty_iff hull_subset inf_commute interior_cball subsetCE subsetI) qed (auto simp: assms con bo) ultimately show ?thesis using homotopy_eqv_simple_connectedness simply_connected_sphere_gen [OF con bo] by blast next case False show ?thesis apply (rule contractible_imp_simply_connected) apply (rule contractible_convex_tweak_boundary_points [OF \convex S\]) apply (simp add: False rel_interior_subset subset_Diff_insert) by (meson Diff_subset closure_subset subset_trans) qed corollary simply_connected_punctured_universe: fixes a :: "'a::euclidean_space" assumes "3 \ DIM('a)" shows "simply_connected(- {a})" proof - have [simp]: "affine hull cball a 1 = UNIV" apply auto by (metis UNIV_I aff_dim_cball aff_dim_lt_full zero_less_one not_less_iff_gr_or_eq) have "simply_connected (rel_frontier (cball a 1)) = simply_connected (affine hull cball a 1 - {a})" apply (rule homotopy_eqv_simple_connectedness) apply (rule homotopy_eqv_rel_frontier_punctured_affine_hull) apply (force simp: rel_interior_cball intro: homotopy_eqv_simple_connectedness homotopy_eqv_rel_frontier_punctured_affine_hull)+ done then show ?thesis using simply_connected_sphere [of a 1, OF assms] by (auto simp: Compl_eq_Diff_UNIV) qed subsection\The power, squaring and exponential functions as covering maps\ proposition covering_space_power_punctured_plane: assumes "0 < n" shows "covering_space (- {0}) (\z::complex. z^n) (- {0})" proof - consider "n = 1" | "2 \ n" using assms by linarith then obtain e where "0 < e" and e: "\w z. cmod(w - z) < e * cmod z \ (w^n = z^n \ w = z)" proof cases assume "n = 1" then show ?thesis by (rule_tac e=1 in that) auto next assume "2 \ n" have eq_if_pow_eq: "w = z" if lt: "cmod (w - z) < 2 * sin (pi / real n) * cmod z" and eq: "w^n = z^n" for w z proof (cases "z = 0") case True with eq assms show ?thesis by (auto simp: power_0_left) next case False then have "z \ 0" by auto have "(w/z)^n = 1" by (metis False divide_self_if eq power_divide power_one) then obtain j where j: "w / z = exp (2 * of_real pi * \ * j / n)" and "j < n" using Suc_leI assms \2 \ n\ complex_roots_unity [THEN eqset_imp_iff, of n "w/z"] by force have "cmod (w/z - 1) < 2 * sin (pi / real n)" using lt assms \z \ 0\ by (simp add: field_split_simps norm_divide) then have "cmod (exp (\ * of_real (2 * pi * j / n)) - 1) < 2 * sin (pi / real n)" by (simp add: j field_simps) then have "2 * \sin((2 * pi * j / n) / 2)\ < 2 * sin (pi / real n)" by (simp only: dist_exp_i_1) then have sin_less: "sin((pi * j / n)) < sin (pi / real n)" by (simp add: field_simps) then have "w / z = 1" proof (cases "j = 0") case True then show ?thesis by (auto simp: j) next case False then have "sin (pi / real n) \ sin((pi * j / n))" proof (cases "j / n \ 1/2") case True show ?thesis apply (rule sin_monotone_2pi_le) using \j \ 0 \ \j < n\ True apply (auto simp: field_simps intro: order_trans [of _ 0]) done next case False then have seq: "sin(pi * j / n) = sin(pi * (n - j) / n)" using \j < n\ by (simp add: algebra_simps diff_divide_distrib of_nat_diff) show ?thesis apply (simp only: seq) apply (rule sin_monotone_2pi_le) using \j < n\ False apply (auto simp: field_simps intro: order_trans [of _ 0]) done qed with sin_less show ?thesis by force qed then show ?thesis by simp qed show ?thesis apply (rule_tac e = "2 * sin(pi / n)" in that) apply (force simp: \2 \ n\ sin_pi_divide_n_gt_0) apply (meson eq_if_pow_eq) done qed have zn1: "continuous_on (- {0}) (\z::complex. z^n)" by (rule continuous_intros)+ have zn2: "(\z::complex. z^n) ` (- {0}) = - {0}" using assms by (auto simp: image_def elim: exists_complex_root_nonzero [where n = n]) have zn3: "\T. z^n \ T \ open T \ 0 \ T \ (\v. \v = -{0} \ (\z. z ^ n) -` T \ (\u\v. open u \ 0 \ u) \ pairwise disjnt v \ (\u\v. Ex (homeomorphism u T (\z. z^n))))" if "z \ 0" for z::complex proof - define d where "d \ min (1/2) (e/4) * norm z" have "0 < d" by (simp add: d_def \0 < e\ \z \ 0\) have iff_x_eq_y: "x^n = y^n \ x = y" if eq: "w^n = z^n" and x: "x \ ball w d" and y: "y \ ball w d" for w x y proof - have [simp]: "norm z = norm w" using that by (simp add: assms power_eq_imp_eq_norm) show ?thesis proof (cases "w = 0") case True with \z \ 0\ assms eq show ?thesis by (auto simp: power_0_left) next case False have "cmod (x - y) < 2*d" using x y by (simp add: dist_norm [symmetric]) (metis dist_commute mult_2 dist_triangle_less_add) also have "... \ 2 * e / 4 * norm w" using \e > 0\ by (simp add: d_def min_mult_distrib_right) also have "... = e * (cmod w / 2)" by simp also have "... \ e * cmod y" apply (rule mult_left_mono) using \e > 0\ y apply (simp_all add: dist_norm d_def min_mult_distrib_right del: divide_const_simps) apply (metis dist_0_norm dist_complex_def dist_triangle_half_l linorder_not_less order_less_irrefl) done finally have "cmod (x - y) < e * cmod y" . then show ?thesis by (rule e) qed qed then have inj: "inj_on (\w. w^n) (ball z d)" by (simp add: inj_on_def) have cont: "continuous_on (ball z d) (\w. w ^ n)" by (intro continuous_intros) have noncon: "\ (\w::complex. w^n) constant_on UNIV" by (metis UNIV_I assms constant_on_def power_one zero_neq_one zero_power) have im_eq: "(\w. w^n) ` ball z' d = (\w. w^n) ` ball z d" if z': "z'^n = z^n" for z' proof - have nz': "norm z' = norm z" using that assms power_eq_imp_eq_norm by blast have "(w \ (\w. w^n) ` ball z' d) = (w \ (\w. w^n) ` ball z d)" for w proof (cases "w=0") case True with assms show ?thesis by (simp add: image_def ball_def nz') next case False have "z' \ 0" using \z \ 0\ nz' by force have [simp]: "(z*x / z')^n = x^n" if "x \ 0" for x using z' that by (simp add: field_simps \z \ 0\) have [simp]: "cmod (z - z * x / z') = cmod (z' - x)" if "x \ 0" for x proof - have "cmod (z - z * x / z') = cmod z * cmod (1 - x / z')" by (metis (no_types) ab_semigroup_mult_class.mult_ac(1) divide_complex_def mult.right_neutral norm_mult right_diff_distrib') also have "... = cmod z' * cmod (1 - x / z')" by (simp add: nz') also have "... = cmod (z' - x)" by (simp add: \z' \ 0\ diff_divide_eq_iff norm_divide) finally show ?thesis . qed have [simp]: "(z'*x / z)^n = x^n" if "x \ 0" for x using z' that by (simp add: field_simps \z \ 0\) have [simp]: "cmod (z' - z' * x / z) = cmod (z - x)" if "x \ 0" for x proof - have "cmod (z * (1 - x * inverse z)) = cmod (z - x)" by (metis \z \ 0\ diff_divide_distrib divide_complex_def divide_self_if nonzero_eq_divide_eq semiring_normalization_rules(7)) then show ?thesis by (metis (no_types) mult.assoc divide_complex_def mult.right_neutral norm_mult nz' right_diff_distrib') qed show ?thesis unfolding image_def ball_def apply safe apply simp_all apply (rule_tac x="z/z' * x" in exI) using assms False apply (simp add: dist_norm) apply (rule_tac x="z'/z * x" in exI) using assms False apply (simp add: dist_norm) done qed then show ?thesis by blast qed have ex_ball: "\B. (\z'. B = ball z' d \ z'^n = z^n) \ x \ B" if "x \ 0" and eq: "x^n = w^n" and dzw: "dist z w < d" for x w proof - have "w \ 0" by (metis assms power_eq_0_iff that(1) that(2)) have [simp]: "cmod x = cmod w" using assms power_eq_imp_eq_norm eq by blast have [simp]: "cmod (x * z / w - x) = cmod (z - w)" proof - have "cmod (x * z / w - x) = cmod x * cmod (z / w - 1)" by (metis (no_types) mult.right_neutral norm_mult right_diff_distrib' times_divide_eq_right) also have "... = cmod w * cmod (z / w - 1)" by simp also have "... = cmod (z - w)" by (simp add: \w \ 0\ divide_diff_eq_iff nonzero_norm_divide) finally show ?thesis . qed show ?thesis apply (rule_tac x="ball (z / w * x) d" in exI) using \d > 0\ that apply (simp add: ball_eq_ball_iff) apply (simp add: \z \ 0\ \w \ 0\ field_simps) apply (simp add: dist_norm) done qed show ?thesis proof (rule exI, intro conjI) show "z ^ n \ (\w. w ^ n) ` ball z d" using \d > 0\ by simp show "open ((\w. w ^ n) ` ball z d)" by (rule invariance_of_domain [OF cont open_ball inj]) show "0 \ (\w. w ^ n) ` ball z d" using \z \ 0\ assms by (force simp: d_def) show "\v. \v = - {0} \ (\z. z ^ n) -` (\w. w ^ n) ` ball z d \ (\u\v. open u \ 0 \ u) \ disjoint v \ (\u\v. Ex (homeomorphism u ((\w. w ^ n) ` ball z d) (\z. z ^ n)))" proof (rule exI, intro ballI conjI) show "\{ball z' d |z'. z'^n = z^n} = - {0} \ (\z. z ^ n) -` (\w. w ^ n) ` ball z d" (is "?l = ?r") proof show "?l \ ?r" apply auto apply (simp add: assms d_def power_eq_imp_eq_norm that) by (metis im_eq image_eqI mem_ball) show "?r \ ?l" by auto (meson ex_ball) qed show "\u. u \ {ball z' d |z'. z' ^ n = z ^ n} \ 0 \ u" by (force simp add: assms d_def power_eq_imp_eq_norm that) show "disjoint {ball z' d |z'. z' ^ n = z ^ n}" proof (clarsimp simp add: pairwise_def disjnt_iff) fix \ \ x assume "\^n = z^n" "\^n = z^n" "ball \ d \ ball \ d" and "dist \ x < d" "dist \ x < d" then have "dist \ \ < d+d" using dist_triangle_less_add by blast then have "cmod (\ - \) < 2*d" by (simp add: dist_norm) also have "... \ e * cmod z" using mult_right_mono \0 < e\ that by (auto simp: d_def) finally have "cmod (\ - \) < e * cmod z" . with e have "\ = \" by (metis \\^n = z^n\ \\^n = z^n\ assms power_eq_imp_eq_norm) then show "False" using \ball \ d \ ball \ d\ by blast qed show "Ex (homeomorphism u ((\w. w ^ n) ` ball z d) (\z. z ^ n))" if "u \ {ball z' d |z'. z' ^ n = z ^ n}" for u proof (rule invariance_of_domain_homeomorphism [of "u" "\z. z^n"]) show "open u" using that by auto show "continuous_on u (\z. z ^ n)" by (intro continuous_intros) show "inj_on (\z. z ^ n) u" using that by (auto simp: iff_x_eq_y inj_on_def) show "\g. homeomorphism u ((\z. z ^ n) ` u) (\z. z ^ n) g \ Ex (homeomorphism u ((\w. w ^ n) ` ball z d) (\z. z ^ n))" using im_eq that by clarify metis qed auto qed auto qed qed show ?thesis using assms apply (simp add: covering_space_def zn1 zn2) apply (subst zn2 [symmetric]) apply (simp add: openin_open_eq open_Compl) apply (blast intro: zn3) done qed corollary covering_space_square_punctured_plane: "covering_space (- {0}) (\z::complex. z^2) (- {0})" by (simp add: covering_space_power_punctured_plane) proposition covering_space_exp_punctured_plane: "covering_space UNIV (\z::complex. exp z) (- {0})" proof (simp add: covering_space_def, intro conjI ballI) show "continuous_on UNIV (\z::complex. exp z)" by (rule continuous_on_exp [OF continuous_on_id]) show "range exp = - {0::complex}" by auto (metis exp_Ln range_eqI) show "\T. z \ T \ openin (top_of_set (- {0})) T \ (\v. \v = exp -` T \ (\u\v. open u) \ disjoint v \ (\u\v. \q. homeomorphism u T exp q))" if "z \ - {0::complex}" for z proof - have "z \ 0" using that by auto have inj_exp: "inj_on exp (ball (Ln z) 1)" apply (rule inj_on_subset [OF inj_on_exp_pi [of "Ln z"]]) using pi_ge_two by (simp add: ball_subset_ball_iff) define \ where "\ \ range (\n. (\x. x + of_real (2 * of_int n * pi) * \) ` (ball(Ln z) 1))" show ?thesis proof (intro exI conjI) show "z \ exp ` (ball(Ln z) 1)" by (metis \z \ 0\ centre_in_ball exp_Ln rev_image_eqI zero_less_one) have "open (- {0::complex})" by blast moreover have "inj_on exp (ball (Ln z) 1)" apply (rule inj_on_subset [OF inj_on_exp_pi [of "Ln z"]]) using pi_ge_two by (simp add: ball_subset_ball_iff) ultimately show "openin (top_of_set (- {0})) (exp ` ball (Ln z) 1)" by (auto simp: openin_open_eq invariance_of_domain continuous_on_exp [OF continuous_on_id]) show "\\ = exp -` exp ` ball (Ln z) 1" by (force simp: \_def Complex_Transcendental.exp_eq image_iff) show "\V\\. open V" by (auto simp: \_def inj_on_def continuous_intros invariance_of_domain) have xy: "2 \ cmod (2 * of_int x * of_real pi * \ - 2 * of_int y * of_real pi * \)" if "x < y" for x y proof - have "1 \ abs (x - y)" using that by linarith then have "1 \ cmod (of_int x - of_int y) * 1" by (metis mult.right_neutral norm_of_int of_int_1_le_iff of_int_abs of_int_diff) also have "... \ cmod (of_int x - of_int y) * of_real pi" apply (rule mult_left_mono) using pi_ge_two by auto also have "... \ cmod ((of_int x - of_int y) * of_real pi * \)" by (simp add: norm_mult) also have "... \ cmod (of_int x * of_real pi * \ - of_int y * of_real pi * \)" by (simp add: algebra_simps) finally have "1 \ cmod (of_int x * of_real pi * \ - of_int y * of_real pi * \)" . then have "2 * 1 \ cmod (2 * (of_int x * of_real pi * \ - of_int y * of_real pi * \))" by (metis mult_le_cancel_left_pos norm_mult_numeral1 zero_less_numeral) then show ?thesis by (simp add: algebra_simps) qed show "disjoint \" apply (clarsimp simp add: \_def pairwise_def disjnt_def add.commute [of _ "x*y" for x y] ball_eq_ball_iff) apply (rule disjoint_ballI) apply (auto simp: dist_norm neq_iff) by (metis norm_minus_commute xy)+ show "\u\\. \q. homeomorphism u (exp ` ball (Ln z) 1) exp q" proof fix u assume "u \ \" then obtain n where n: "u = (\x. x + of_real (2 * of_int n * pi) * \) ` (ball(Ln z) 1)" by (auto simp: \_def) have "compact (cball (Ln z) 1)" by simp moreover have "continuous_on (cball (Ln z) 1) exp" by (rule continuous_on_exp [OF continuous_on_id]) moreover have "inj_on exp (cball (Ln z) 1)" apply (rule inj_on_subset [OF inj_on_exp_pi [of "Ln z"]]) using pi_ge_two by (simp add: cball_subset_ball_iff) ultimately obtain \ where hom: "homeomorphism (cball (Ln z) 1) (exp ` cball (Ln z) 1) exp \" using homeomorphism_compact by blast have eq1: "exp ` u = exp ` ball (Ln z) 1" unfolding n apply (auto simp: algebra_simps) apply (rename_tac w) apply (rule_tac x = "w + \ * (of_int n * (of_real pi * 2))" in image_eqI) apply (auto simp: image_iff) done have \exp: "\ (exp x) + 2 * of_int n * of_real pi * \ = x" if "x \ u" for x proof - have "exp x = exp (x - 2 * of_int n * of_real pi * \)" by (simp add: exp_eq) then have "\ (exp x) = \ (exp (x - 2 * of_int n * of_real pi * \))" by simp also have "... = x - 2 * of_int n * of_real pi * \" apply (rule homeomorphism_apply1 [OF hom]) using \x \ u\ by (auto simp: n) finally show ?thesis by simp qed have exp2n: "exp (\ (exp x) + 2 * of_int n * complex_of_real pi * \) = exp x" if "dist (Ln z) x < 1" for x using that by (auto simp: exp_eq homeomorphism_apply1 [OF hom]) have cont: "continuous_on (exp ` ball (Ln z) 1) (\x. \ x + 2 * of_int n * complex_of_real pi * \)" apply (intro continuous_intros) apply (rule continuous_on_subset [OF homeomorphism_cont2 [OF hom]]) apply (force simp:) done show "\q. homeomorphism u (exp ` ball (Ln z) 1) exp q" apply (rule_tac x="(\x. x + of_real(2 * n * pi) * \) \ \" in exI) unfolding homeomorphism_def apply (intro conjI ballI eq1 continuous_on_exp [OF continuous_on_id]) apply (auto simp: \exp exp2n cont n) apply (simp add: homeomorphism_apply1 [OF hom]) using hom homeomorphism_apply1 apply (force simp: image_iff) done qed qed qed qed subsection\Hence the Borsukian results about mappings into circles\(*FIX ME title *) lemma inessential_eq_continuous_logarithm: fixes f :: "'a::real_normed_vector \ complex" shows "(\a. homotopic_with_canon (\h. True) S (-{0}) f (\t. a)) \ (\g. continuous_on S g \ (\x \ S. f x = exp(g x)))" (is "?lhs \ ?rhs") proof assume ?lhs thus ?rhs by (metis covering_space_lift_inessential_function covering_space_exp_punctured_plane) next assume ?rhs then obtain g where contg: "continuous_on S g" and f: "\x. x \ S \ f x = exp(g x)" by metis obtain a where "homotopic_with_canon (\h. True) S (- {of_real 0}) (exp \ g) (\x. a)" proof (rule nullhomotopic_through_contractible [OF contg subset_UNIV _ _ contractible_UNIV]) show "continuous_on (UNIV::complex set) exp" by (intro continuous_intros) show "range exp \ - {0}" by auto qed force then have "homotopic_with_canon (\h. True) S (- {0}) f (\t. a)" using f homotopic_with_eq by fastforce then show ?lhs .. qed corollary inessential_imp_continuous_logarithm_circle: fixes f :: "'a::real_normed_vector \ complex" assumes "homotopic_with_canon (\h. True) S (sphere 0 1) f (\t. a)" obtains g where "continuous_on S g" and "\x. x \ S \ f x = exp(g x)" proof - have "homotopic_with_canon (\h. True) S (- {0}) f (\t. a)" using assms homotopic_with_subset_right by fastforce then show ?thesis by (metis inessential_eq_continuous_logarithm that) qed lemma inessential_eq_continuous_logarithm_circle: fixes f :: "'a::real_normed_vector \ complex" shows "(\a. homotopic_with_canon (\h. True) S (sphere 0 1) f (\t. a)) \ (\g. continuous_on S g \ (\x \ S. f x = exp(\ * of_real(g x))))" (is "?lhs \ ?rhs") proof assume L: ?lhs then obtain g where contg: "continuous_on S g" and g: "\x. x \ S \ f x = exp(g x)" using inessential_imp_continuous_logarithm_circle by blast have "f ` S \ sphere 0 1" by (metis L homotopic_with_imp_subset1) then have "\x. x \ S \ Re (g x) = 0" using g by auto then show ?rhs apply (rule_tac x="Im \ g" in exI) apply (intro conjI contg continuous_intros) apply (auto simp: Euler g) done next assume ?rhs then obtain g where contg: "continuous_on S g" and g: "\x. x \ S \ f x = exp(\* of_real(g x))" by metis obtain a where "homotopic_with_canon (\h. True) S (sphere 0 1) ((exp \ (\z. \*z)) \ (of_real \ g)) (\x. a)" proof (rule nullhomotopic_through_contractible) show "continuous_on S (complex_of_real \ g)" by (intro conjI contg continuous_intros) show "(complex_of_real \ g) ` S \ \" by auto show "continuous_on \ (exp \ (*)\)" by (intro continuous_intros) show "(exp \ (*)\) ` \ \ sphere 0 1" by (auto simp: complex_is_Real_iff) qed (auto simp: convex_Reals convex_imp_contractible) moreover have "\x. x \ S \ (exp \ (*)\ \ (complex_of_real \ g)) x = f x" by (simp add: g) ultimately have "homotopic_with_canon (\h. True) S (sphere 0 1) f (\t. a)" using homotopic_with_eq by force then show ?lhs .. qed proposition homotopic_with_sphere_times: fixes f :: "'a::real_normed_vector \ complex" assumes hom: "homotopic_with_canon (\x. True) S (sphere 0 1) f g" and conth: "continuous_on S h" and hin: "\x. x \ S \ h x \ sphere 0 1" shows "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x * h x) (\x. g x * h x)" proof - obtain k where contk: "continuous_on ({0..1::real} \ S) k" and kim: "k ` ({0..1} \ S) \ sphere 0 1" and k0: "\x. k(0, x) = f x" and k1: "\x. k(1, x) = g x" using hom by (auto simp: homotopic_with_def) show ?thesis apply (simp add: homotopic_with) apply (rule_tac x="\z. k z*(h \ snd)z" in exI) apply (intro conjI contk continuous_intros) apply (simp add: conth) using kim hin apply (force simp: norm_mult k0 k1)+ done qed proposition homotopic_circlemaps_divide: fixes f :: "'a::real_normed_vector \ complex" shows "homotopic_with_canon (\x. True) S (sphere 0 1) f g \ continuous_on S f \ f ` S \ sphere 0 1 \ continuous_on S g \ g ` S \ sphere 0 1 \ (\c. homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. c))" proof - have "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. 1)" if "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. c)" for c proof - have "S = {} \ path_component (sphere 0 1) 1 c" using homotopic_with_imp_subset2 [OF that] path_connected_sphere [of "0::complex" 1] by (auto simp: path_connected_component) then have "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. 1) (\x. c)" by (simp add: homotopic_constant_maps) then show ?thesis using homotopic_with_symD homotopic_with_trans that by blast qed then have *: "(\c. homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. c)) \ homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. 1)" by auto have "homotopic_with_canon (\x. True) S (sphere 0 1) f g \ continuous_on S f \ f ` S \ sphere 0 1 \ continuous_on S g \ g ` S \ sphere 0 1 \ homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. 1)" (is "?lhs \ ?rhs") proof assume L: ?lhs have geq1 [simp]: "\x. x \ S \ cmod (g x) = 1" using homotopic_with_imp_subset2 [OF L] by (simp add: image_subset_iff) have cont: "continuous_on S (inverse \ g)" apply (rule continuous_intros) using homotopic_with_imp_continuous [OF L] apply blast apply (rule continuous_on_subset [of "sphere 0 1", OF continuous_on_inverse]) apply (auto) done have "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. 1)" using homotopic_with_sphere_times [OF L cont] apply (rule homotopic_with_eq) apply (auto simp: division_ring_class.divide_inverse norm_inverse) by (metis geq1 norm_zero right_inverse zero_neq_one) with L show ?rhs by (auto simp: homotopic_with_imp_continuous dest: homotopic_with_imp_subset1 homotopic_with_imp_subset2) next assume ?rhs then show ?lhs by (elim conjE homotopic_with_eq [OF homotopic_with_sphere_times]; force) qed then show ?thesis by (simp add: *) qed subsection\Upper and lower hemicontinuous functions\ text\And relation in the case of preimage map to open and closed maps, and fact that upper and lower hemicontinuity together imply continuity in the sense of the Hausdorff metric (at points where the function gives a bounded and nonempty set).\ text\Many similar proofs below.\ lemma upper_hemicontinuous: assumes "\x. x \ S \ f x \ T" shows "((\U. openin (top_of_set T) U \ openin (top_of_set S) {x \ S. f x \ U}) \ (\U. closedin (top_of_set T) U \ closedin (top_of_set S) {x \ S. f x \ U \ {}}))" (is "?lhs = ?rhs") proof (intro iffI allI impI) fix U assume * [rule_format]: ?lhs and "closedin (top_of_set T) U" then have "openin (top_of_set T) (T - U)" by (simp add: openin_diff) then have "openin (top_of_set S) {x \ S. f x \ T - U}" using * [of "T-U"] by blast moreover have "S - {x \ S. f x \ T - U} = {x \ S. f x \ U \ {}}" using assms by blast ultimately show "closedin (top_of_set S) {x \ S. f x \ U \ {}}" by (simp add: openin_closedin_eq) next fix U assume * [rule_format]: ?rhs and "openin (top_of_set T) U" then have "closedin (top_of_set T) (T - U)" by (simp add: closedin_diff) then have "closedin (top_of_set S) {x \ S. f x \ (T - U) \ {}}" using * [of "T-U"] by blast moreover have "{x \ S. f x \ (T - U) \ {}} = S - {x \ S. f x \ U}" using assms by auto ultimately show "openin (top_of_set S) {x \ S. f x \ U}" by (simp add: openin_closedin_eq) qed lemma lower_hemicontinuous: assumes "\x. x \ S \ f x \ T" shows "((\U. closedin (top_of_set T) U \ closedin (top_of_set S) {x \ S. f x \ U}) \ (\U. openin (top_of_set T) U \ openin (top_of_set S) {x \ S. f x \ U \ {}}))" (is "?lhs = ?rhs") proof (intro iffI allI impI) fix U assume * [rule_format]: ?lhs and "openin (top_of_set T) U" then have "closedin (top_of_set T) (T - U)" by (simp add: closedin_diff) then have "closedin (top_of_set S) {x \ S. f x \ T-U}" using * [of "T-U"] by blast moreover have "{x \ S. f x \ T-U} = S - {x \ S. f x \ U \ {}}" using assms by auto ultimately show "openin (top_of_set S) {x \ S. f x \ U \ {}}" by (simp add: openin_closedin_eq) next fix U assume * [rule_format]: ?rhs and "closedin (top_of_set T) U" then have "openin (top_of_set T) (T - U)" by (simp add: openin_diff) then have "openin (top_of_set S) {x \ S. f x \ (T - U) \ {}}" using * [of "T-U"] by blast moreover have "S - {x \ S. f x \ (T - U) \ {}} = {x \ S. f x \ U}" using assms by blast ultimately show "closedin (top_of_set S) {x \ S. f x \ U}" by (simp add: openin_closedin_eq) qed lemma open_map_iff_lower_hemicontinuous_preimage: assumes "f ` S \ T" shows "((\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)) \ (\U. closedin (top_of_set S) U \ closedin (top_of_set T) {y \ T. {x. x \ S \ f x = y} \ U}))" (is "?lhs = ?rhs") proof (intro iffI allI impI) fix U assume * [rule_format]: ?lhs and "closedin (top_of_set S) U" then have "openin (top_of_set S) (S - U)" by (simp add: openin_diff) then have "openin (top_of_set T) (f ` (S - U))" using * [of "S-U"] by blast moreover have "T - (f ` (S - U)) = {y \ T. {x \ S. f x = y} \ U}" using assms by blast ultimately show "closedin (top_of_set T) {y \ T. {x \ S. f x = y} \ U}" by (simp add: openin_closedin_eq) next fix U assume * [rule_format]: ?rhs and opeSU: "openin (top_of_set S) U" then have "closedin (top_of_set S) (S - U)" by (simp add: closedin_diff) then have "closedin (top_of_set T) {y \ T. {x \ S. f x = y} \ S - U}" using * [of "S-U"] by blast moreover have "{y \ T. {x \ S. f x = y} \ S - U} = T - (f ` U)" using assms openin_imp_subset [OF opeSU] by auto ultimately show "openin (top_of_set T) (f ` U)" using assms openin_imp_subset [OF opeSU] by (force simp: openin_closedin_eq) qed lemma closed_map_iff_upper_hemicontinuous_preimage: assumes "f ` S \ T" shows "((\U. closedin (top_of_set S) U \ closedin (top_of_set T) (f ` U)) \ (\U. openin (top_of_set S) U \ openin (top_of_set T) {y \ T. {x. x \ S \ f x = y} \ U}))" (is "?lhs = ?rhs") proof (intro iffI allI impI) fix U assume * [rule_format]: ?lhs and opeSU: "openin (top_of_set S) U" then have "closedin (top_of_set S) (S - U)" by (simp add: closedin_diff) then have "closedin (top_of_set T) (f ` (S - U))" using * [of "S-U"] by blast moreover have "f ` (S - U) = T - {y \ T. {x. x \ S \ f x = y} \ U}" using assms openin_imp_subset [OF opeSU] by auto ultimately show "openin (top_of_set T) {y \ T. {x. x \ S \ f x = y} \ U}" using assms openin_imp_subset [OF opeSU] by (force simp: openin_closedin_eq) next fix U assume * [rule_format]: ?rhs and cloSU: "closedin (top_of_set S) U" then have "openin (top_of_set S) (S - U)" by (simp add: openin_diff) then have "openin (top_of_set T) {y \ T. {x \ S. f x = y} \ S - U}" using * [of "S-U"] by blast moreover have "(f ` U) = T - {y \ T. {x \ S. f x = y} \ S - U}" using assms closedin_imp_subset [OF cloSU] by auto ultimately show "closedin (top_of_set T) (f ` U)" by (simp add: openin_closedin_eq) qed proposition upper_lower_hemicontinuous_explicit: fixes T :: "('b::{real_normed_vector,heine_borel}) set" assumes fST: "\x. x \ S \ f x \ T" and ope: "\U. openin (top_of_set T) U \ openin (top_of_set S) {x \ S. f x \ U}" and clo: "\U. closedin (top_of_set T) U \ closedin (top_of_set S) {x \ S. f x \ U}" and "x \ S" "0 < e" and bofx: "bounded(f x)" and fx_ne: "f x \ {}" obtains d where "0 < d" "\x'. \x' \ S; dist x x' < d\ \ (\y \ f x. \y'. y' \ f x' \ dist y y' < e) \ (\y' \ f x'. \y. y \ f x \ dist y' y < e)" proof - have "openin (top_of_set T) (T \ (\a\f x. \b\ball 0 e. {a + b}))" by (auto simp: open_sums openin_open_Int) with ope have "openin (top_of_set S) {u \ S. f u \ T \ (\a\f x. \b\ball 0 e. {a + b})}" by blast with \0 < e\ \x \ S\ obtain d1 where "d1 > 0" and d1: "\x'. \x' \ S; dist x' x < d1\ \ f x' \ T \ f x' \ (\a \ f x. \b \ ball 0 e. {a + b})" by (force simp: openin_euclidean_subtopology_iff dest: fST) have oo: "\U. openin (top_of_set T) U \ openin (top_of_set S) {x \ S. f x \ U \ {}}" apply (rule lower_hemicontinuous [THEN iffD1, rule_format]) using fST clo by auto have "compact (closure(f x))" by (simp add: bofx) moreover have "closure(f x) \ (\a \ f x. ball a (e/2))" using \0 < e\ by (force simp: closure_approachable simp del: divide_const_simps) ultimately obtain C where "C \ f x" "finite C" "closure(f x) \ (\a \ C. ball a (e/2))" apply (rule compactE, force) by (metis finite_subset_image) then have fx_cover: "f x \ (\a \ C. ball a (e/2))" by (meson closure_subset order_trans) with fx_ne have "C \ {}" by blast have xin: "x \ (\a \ C. {x \ S. f x \ T \ ball a (e/2) \ {}})" using \x \ S\ \0 < e\ fST \C \ f x\ by force have "openin (top_of_set S) {x \ S. f x \ (T \ ball a (e/2)) \ {}}" for a by (simp add: openin_open_Int oo) then have "openin (top_of_set S) (\a \ C. {x \ S. f x \ T \ ball a (e/2) \ {}})" by (simp add: Int_assoc openin_INT2 [OF \finite C\ \C \ {}\]) with xin obtain d2 where "d2>0" and d2: "\u v. \u \ S; dist u x < d2; v \ C\ \ f u \ T \ ball v (e/2) \ {}" unfolding openin_euclidean_subtopology_iff using xin by fastforce show ?thesis proof (intro that conjI ballI) show "0 < min d1 d2" using \0 < d1\ \0 < d2\ by linarith next fix x' y assume "x' \ S" "dist x x' < min d1 d2" "y \ f x" then have dd2: "dist x' x < d2" by (auto simp: dist_commute) obtain a where "a \ C" "y \ ball a (e/2)" using fx_cover \y \ f x\ by auto then show "\y'. y' \ f x' \ dist y y' < e" using d2 [OF \x' \ S\ dd2] dist_triangle_half_r by fastforce next fix x' y' assume "x' \ S" "dist x x' < min d1 d2" "y' \ f x'" then have "dist x' x < d1" by (auto simp: dist_commute) then have "y' \ (\a\f x. \b\ball 0 e. {a + b})" using d1 [OF \x' \ S\] \y' \ f x'\ by force then show "\y. y \ f x \ dist y' y < e" apply auto by (metis add_diff_cancel_left' dist_norm) qed qed subsection\Complex logs exist on various "well-behaved" sets\ lemma continuous_logarithm_on_contractible: fixes f :: "'a::real_normed_vector \ complex" assumes "continuous_on S f" "contractible S" "\z. z \ S \ f z \ 0" obtains g where "continuous_on S g" "\x. x \ S \ f x = exp(g x)" proof - obtain c where hom: "homotopic_with_canon (\h. True) S (-{0}) f (\x. c)" using nullhomotopic_from_contractible assms by (metis imageE subset_Compl_singleton) then show ?thesis by (metis inessential_eq_continuous_logarithm that) qed lemma continuous_logarithm_on_simply_connected: fixes f :: "'a::real_normed_vector \ complex" assumes contf: "continuous_on S f" and S: "simply_connected S" "locally path_connected S" and f: "\z. z \ S \ f z \ 0" obtains g where "continuous_on S g" "\x. x \ S \ f x = exp(g x)" using covering_space_lift [OF covering_space_exp_punctured_plane S contf] by (metis (full_types) f imageE subset_Compl_singleton) lemma continuous_logarithm_on_cball: fixes f :: "'a::real_normed_vector \ complex" assumes "continuous_on (cball a r) f" and "\z. z \ cball a r \ f z \ 0" obtains h where "continuous_on (cball a r) h" "\z. z \ cball a r \ f z = exp(h z)" using assms continuous_logarithm_on_contractible convex_imp_contractible by blast lemma continuous_logarithm_on_ball: fixes f :: "'a::real_normed_vector \ complex" assumes "continuous_on (ball a r) f" and "\z. z \ ball a r \ f z \ 0" obtains h where "continuous_on (ball a r) h" "\z. z \ ball a r \ f z = exp(h z)" using assms continuous_logarithm_on_contractible convex_imp_contractible by blast lemma continuous_sqrt_on_contractible: fixes f :: "'a::real_normed_vector \ complex" assumes "continuous_on S f" "contractible S" and "\z. z \ S \ f z \ 0" obtains g where "continuous_on S g" "\x. x \ S \ f x = (g x) ^ 2" proof - obtain g where contg: "continuous_on S g" and feq: "\x. x \ S \ f x = exp(g x)" using continuous_logarithm_on_contractible [OF assms] by blast show ?thesis proof show "continuous_on S (\z. exp (g z / 2))" by (rule continuous_on_compose2 [of UNIV exp]; intro continuous_intros contg subset_UNIV) auto show "\x. x \ S \ f x = (exp (g x / 2))\<^sup>2" by (metis exp_double feq nonzero_mult_div_cancel_left times_divide_eq_right zero_neq_numeral) qed qed lemma continuous_sqrt_on_simply_connected: fixes f :: "'a::real_normed_vector \ complex" assumes contf: "continuous_on S f" and S: "simply_connected S" "locally path_connected S" and f: "\z. z \ S \ f z \ 0" obtains g where "continuous_on S g" "\x. x \ S \ f x = (g x) ^ 2" proof - obtain g where contg: "continuous_on S g" and feq: "\x. x \ S \ f x = exp(g x)" using continuous_logarithm_on_simply_connected [OF assms] by blast show ?thesis proof show "continuous_on S (\z. exp (g z / 2))" by (rule continuous_on_compose2 [of UNIV exp]; intro continuous_intros contg subset_UNIV) auto show "\x. x \ S \ f x = (exp (g x / 2))\<^sup>2" by (metis exp_double feq nonzero_mult_div_cancel_left times_divide_eq_right zero_neq_numeral) qed qed subsection\Another simple case where sphere maps are nullhomotopic\ lemma inessential_spheremap_2_aux: fixes f :: "'a::euclidean_space \ complex" assumes 2: "2 < DIM('a)" and contf: "continuous_on (sphere a r) f" and fim: "f `(sphere a r) \ (sphere 0 1)" obtains c where "homotopic_with_canon (\z. True) (sphere a r) (sphere 0 1) f (\x. c)" proof - obtain g where contg: "continuous_on (sphere a r) g" and feq: "\x. x \ sphere a r \ f x = exp(g x)" proof (rule continuous_logarithm_on_simply_connected [OF contf]) show "simply_connected (sphere a r)" using 2 by (simp add: simply_connected_sphere_eq) show "locally path_connected (sphere a r)" by (simp add: locally_path_connected_sphere) show "\z. z \ sphere a r \ f z \ 0" using fim by force qed auto have "\g. continuous_on (sphere a r) g \ (\x\sphere a r. f x = exp (\ * complex_of_real (g x)))" proof (intro exI conjI) show "continuous_on (sphere a r) (Im \ g)" by (intro contg continuous_intros continuous_on_compose) show "\x\sphere a r. f x = exp (\ * complex_of_real ((Im \ g) x))" using exp_eq_polar feq fim norm_exp_eq_Re by auto qed with inessential_eq_continuous_logarithm_circle that show ?thesis by metis qed lemma inessential_spheremap_2: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes a2: "2 < DIM('a)" and b2: "DIM('b) = 2" and contf: "continuous_on (sphere a r) f" and fim: "f `(sphere a r) \ (sphere b s)" obtains c where "homotopic_with_canon (\z. True) (sphere a r) (sphere b s) f (\x. c)" proof (cases "s \ 0") case True then show ?thesis using contf contractible_sphere fim nullhomotopic_into_contractible that by blast next case False then have "sphere b s homeomorphic sphere (0::complex) 1" using assms by (simp add: homeomorphic_spheres_gen) then obtain h k where hk: "homeomorphism (sphere b s) (sphere (0::complex) 1) h k" by (auto simp: homeomorphic_def) then have conth: "continuous_on (sphere b s) h" and contk: "continuous_on (sphere 0 1) k" and him: "h ` sphere b s \ sphere 0 1" and kim: "k ` sphere 0 1 \ sphere b s" by (simp_all add: homeomorphism_def) obtain c where "homotopic_with_canon (\z. True) (sphere a r) (sphere 0 1) (h \ f) (\x. c)" proof (rule inessential_spheremap_2_aux [OF a2]) show "continuous_on (sphere a r) (h \ f)" by (meson continuous_on_compose [OF contf] conth continuous_on_subset fim) show "(h \ f) ` sphere a r \ sphere 0 1" using fim him by force qed auto then have "homotopic_with_canon (\f. True) (sphere a r) (sphere b s) (k \ (h \ f)) (k \ (\x. c))" by (rule homotopic_with_compose_continuous_left [OF _ contk kim]) then have "homotopic_with_canon (\z. True) (sphere a r) (sphere b s) f (\x. k c)" apply (rule homotopic_with_eq, auto) by (metis fim hk homeomorphism_def image_subset_iff mem_sphere) then show ?thesis by (metis that) qed subsection\Holomorphic logarithms and square roots\ lemma contractible_imp_holomorphic_log: assumes holf: "f holomorphic_on S" and S: "contractible S" and fnz: "\z. z \ S \ f z \ 0" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = exp(g z)" proof - have contf: "continuous_on S f" by (simp add: holf holomorphic_on_imp_continuous_on) obtain g where contg: "continuous_on S g" and feq: "\x. x \ S \ f x = exp (g x)" by (metis continuous_logarithm_on_contractible [OF contf S fnz]) have "g field_differentiable at z within S" if "f field_differentiable at z within S" "z \ S" for z proof - obtain f' where f': "((\y. (f y - f z) / (y - z)) \ f') (at z within S)" using \f field_differentiable at z within S\ by (auto simp: field_differentiable_def has_field_derivative_iff) then have ee: "((\x. (exp(g x) - exp(g z)) / (x - z)) \ f') (at z within S)" by (simp add: feq \z \ S\ Lim_transform_within [OF _ zero_less_one]) have "(((\y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \ g) \ exp (g z)) (at z within S)" proof (rule tendsto_compose_at) show "(g \ g z) (at z within S)" using contg continuous_on \z \ S\ by blast show "(\y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \g z\ exp (g z)" apply (subst Lim_at_zero) apply (simp add: DERIV_D cong: if_cong Lim_cong_within) done qed auto then have dd: "((\x. if g x = g z then exp(g z) else (exp(g x) - exp(g z)) / (g x - g z)) \ exp(g z)) (at z within S)" by (simp add: o_def) have "continuous (at z within S) g" using contg continuous_on_eq_continuous_within \z \ S\ by blast then have "(\\<^sub>F x in at z within S. dist (g x) (g z) < 2*pi)" by (simp add: continuous_within tendsto_iff) then have "\\<^sub>F x in at z within S. exp (g x) = exp (g z) \ g x \ g z \ x = z" apply (rule eventually_mono) apply (auto simp: exp_eq dist_norm norm_mult) done then have "((\y. (g y - g z) / (y - z)) \ f' / exp (g z)) (at z within S)" by (auto intro!: Lim_transform_eventually [OF tendsto_divide [OF ee dd]]) then show ?thesis by (auto simp: field_differentiable_def has_field_derivative_iff) qed then have "g holomorphic_on S" using holf holomorphic_on_def by auto then show ?thesis using feq that by auto qed (*Identical proofs*) lemma simply_connected_imp_holomorphic_log: assumes holf: "f holomorphic_on S" and S: "simply_connected S" "locally path_connected S" and fnz: "\z. z \ S \ f z \ 0" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = exp(g z)" proof - have contf: "continuous_on S f" by (simp add: holf holomorphic_on_imp_continuous_on) obtain g where contg: "continuous_on S g" and feq: "\x. x \ S \ f x = exp (g x)" by (metis continuous_logarithm_on_simply_connected [OF contf S fnz]) have "g field_differentiable at z within S" if "f field_differentiable at z within S" "z \ S" for z proof - obtain f' where f': "((\y. (f y - f z) / (y - z)) \ f') (at z within S)" using \f field_differentiable at z within S\ by (auto simp: field_differentiable_def has_field_derivative_iff) then have ee: "((\x. (exp(g x) - exp(g z)) / (x - z)) \ f') (at z within S)" by (simp add: feq \z \ S\ Lim_transform_within [OF _ zero_less_one]) have "(((\y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \ g) \ exp (g z)) (at z within S)" proof (rule tendsto_compose_at) show "(g \ g z) (at z within S)" using contg continuous_on \z \ S\ by blast show "(\y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \g z\ exp (g z)" apply (subst Lim_at_zero) apply (simp add: DERIV_D cong: if_cong Lim_cong_within) done qed auto then have dd: "((\x. if g x = g z then exp(g z) else (exp(g x) - exp(g z)) / (g x - g z)) \ exp(g z)) (at z within S)" by (simp add: o_def) have "continuous (at z within S) g" using contg continuous_on_eq_continuous_within \z \ S\ by blast then have "(\\<^sub>F x in at z within S. dist (g x) (g z) < 2*pi)" by (simp add: continuous_within tendsto_iff) then have "\\<^sub>F x in at z within S. exp (g x) = exp (g z) \ g x \ g z \ x = z" apply (rule eventually_mono) apply (auto simp: exp_eq dist_norm norm_mult) done then have "((\y. (g y - g z) / (y - z)) \ f' / exp (g z)) (at z within S)" by (auto intro!: Lim_transform_eventually [OF tendsto_divide [OF ee dd]]) then show ?thesis by (auto simp: field_differentiable_def has_field_derivative_iff) qed then have "g holomorphic_on S" using holf holomorphic_on_def by auto then show ?thesis using feq that by auto qed lemma contractible_imp_holomorphic_sqrt: assumes holf: "f holomorphic_on S" and S: "contractible S" and fnz: "\z. z \ S \ f z \ 0" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = g z ^ 2" proof - obtain g where holg: "g holomorphic_on S" and feq: "\z. z \ S \ f z = exp(g z)" using contractible_imp_holomorphic_log [OF assms] by blast show ?thesis proof show "exp \ (\z. z / 2) \ g holomorphic_on S" by (intro holomorphic_on_compose holg holomorphic_intros) auto show "\z. z \ S \ f z = ((exp \ (\z. z / 2) \ g) z)\<^sup>2" apply (auto simp: feq) by (metis eq_divide_eq_numeral1(1) exp_double mult.commute zero_neq_numeral) qed qed lemma simply_connected_imp_holomorphic_sqrt: assumes holf: "f holomorphic_on S" and S: "simply_connected S" "locally path_connected S" and fnz: "\z. z \ S \ f z \ 0" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = g z ^ 2" proof - obtain g where holg: "g holomorphic_on S" and feq: "\z. z \ S \ f z = exp(g z)" using simply_connected_imp_holomorphic_log [OF assms] by blast show ?thesis proof show "exp \ (\z. z / 2) \ g holomorphic_on S" by (intro holomorphic_on_compose holg holomorphic_intros) auto show "\z. z \ S \ f z = ((exp \ (\z. z / 2) \ g) z)\<^sup>2" apply (auto simp: feq) by (metis eq_divide_eq_numeral1(1) exp_double mult.commute zero_neq_numeral) qed qed text\ Related theorems about holomorphic inverse cosines.\ lemma contractible_imp_holomorphic_arccos: assumes holf: "f holomorphic_on S" and S: "contractible S" and non1: "\z. z \ S \ f z \ 1 \ f z \ -1" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = cos(g z)" proof - have hol1f: "(\z. 1 - f z ^ 2) holomorphic_on S" by (intro holomorphic_intros holf) obtain g where holg: "g holomorphic_on S" and eq: "\z. z \ S \ 1 - (f z)\<^sup>2 = (g z)\<^sup>2" using contractible_imp_holomorphic_sqrt [OF hol1f S] by (metis eq_iff_diff_eq_0 non1 power2_eq_1_iff) have holfg: "(\z. f z + \*g z) holomorphic_on S" by (intro holf holg holomorphic_intros) have "\z. z \ S \ f z + \*g z \ 0" by (metis Arccos_body_lemma eq add.commute add.inverse_unique complex_i_mult_minus power2_csqrt power2_eq_iff) then obtain h where holh: "h holomorphic_on S" and fgeq: "\z. z \ S \ f z + \*g z = exp (h z)" using contractible_imp_holomorphic_log [OF holfg S] by metis show ?thesis proof show "(\z. -\*h z) holomorphic_on S" by (intro holh holomorphic_intros) show "f z = cos (- \*h z)" if "z \ S" for z proof - have "(f z + \*g z)*(f z - \*g z) = 1" using that eq by (auto simp: algebra_simps power2_eq_square) then have "f z - \*g z = inverse (f z + \*g z)" using inverse_unique by force also have "... = exp (- h z)" by (simp add: exp_minus fgeq that) finally have "f z = exp (- h z) + \*g z" by (simp add: diff_eq_eq) then show ?thesis apply (simp add: cos_exp_eq) by (metis fgeq add.assoc mult_2_right that) qed qed qed lemma contractible_imp_holomorphic_arccos_bounded: assumes holf: "f holomorphic_on S" and S: "contractible S" and "a \ S" and non1: "\z. z \ S \ f z \ 1 \ f z \ -1" obtains g where "g holomorphic_on S" "norm(g a) \ pi + norm(f a)" "\z. z \ S \ f z = cos(g z)" proof - obtain g where holg: "g holomorphic_on S" and feq: "\z. z \ S \ f z = cos (g z)" using contractible_imp_holomorphic_arccos [OF holf S non1] by blast obtain b where "cos b = f a" "norm b \ pi + norm (f a)" using cos_Arccos norm_Arccos_bounded by blast then have "cos b = cos (g a)" by (simp add: \a \ S\ feq) then consider n where "n \ \" "b = g a + of_real(2*n*pi)" | n where "n \ \" "b = -g a + of_real(2*n*pi)" by (auto simp: complex_cos_eq) then show ?thesis proof cases case 1 show ?thesis proof show "(\z. g z + of_real(2*n*pi)) holomorphic_on S" by (intro holomorphic_intros holg) show "cmod (g a + of_real(2*n*pi)) \ pi + cmod (f a)" using "1" \cmod b \ pi + cmod (f a)\ by blast show "\z. z \ S \ f z = cos (g z + complex_of_real (2*n*pi))" by (metis \n \ \\ complex_cos_eq feq) qed next case 2 show ?thesis proof show "(\z. -g z + of_real(2*n*pi)) holomorphic_on S" by (intro holomorphic_intros holg) show "cmod (-g a + of_real(2*n*pi)) \ pi + cmod (f a)" using "2" \cmod b \ pi + cmod (f a)\ by blast show "\z. z \ S \ f z = cos (-g z + complex_of_real (2*n*pi))" by (metis \n \ \\ complex_cos_eq feq) qed qed qed subsection\The "Borsukian" property of sets\ text\This doesn't have a standard name. Kuratowski uses ``contractible with respect to \[S\<^sup>1]\'' while Whyburn uses ``property b''. It's closely related to unicoherence.\ definition\<^marker>\tag important\ Borsukian where "Borsukian S \ \f. continuous_on S f \ f ` S \ (- {0::complex}) \ (\a. homotopic_with_canon (\h. True) S (- {0}) f (\x. a))" lemma Borsukian_retraction_gen: assumes "Borsukian S" "continuous_on S h" "h ` S = T" "continuous_on T k" "k ` T \ S" "\y. y \ T \ h(k y) = y" shows "Borsukian T" proof - interpret R: Retracts S h T k using assms by (simp add: Retracts.intro) show ?thesis using assms apply (simp add: Borsukian_def, clarify) apply (rule R.cohomotopically_trivial_retraction_null_gen [OF TrueI TrueI refl, of "-{0}"], auto) done qed lemma retract_of_Borsukian: "\Borsukian T; S retract_of T\ \ Borsukian S" apply (auto simp: retract_of_def retraction_def) apply (erule (1) Borsukian_retraction_gen) apply (meson retraction retraction_def) apply (auto) done lemma homeomorphic_Borsukian: "\Borsukian S; S homeomorphic T\ \ Borsukian T" using Borsukian_retraction_gen order_refl by (fastforce simp add: homeomorphism_def homeomorphic_def) lemma homeomorphic_Borsukian_eq: "S homeomorphic T \ Borsukian S \ Borsukian T" by (meson homeomorphic_Borsukian homeomorphic_sym) lemma Borsukian_translation: fixes S :: "'a::real_normed_vector set" shows "Borsukian (image (\x. a + x) S) \ Borsukian S" apply (rule homeomorphic_Borsukian_eq) using homeomorphic_translation homeomorphic_sym by blast lemma Borsukian_injective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "Borsukian(f ` S) \ Borsukian S" apply (rule homeomorphic_Borsukian_eq) using assms homeomorphic_sym linear_homeomorphic_image by blast lemma homotopy_eqv_Borsukianness: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" assumes "S homotopy_eqv T" shows "(Borsukian S \ Borsukian T)" by (meson Borsukian_def assms homotopy_eqv_cohomotopic_triviality_null) lemma Borsukian_alt: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f g. continuous_on S f \ f ` S \ -{0} \ continuous_on S g \ g ` S \ -{0} \ homotopic_with_canon (\h. True) S (- {0::complex}) f g)" unfolding Borsukian_def homotopic_triviality by (simp add: path_connected_punctured_universe) lemma Borsukian_continuous_logarithm: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f. continuous_on S f \ f ` S \ (- {0::complex}) \ (\g. continuous_on S g \ (\x \ S. f x = exp(g x))))" by (simp add: Borsukian_def inessential_eq_continuous_logarithm) lemma Borsukian_continuous_logarithm_circle: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f. continuous_on S f \ f ` S \ sphere (0::complex) 1 \ (\g. continuous_on S g \ (\x \ S. f x = exp(g x))))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (force simp: Borsukian_continuous_logarithm) next assume RHS [rule_format]: ?rhs show ?lhs proof (clarsimp simp: Borsukian_continuous_logarithm) fix f :: "'a \ complex" assume contf: "continuous_on S f" and 0: "0 \ f ` S" then have "continuous_on S (\x. f x / complex_of_real (cmod (f x)))" by (intro continuous_intros) auto moreover have "(\x. f x / complex_of_real (cmod (f x))) ` S \ sphere 0 1" using 0 by (auto simp: norm_divide) ultimately obtain g where contg: "continuous_on S g" and fg: "\x \ S. f x / complex_of_real (cmod (f x)) = exp(g x)" using RHS [of "\x. f x / of_real(norm(f x))"] by auto show "\g. continuous_on S g \ (\x\S. f x = exp (g x))" proof (intro exI ballI conjI) show "continuous_on S (\x. (Ln \ of_real \ norm \ f)x + g x)" by (intro continuous_intros contf contg conjI) (use "0" in auto) show "f x = exp ((Ln \ complex_of_real \ cmod \ f) x + g x)" if "x \ S" for x using 0 that apply (clarsimp simp: exp_add) apply (subst exp_Ln, force) by (metis eq_divide_eq exp_not_eq_zero fg mult.commute) qed qed qed lemma Borsukian_continuous_logarithm_circle_real: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f. continuous_on S f \ f ` S \ sphere (0::complex) 1 \ (\g. continuous_on S (complex_of_real \ g) \ (\x \ S. f x = exp(\ * of_real(g x)))))" (is "?lhs = ?rhs") proof assume LHS: ?lhs show ?rhs proof (clarify) fix f :: "'a \ complex" assume "continuous_on S f" and f01: "f ` S \ sphere 0 1" then obtain g where contg: "continuous_on S g" and "\x. x \ S \ f x = exp(g x)" using LHS by (auto simp: Borsukian_continuous_logarithm_circle) then have "\x\S. f x = exp (\ * complex_of_real ((Im \ g) x))" using f01 apply (simp add: image_iff subset_iff) by (metis cis_conv_exp exp_eq_polar mult.left_neutral norm_exp_eq_Re of_real_1) then show "\g. continuous_on S (complex_of_real \ g) \ (\x\S. f x = exp (\ * complex_of_real (g x)))" by (rule_tac x="Im \ g" in exI) (force intro: continuous_intros contg) qed next assume RHS [rule_format]: ?rhs show ?lhs proof (clarsimp simp: Borsukian_continuous_logarithm_circle) fix f :: "'a \ complex" assume "continuous_on S f" and f01: "f ` S \ sphere 0 1" then obtain g where contg: "continuous_on S (complex_of_real \ g)" and "\x. x \ S \ f x = exp(\ * of_real(g x))" by (metis RHS) then show "\g. continuous_on S g \ (\x\S. f x = exp (g x))" by (rule_tac x="\x. \* of_real(g x)" in exI) (auto simp: continuous_intros contg) qed qed lemma Borsukian_circle: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f. continuous_on S f \ f ` S \ sphere (0::complex) 1 \ (\a. homotopic_with_canon (\h. True) S (sphere (0::complex) 1) f (\x. a)))" by (simp add: inessential_eq_continuous_logarithm_circle Borsukian_continuous_logarithm_circle_real) lemma contractible_imp_Borsukian: "contractible S \ Borsukian S" by (meson Borsukian_def nullhomotopic_from_contractible) lemma simply_connected_imp_Borsukian: fixes S :: "'a::real_normed_vector set" shows "\simply_connected S; locally path_connected S\ \ Borsukian S" apply (simp add: Borsukian_continuous_logarithm) by (metis (no_types, lifting) continuous_logarithm_on_simply_connected image_iff) lemma starlike_imp_Borsukian: fixes S :: "'a::real_normed_vector set" shows "starlike S \ Borsukian S" by (simp add: contractible_imp_Borsukian starlike_imp_contractible) lemma Borsukian_empty: "Borsukian {}" by (auto simp: contractible_imp_Borsukian) lemma Borsukian_UNIV: "Borsukian (UNIV :: 'a::real_normed_vector set)" by (auto simp: contractible_imp_Borsukian) lemma convex_imp_Borsukian: fixes S :: "'a::real_normed_vector set" shows "convex S \ Borsukian S" by (meson Borsukian_def convex_imp_contractible nullhomotopic_from_contractible) proposition Borsukian_sphere: fixes a :: "'a::euclidean_space" shows "3 \ DIM('a) \ Borsukian (sphere a r)" apply (rule simply_connected_imp_Borsukian) using simply_connected_sphere apply blast using ENR_imp_locally_path_connected ENR_sphere by blast proposition Borsukian_open_Un: fixes S :: "'a::real_normed_vector set" assumes opeS: "openin (top_of_set (S \ T)) S" and opeT: "openin (top_of_set (S \ T)) T" and BS: "Borsukian S" and BT: "Borsukian T" and ST: "connected(S \ T)" shows "Borsukian(S \ T)" proof (clarsimp simp add: Borsukian_continuous_logarithm) fix f :: "'a \ complex" assume contf: "continuous_on (S \ T) f" and 0: "0 \ f ` (S \ T)" then have contfS: "continuous_on S f" and contfT: "continuous_on T f" using continuous_on_subset by auto have "\continuous_on S f; f ` S \ -{0}\ \ \g. continuous_on S g \ (\x \ S. f x = exp(g x))" using BS by (auto simp: Borsukian_continuous_logarithm) then obtain g where contg: "continuous_on S g" and fg: "\x. x \ S \ f x = exp(g x)" using "0" contfS by blast have "\continuous_on T f; f ` T \ -{0}\ \ \g. continuous_on T g \ (\x \ T. f x = exp(g x))" using BT by (auto simp: Borsukian_continuous_logarithm) then obtain h where conth: "continuous_on T h" and fh: "\x. x \ T \ f x = exp(h x)" using "0" contfT by blast show "\g. continuous_on (S \ T) g \ (\x\S \ T. f x = exp (g x))" proof (cases "S \ T = {}") case True show ?thesis proof (intro exI conjI) show "continuous_on (S \ T) (\x. if x \ S then g x else h x)" apply (rule continuous_on_cases_local_open [OF opeS opeT contg conth]) using True by blast show "\x\S \ T. f x = exp (if x \ S then g x else h x)" using fg fh by auto qed next case False have "(\x. g x - h x) constant_on S \ T" proof (rule continuous_discrete_range_constant [OF ST]) show "continuous_on (S \ T) (\x. g x - h x)" apply (intro continuous_intros) apply (meson contg continuous_on_subset inf_le1) by (meson conth continuous_on_subset inf_sup_ord(2)) show "\e>0. \y. y \ S \ T \ g y - h y \ g x - h x \ e \ cmod (g y - h y - (g x - h x))" if "x \ S \ T" for x proof - have "g y - g x = h y - h x" if "y \ S" "y \ T" "cmod (g y - g x - (h y - h x)) < 2 * pi" for y proof (rule exp_complex_eqI) have "\Im (g y - g x) - Im (h y - h x)\ \ cmod (g y - g x - (h y - h x))" by (metis abs_Im_le_cmod minus_complex.simps(2)) then show "\Im (g y - g x) - Im (h y - h x)\ < 2 * pi" using that by linarith have "exp (g x) = exp (h x)" "exp (g y) = exp (h y)" using fg fh that \x \ S \ T\ by fastforce+ then show "exp (g y - g x) = exp (h y - h x)" by (simp add: exp_diff) qed then show ?thesis by (rule_tac x="2*pi" in exI) (fastforce simp add: algebra_simps) qed qed then obtain a where a: "\x. x \ S \ T \ g x - h x = a" by (auto simp: constant_on_def) with False have "exp a = 1" by (metis IntI disjoint_iff_not_equal divide_self_if exp_diff exp_not_eq_zero fg fh) with a show ?thesis apply (rule_tac x="\x. if x \ S then g x else a + h x" in exI) apply (intro continuous_on_cases_local_open opeS opeT contg conth continuous_intros conjI) apply (auto simp: algebra_simps fg fh exp_add) done qed qed text\The proof is a duplicate of that of \Borsukian_open_Un\.\ lemma Borsukian_closed_Un: fixes S :: "'a::real_normed_vector set" assumes cloS: "closedin (top_of_set (S \ T)) S" and cloT: "closedin (top_of_set (S \ T)) T" and BS: "Borsukian S" and BT: "Borsukian T" and ST: "connected(S \ T)" shows "Borsukian(S \ T)" proof (clarsimp simp add: Borsukian_continuous_logarithm) fix f :: "'a \ complex" assume contf: "continuous_on (S \ T) f" and 0: "0 \ f ` (S \ T)" then have contfS: "continuous_on S f" and contfT: "continuous_on T f" using continuous_on_subset by auto have "\continuous_on S f; f ` S \ -{0}\ \ \g. continuous_on S g \ (\x \ S. f x = exp(g x))" using BS by (auto simp: Borsukian_continuous_logarithm) then obtain g where contg: "continuous_on S g" and fg: "\x. x \ S \ f x = exp(g x)" using "0" contfS by blast have "\continuous_on T f; f ` T \ -{0}\ \ \g. continuous_on T g \ (\x \ T. f x = exp(g x))" using BT by (auto simp: Borsukian_continuous_logarithm) then obtain h where conth: "continuous_on T h" and fh: "\x. x \ T \ f x = exp(h x)" using "0" contfT by blast show "\g. continuous_on (S \ T) g \ (\x\S \ T. f x = exp (g x))" proof (cases "S \ T = {}") case True show ?thesis proof (intro exI conjI) show "continuous_on (S \ T) (\x. if x \ S then g x else h x)" apply (rule continuous_on_cases_local [OF cloS cloT contg conth]) using True by blast show "\x\S \ T. f x = exp (if x \ S then g x else h x)" using fg fh by auto qed next case False have "(\x. g x - h x) constant_on S \ T" proof (rule continuous_discrete_range_constant [OF ST]) show "continuous_on (S \ T) (\x. g x - h x)" apply (intro continuous_intros) apply (meson contg continuous_on_subset inf_le1) by (meson conth continuous_on_subset inf_sup_ord(2)) show "\e>0. \y. y \ S \ T \ g y - h y \ g x - h x \ e \ cmod (g y - h y - (g x - h x))" if "x \ S \ T" for x proof - have "g y - g x = h y - h x" if "y \ S" "y \ T" "cmod (g y - g x - (h y - h x)) < 2 * pi" for y proof (rule exp_complex_eqI) have "\Im (g y - g x) - Im (h y - h x)\ \ cmod (g y - g x - (h y - h x))" by (metis abs_Im_le_cmod minus_complex.simps(2)) then show "\Im (g y - g x) - Im (h y - h x)\ < 2 * pi" using that by linarith have "exp (g x) = exp (h x)" "exp (g y) = exp (h y)" using fg fh that \x \ S \ T\ by fastforce+ then show "exp (g y - g x) = exp (h y - h x)" by (simp add: exp_diff) qed then show ?thesis by (rule_tac x="2*pi" in exI) (fastforce simp add: algebra_simps) qed qed then obtain a where a: "\x. x \ S \ T \ g x - h x = a" by (auto simp: constant_on_def) with False have "exp a = 1" by (metis IntI disjoint_iff_not_equal divide_self_if exp_diff exp_not_eq_zero fg fh) with a show ?thesis apply (rule_tac x="\x. if x \ S then g x else a + h x" in exI) apply (intro continuous_on_cases_local cloS cloT contg conth continuous_intros conjI) apply (auto simp: algebra_simps fg fh exp_add) done qed qed lemma Borsukian_separation_compact: fixes S :: "complex set" assumes "compact S" shows "Borsukian S \ connected(- S)" by (simp add: Borsuk_separation_theorem Borsukian_circle assms) lemma Borsukian_monotone_image_compact: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "Borsukian S" and contf: "continuous_on S f" and fim: "f ` S = T" and "compact S" and conn: "\y. y \ T \ connected {x. x \ S \ f x = y}" shows "Borsukian T" proof (clarsimp simp add: Borsukian_continuous_logarithm) fix g :: "'b \ complex" assume contg: "continuous_on T g" and 0: "0 \ g ` T" have "continuous_on S (g \ f)" using contf contg continuous_on_compose fim by blast moreover have "(g \ f) ` S \ -{0}" using fim 0 by auto ultimately obtain h where conth: "continuous_on S h" and gfh: "\x. x \ S \ (g \ f) x = exp(h x)" using \Borsukian S\ by (auto simp: Borsukian_continuous_logarithm) have "\y. \x. y \ T \ x \ S \ f x = y" using fim by auto then obtain f' where f': "\y. y \ T \ f' y \ S \ f (f' y) = y" by metis have *: "(\x. h x - h(f' y)) constant_on {x. x \ S \ f x = y}" if "y \ T" for y proof (rule continuous_discrete_range_constant [OF conn [OF that], of "\x. h x - h (f' y)"], simp_all add: algebra_simps) show "continuous_on {x \ S. f x = y} (\x. h x - h (f' y))" by (intro continuous_intros continuous_on_subset [OF conth]) auto show "\e>0. \u. u \ S \ f u = y \ h u \ h x \ e \ cmod (h u - h x)" if x: "x \ S \ f x = y" for x proof - have "h u = h x" if "u \ S" "f u = y" "cmod (h u - h x) < 2 * pi" for u proof (rule exp_complex_eqI) have "\Im (h u) - Im (h x)\ \ cmod (h u - h x)" by (metis abs_Im_le_cmod minus_complex.simps(2)) then show "\Im (h u) - Im (h x)\ < 2 * pi" using that by linarith show "exp (h u) = exp (h x)" by (simp add: gfh [symmetric] x that) qed then show ?thesis by (rule_tac x="2*pi" in exI) (fastforce simp add: algebra_simps) qed qed have "h x = h (f' (f x))" if "x \ S" for x using * [of "f x"] fim that unfolding constant_on_def by clarsimp (metis f' imageI right_minus_eq) moreover have "\x. x \ T \ \u. u \ S \ x = f u \ h (f' x) = h u" using f' by fastforce ultimately have eq: "((\x. (x, (h \ f') x)) ` T) = {p. \x. x \ S \ (x, p) \ (S \ UNIV) \ ((\z. snd z - ((f \ fst) z, (h \ fst) z)) -` {0})}" using fim by (auto simp: image_iff) show "\h. continuous_on T h \ (\x\T. g x = exp (h x))" proof (intro exI conjI) show "continuous_on T (h \ f')" proof (rule continuous_from_closed_graph [of "h ` S"]) show "compact (h ` S)" by (simp add: \compact S\ compact_continuous_image conth) show "(h \ f') ` T \ h ` S" by (auto simp: f') show "closed ((\x. (x, (h \ f') x)) ` T)" apply (subst eq) apply (intro closed_compact_projection [OF \compact S\] continuous_closed_preimage continuous_intros continuous_on_subset [OF contf] continuous_on_subset [OF conth]) apply (auto simp: \compact S\ closed_Times compact_imp_closed) done qed qed (use f' gfh in fastforce) qed lemma Borsukian_open_map_image_compact: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "Borsukian S" and contf: "continuous_on S f" and fim: "f ` S = T" and "compact S" and ope: "\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)" shows "Borsukian T" proof (clarsimp simp add: Borsukian_continuous_logarithm_circle_real) fix g :: "'b \ complex" assume contg: "continuous_on T g" and gim: "g ` T \ sphere 0 1" have "continuous_on S (g \ f)" using contf contg continuous_on_compose fim by blast moreover have "(g \ f) ` S \ sphere 0 1" using fim gim by auto ultimately obtain h where cont_cxh: "continuous_on S (complex_of_real \ h)" and gfh: "\x. x \ S \ (g \ f) x = exp(\ * of_real(h x))" using \Borsukian S\ Borsukian_continuous_logarithm_circle_real by metis then have conth: "continuous_on S h" by simp have "\x. x \ S \ f x = y \ (\x' \ S. f x' = y \ h x \ h x')" if "y \ T" for y proof - have 1: "compact (h ` {x \ S. f x = y})" proof (rule compact_continuous_image) show "continuous_on {x \ S. f x = y} h" by (rule continuous_on_subset [OF conth]) auto have "compact (S \ f -` {y})" by (rule proper_map_from_compact [OF contf _ \compact S\, of T]) (simp_all add: fim that) then show "compact {x \ S. f x = y}" by (auto simp: vimage_def Int_def) qed have 2: "h ` {x \ S. f x = y} \ {}" using fim that by auto have "\s \ h ` {x \ S. f x = y}. \t \ h ` {x \ S. f x = y}. s \ t" using compact_attains_inf [OF 1 2] by blast then show ?thesis by auto qed then obtain k where kTS: "\y. y \ T \ k y \ S" and fk: "\y. y \ T \ f (k y) = y " and hle: "\x' y. \y \ T; x' \ S; f x' = y\ \ h (k y) \ h x'" by metis have "continuous_on T (h \ k)" proof (clarsimp simp add: continuous_on_iff) fix y and e::real assume "y \ T" "0 < e" moreover have "uniformly_continuous_on S (complex_of_real \ h)" using \compact S\ cont_cxh compact_uniformly_continuous by blast ultimately obtain d where "0 < d" and d: "\x x'. \x\S; x'\S; dist x' x < d\ \ dist (h x') (h x) < e" by (force simp: uniformly_continuous_on_def) obtain \ where "0 < \" and \: "\x'. \x' \ T; dist y x' < \\ \ (\v \ {z \ S. f z = y}. \v'. v' \ {z \ S. f z = x'} \ dist v v' < d) \ (\v' \ {z \ S. f z = x'}. \v. v \ {z \ S. f z = y} \ dist v' v < d)" proof (rule upper_lower_hemicontinuous_explicit [of T "\y. {z \ S. f z = y}" S]) show "\U. openin (top_of_set S) U \ openin (top_of_set T) {x \ T. {z \ S. f z = x} \ U}" using closed_map_iff_upper_hemicontinuous_preimage [OF fim [THEN equalityD1]] by (simp add: Abstract_Topology_2.continuous_imp_closed_map \compact S\ contf fim) show "\U. closedin (top_of_set S) U \ closedin (top_of_set T) {x \ T. {z \ S. f z = x} \ U}" using ope open_map_iff_lower_hemicontinuous_preimage [OF fim [THEN equalityD1]] by meson show "bounded {z \ S. f z = y}" by (metis (no_types, lifting) compact_imp_bounded [OF \compact S\] bounded_subset mem_Collect_eq subsetI) qed (use \y \ T\ \0 < d\ fk kTS in \force+\) have "dist (h (k y')) (h (k y)) < e" if "y' \ T" "dist y y' < \" for y' proof - have k1: "k y \ S" "f (k y) = y" and k2: "k y' \ S" "f (k y') = y'" by (auto simp: \y \ T\ \y' \ T\ kTS fk) have 1: "\v. \v \ S; f v = y\ \ \v'. v' \ {z \ S. f z = y'} \ dist v v' < d" and 2: "\v'. \v' \ S; f v' = y'\ \ \v. v \ {z \ S. f z = y} \ dist v' v < d" using \ [OF that] by auto then obtain w' w where "w' \ S" "f w' = y'" "dist (k y) w' < d" and "w \ S" "f w = y" "dist (k y') w < d" using 1 [OF k1] 2 [OF k2] by auto then show ?thesis using d [of w "k y'"] d [of w' "k y"] k1 k2 \y' \ T\ \y \ T\ hle by (fastforce simp: dist_norm abs_diff_less_iff algebra_simps) qed then show "\d>0. \x'\T. dist x' y < d \ dist (h (k x')) (h (k y)) < e" using \0 < \\ by (auto simp: dist_commute) qed then show "\h. continuous_on T h \ (\x\T. g x = exp (\ * complex_of_real (h x)))" using fk gfh kTS by force qed text\If two points are separated by a closed set, there's a minimal one.\ proposition closed_irreducible_separator: fixes a :: "'a::real_normed_vector" assumes "closed S" and ab: "\ connected_component (- S) a b" obtains T where "T \ S" "closed T" "T \ {}" "\ connected_component (- T) a b" "\U. U \ T \ connected_component (- U) a b" proof (cases "a \ S \ b \ S") case True then show ?thesis proof assume *: "a \ S" show ?thesis proof show "{a} \ S" using * by blast show "\ connected_component (- {a}) a b" using connected_component_in by auto show "\U. U \ {a} \ connected_component (- U) a b" by (metis connected_component_UNIV UNIV_I compl_bot_eq connected_component_eq_eq less_le_not_le subset_singletonD) qed auto next assume *: "b \ S" show ?thesis proof show "{b} \ S" using * by blast show "\ connected_component (- {b}) a b" using connected_component_in by auto show "\U. U \ {b} \ connected_component (- U) a b" by (metis connected_component_UNIV UNIV_I compl_bot_eq connected_component_eq_eq less_le_not_le subset_singletonD) qed auto qed next case False define A where "A \ connected_component_set (- S) a" define B where "B \ connected_component_set (- (closure A)) b" have "a \ A" using False A_def by auto have "b \ B" unfolding A_def B_def closure_Un_frontier using ab False \closed S\ frontier_complement frontier_of_connected_component_subset frontier_subset_closed by force have "frontier B \ frontier (connected_component_set (- closure A) b)" using B_def by blast also have frsub: "... \ frontier A" proof - have "\A. closure (- closure (- A)) \ closure A" by (metis (no_types) closure_mono closure_subset compl_le_compl_iff double_compl) then show ?thesis by (metis (no_types) closure_closure double_compl frontier_closures frontier_of_connected_component_subset le_inf_iff subset_trans) qed finally have frBA: "frontier B \ frontier A" . show ?thesis proof show "frontier B \ S" proof - have "frontier S \ S" by (simp add: \closed S\ frontier_subset_closed) then show ?thesis using frsub frontier_complement frontier_of_connected_component_subset unfolding A_def B_def by blast qed show "closed (frontier B)" by simp show "\ connected_component (- frontier B) a b" unfolding connected_component_def proof clarify fix T assume "connected T" and TB: "T \ - frontier B" and "a \ T" and "b \ T" have "a \ B" by (metis A_def B_def ComplD \a \ A\ assms(1) closed_open connected_component_subset in_closure_connected_component subsetD) have "T \ B \ {}" using \b \ B\ \b \ T\ by blast moreover have "T - B \ {}" using \a \ B\ \a \ T\ by blast ultimately show "False" using connected_Int_frontier [of T B] TB \connected T\ by blast qed moreover have "connected_component (- frontier B) a b" if "frontier B = {}" apply (simp add: that) using connected_component_eq_UNIV by blast ultimately show "frontier B \ {}" by blast show "connected_component (- U) a b" if "U \ frontier B" for U proof - obtain p where Usub: "U \ frontier B" and p: "p \ frontier B" "p \ U" using \U \ frontier B\ by blast show ?thesis unfolding connected_component_def proof (intro exI conjI) have "connected ((insert p A) \ (insert p B))" proof (rule connected_Un) show "connected (insert p A)" by (metis A_def IntD1 frBA \p \ frontier B\ closure_insert closure_subset connected_connected_component connected_intermediate_closure frontier_closures insert_absorb subsetCE subset_insertI) show "connected (insert p B)" by (metis B_def IntD1 \p \ frontier B\ closure_insert closure_subset connected_connected_component connected_intermediate_closure frontier_closures insert_absorb subset_insertI) qed blast then show "connected (insert p (B \ A))" by (simp add: sup.commute) have "A \ - U" using A_def Usub \frontier B \ S\ connected_component_subset by fastforce moreover have "B \ - U" using B_def Usub connected_component_subset frBA frontier_closures by fastforce ultimately show "insert p (B \ A) \ - U" using p by auto qed (auto simp: \a \ A\ \b \ B\) qed qed qed lemma frontier_minimal_separating_closed_pointwise: fixes S :: "'a::real_normed_vector set" assumes S: "closed S" "a \ S" and nconn: "\ connected_component (- S) a b" and conn: "\T. \closed T; T \ S\ \ connected_component (- T) a b" shows "frontier(connected_component_set (- S) a) = S" (is "?F = S") proof - have "?F \ S" by (simp add: S componentsI frontier_of_components_closed_complement) moreover have False if "?F \ S" proof - have "connected_component (- ?F) a b" by (simp add: conn that) then obtain T where "connected T" "T \ -?F" "a \ T" "b \ T" by (auto simp: connected_component_def) moreover have "T \ ?F \ {}" proof (rule connected_Int_frontier [OF \connected T\]) show "T \ connected_component_set (- S) a \ {}" using \a \ S\ \a \ T\ by fastforce show "T - connected_component_set (- S) a \ {}" using \b \ T\ nconn by blast qed ultimately show ?thesis by blast qed ultimately show ?thesis by blast qed subsection\Unicoherence (closed)\ definition\<^marker>\tag important\ unicoherent where "unicoherent U \ \S T. connected S \ connected T \ S \ T = U \ closedin (top_of_set U) S \ closedin (top_of_set U) T \ connected (S \ T)" lemma unicoherentI [intro?]: assumes "\S T. \connected S; connected T; U = S \ T; closedin (top_of_set U) S; closedin (top_of_set U) T\ \ connected (S \ T)" shows "unicoherent U" using assms unfolding unicoherent_def by blast lemma unicoherentD: assumes "unicoherent U" "connected S" "connected T" "U = S \ T" "closedin (top_of_set U) S" "closedin (top_of_set U) T" shows "connected (S \ T)" using assms unfolding unicoherent_def by blast proposition homeomorphic_unicoherent: assumes ST: "S homeomorphic T" and S: "unicoherent S" shows "unicoherent T" proof - obtain f g where gf: "\x. x \ S \ g (f x) = x" and fim: "T = f ` S" and gfim: "g ` f ` S = S" and contf: "continuous_on S f" and contg: "continuous_on (f ` S) g" using ST by (auto simp: homeomorphic_def homeomorphism_def) show ?thesis proof fix U V assume "connected U" "connected V" and T: "T = U \ V" and cloU: "closedin (top_of_set T) U" and cloV: "closedin (top_of_set T) V" have "f ` (g ` U \ g ` V) \ U" "f ` (g ` U \ g ` V) \ V" using gf fim T by auto (metis UnCI image_iff)+ moreover have "U \ V \ f ` (g ` U \ g ` V)" using gf fim by (force simp: image_iff T) ultimately have "U \ V = f ` (g ` U \ g ` V)" by blast moreover have "connected (f ` (g ` U \ g ` V))" proof (rule connected_continuous_image) show "continuous_on (g ` U \ g ` V) f" apply (rule continuous_on_subset [OF contf]) using T fim gfim by blast show "connected (g ` U \ g ` V)" proof (intro conjI unicoherentD [OF S]) show "connected (g ` U)" "connected (g ` V)" using \connected U\ cloU \connected V\ cloV by (metis Topological_Spaces.connected_continuous_image closedin_imp_subset contg continuous_on_subset fim)+ show "S = g ` U \ g ` V" using T fim gfim by auto have hom: "homeomorphism T S g f" by (simp add: contf contg fim gf gfim homeomorphism_def) have "closedin (top_of_set T) U" "closedin (top_of_set T) V" by (simp_all add: cloU cloV) then show "closedin (top_of_set S) (g ` U)" "closedin (top_of_set S) (g ` V)" by (blast intro: homeomorphism_imp_closed_map [OF hom])+ qed qed ultimately show "connected (U \ V)" by metis qed qed lemma homeomorphic_unicoherent_eq: "S homeomorphic T \ (unicoherent S \ unicoherent T)" by (meson homeomorphic_sym homeomorphic_unicoherent) lemma unicoherent_translation: fixes S :: "'a::real_normed_vector set" shows "unicoherent (image (\x. a + x) S) \ unicoherent S" using homeomorphic_translation homeomorphic_unicoherent_eq by blast lemma unicoherent_injective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "(unicoherent(f ` S) \ unicoherent S)" using assms homeomorphic_unicoherent_eq linear_homeomorphic_image by blast lemma Borsukian_imp_unicoherent: fixes U :: "'a::euclidean_space set" assumes "Borsukian U" shows "unicoherent U" unfolding unicoherent_def proof clarify fix S T assume "connected S" "connected T" "U = S \ T" and cloS: "closedin (top_of_set (S \ T)) S" and cloT: "closedin (top_of_set (S \ T)) T" show "connected (S \ T)" unfolding connected_closedin_eq proof clarify fix V W assume "closedin (top_of_set (S \ T)) V" and "closedin (top_of_set (S \ T)) W" and VW: "V \ W = S \ T" "V \ W = {}" and "V \ {}" "W \ {}" then have cloV: "closedin (top_of_set U) V" and cloW: "closedin (top_of_set U) W" using \U = S \ T\ cloS cloT closedin_trans by blast+ obtain q where contq: "continuous_on U q" and q01: "\x. x \ U \ q x \ {0..1::real}" and qV: "\x. x \ V \ q x = 0" and qW: "\x. x \ W \ q x = 1" by (rule Urysohn_local [OF cloV cloW \V \ W = {}\, of 0 1]) (fastforce simp: closed_segment_eq_real_ivl) let ?h = "\x. if x \ S then exp(pi * \ * q x) else 1 / exp(pi * \ * q x)" have eqST: "exp(pi * \ * q x) = 1 / exp(pi * \ * q x)" if "x \ S \ T" for x proof - have "x \ V \ W" using that \V \ W = S \ T\ by blast with qV qW show ?thesis by force qed obtain g where contg: "continuous_on U g" and circle: "g ` U \ sphere 0 1" and S: "\x. x \ S \ g x = exp(pi * \ * q x)" and T: "\x. x \ T \ g x = 1 / exp(pi * \ * q x)" proof show "continuous_on U ?h" unfolding \U = S \ T\ proof (rule continuous_on_cases_local [OF cloS cloT]) show "continuous_on S (\x. exp (pi * \ * q x))" apply (intro continuous_intros) using \U = S \ T\ continuous_on_subset contq by blast show "continuous_on T (\x. 1 / exp (pi * \ * q x))" apply (intro continuous_intros) using \U = S \ T\ continuous_on_subset contq by auto qed (use eqST in auto) qed (use eqST in \auto simp: norm_divide\) then obtain h where conth: "continuous_on U h" and heq: "\x. x \ U \ g x = exp (h x)" by (metis Borsukian_continuous_logarithm_circle assms) obtain v w where "v \ V" "w \ W" using \V \ {}\ \W \ {}\ by blast then have vw: "v \ S \ T" "w \ S \ T" using VW by auto have iff: "2 * pi \ cmod (2 * of_int m * of_real pi * \ - 2 * of_int n * of_real pi * \) \ 1 \ abs (m - n)" for m n proof - have "2 * pi \ cmod (2 * of_int m * of_real pi * \ - 2 * of_int n * of_real pi * \) \ 2 * pi \ cmod ((2 * pi * \) * (of_int m - of_int n))" by (simp add: algebra_simps) also have "... \ 2 * pi \ 2 * pi * cmod (of_int m - of_int n)" by (simp add: norm_mult) also have "... \ 1 \ abs (m - n)" by simp (metis norm_of_int of_int_1_le_iff of_int_abs of_int_diff) finally show ?thesis . qed have *: "\n::int. h x - (pi * \ * q x) = (of_int(2*n) * pi) * \" if "x \ S" for x using that S \U = S \ T\ heq exp_eq [symmetric] by (simp add: algebra_simps) moreover have "(\x. h x - (pi * \ * q x)) constant_on S" proof (rule continuous_discrete_range_constant [OF \connected S\]) have "continuous_on S h" "continuous_on S q" using \U = S \ T\ continuous_on_subset conth contq by blast+ then show "continuous_on S (\x. h x - (pi * \ * q x))" by (intro continuous_intros) have "2*pi \ cmod (h y - (pi * \ * q y) - (h x - (pi * \ * q x)))" if "x \ S" "y \ S" and ne: "h y - (pi * \ * q y) \ h x - (pi * \ * q x)" for x y using * [OF \x \ S\] * [OF \y \ S\] ne by (auto simp: iff) then show "\x. x \ S \ \e>0. \y. y \ S \ h y - (pi * \ * q y) \ h x - (pi * \ * q x) \ e \ cmod (h y - (pi * \ * q y) - (h x - (pi * \ * q x)))" by (rule_tac x="2*pi" in exI) auto qed ultimately obtain m where m: "\x. x \ S \ h x - (pi * \ * q x) = (of_int(2*m) * pi) * \" using vw by (force simp: constant_on_def) have *: "\n::int. h x = - (pi * \ * q x) + (of_int(2*n) * pi) * \" if "x \ T" for x unfolding exp_eq [symmetric] using that T \U = S \ T\ by (simp add: exp_minus field_simps heq [symmetric]) moreover have "(\x. h x + (pi * \ * q x)) constant_on T" proof (rule continuous_discrete_range_constant [OF \connected T\]) have "continuous_on T h" "continuous_on T q" using \U = S \ T\ continuous_on_subset conth contq by blast+ then show "continuous_on T (\x. h x + (pi * \ * q x))" by (intro continuous_intros) have "2*pi \ cmod (h y + (pi * \ * q y) - (h x + (pi * \ * q x)))" if "x \ T" "y \ T" and ne: "h y + (pi * \ * q y) \ h x + (pi * \ * q x)" for x y using * [OF \x \ T\] * [OF \y \ T\] ne by (auto simp: iff) then show "\x. x \ T \ \e>0. \y. y \ T \ h y + (pi * \ * q y) \ h x + (pi * \ * q x) \ e \ cmod (h y + (pi * \ * q y) - (h x + (pi * \ * q x)))" by (rule_tac x="2*pi" in exI) auto qed ultimately obtain n where n: "\x. x \ T \ h x + (pi * \ * q x) = (of_int(2*n) * pi) * \" using vw by (force simp: constant_on_def) show "False" using m [of v] m [of w] n [of v] n [of w] vw by (auto simp: algebra_simps \v \ V\ \w \ W\ qV qW) qed qed corollary contractible_imp_unicoherent: fixes U :: "'a::euclidean_space set" assumes "contractible U" shows "unicoherent U" by (simp add: Borsukian_imp_unicoherent assms contractible_imp_Borsukian) corollary convex_imp_unicoherent: fixes U :: "'a::euclidean_space set" assumes "convex U" shows "unicoherent U" by (simp add: Borsukian_imp_unicoherent assms convex_imp_Borsukian) text\If the type class constraint can be relaxed, I don't know how!\ corollary unicoherent_UNIV: "unicoherent (UNIV :: 'a :: euclidean_space set)" by (simp add: convex_imp_unicoherent) lemma unicoherent_monotone_image_compact: fixes T :: "'b :: t2_space set" assumes S: "unicoherent S" "compact S" and contf: "continuous_on S f" and fim: "f ` S = T" and conn: "\y. y \ T \ connected (S \ f -` {y})" shows "unicoherent T" proof fix U V assume UV: "connected U" "connected V" "T = U \ V" and cloU: "closedin (top_of_set T) U" and cloV: "closedin (top_of_set T) V" moreover have "compact T" using \compact S\ compact_continuous_image contf fim by blast ultimately have "closed U" "closed V" by (auto simp: closedin_closed_eq compact_imp_closed) let ?SUV = "(S \ f -` U) \ (S \ f -` V)" have UV_eq: "f ` ?SUV = U \ V" using \T = U \ V\ fim by force+ have "connected (f ` ?SUV)" proof (rule connected_continuous_image) show "continuous_on ?SUV f" by (meson contf continuous_on_subset inf_le1) show "connected ?SUV" proof (rule unicoherentD [OF \unicoherent S\, of "S \ f -` U" "S \ f -` V"]) have "\C. closedin (top_of_set S) C \ closedin (top_of_set T) (f ` C)" by (metis \compact S\ closed_subset closedin_compact closedin_imp_subset compact_continuous_image compact_imp_closed contf continuous_on_subset fim image_mono) then show "connected (S \ f -` U)" "connected (S \ f -` V)" using UV by (auto simp: conn intro: connected_closed_monotone_preimage [OF contf fim]) show "S = (S \ f -` U) \ (S \ f -` V)" using UV fim by blast show "closedin (top_of_set S) (S \ f -` U)" "closedin (top_of_set S) (S \ f -` V)" by (auto simp: continuous_on_imp_closedin cloU cloV contf fim) qed qed with UV_eq show "connected (U \ V)" by simp qed subsection\Several common variants of unicoherence\ lemma connected_frontier_simple: fixes S :: "'a :: euclidean_space set" assumes "connected S" "connected(- S)" shows "connected(frontier S)" unfolding frontier_closures apply (rule unicoherentD [OF unicoherent_UNIV]) apply (simp_all add: assms connected_imp_connected_closure) by (simp add: closure_def) lemma connected_frontier_component_complement: fixes S :: "'a :: euclidean_space set" assumes "connected S" and C: "C \ components(- S)" shows "connected(frontier C)" apply (rule connected_frontier_simple) using C in_components_connected apply blast by (metis assms component_complement_connected) lemma connected_frontier_disjoint: fixes S :: "'a :: euclidean_space set" assumes "connected S" "connected T" "disjnt S T" and ST: "frontier S \ frontier T" shows "connected(frontier S)" proof (cases "S = UNIV") case True then show ?thesis by simp next case False then have "-S \ {}" by blast then obtain C where C: "C \ components(- S)" and "T \ C" by (metis ComplI disjnt_iff subsetI exists_component_superset \disjnt S T\ \connected T\) moreover have "frontier S = frontier C" proof - have "frontier C \ frontier S" using C frontier_complement frontier_of_components_subset by blast moreover have "x \ frontier C" if "x \ frontier S" for x proof - have "x \ closure C" using that unfolding frontier_def by (metis (no_types) Diff_eq ST \T \ C\ closure_mono contra_subsetD frontier_def le_inf_iff that) moreover have "x \ interior C" using that unfolding frontier_def by (metis C Compl_eq_Diff_UNIV Diff_iff subsetD in_components_subset interior_diff interior_mono) ultimately show ?thesis by (auto simp: frontier_def) qed ultimately show ?thesis by blast qed ultimately show ?thesis using \connected S\ connected_frontier_component_complement by auto qed subsection\Some separation results\ lemma separation_by_component_closed_pointwise: fixes S :: "'a :: euclidean_space set" assumes "closed S" "\ connected_component (- S) a b" obtains C where "C \ components S" "\ connected_component(- C) a b" proof (cases "a \ S \ b \ S") case True then show ?thesis using connected_component_in componentsI that by fastforce next case False obtain T where "T \ S" "closed T" "T \ {}" and nab: "\ connected_component (- T) a b" and conn: "\U. U \ T \ connected_component (- U) a b" using closed_irreducible_separator [OF assms] by metis moreover have "connected T" proof - have ab: "frontier(connected_component_set (- T) a) = T" "frontier(connected_component_set (- T) b) = T" using frontier_minimal_separating_closed_pointwise by (metis False \T \ S\ \closed T\ connected_component_sym conn connected_component_eq_empty connected_component_intermediate_subset empty_subsetI nab)+ have "connected (frontier (connected_component_set (- T) a))" proof (rule connected_frontier_disjoint) show "disjnt (connected_component_set (- T) a) (connected_component_set (- T) b)" unfolding disjnt_iff by (metis connected_component_eq connected_component_eq_empty connected_component_idemp mem_Collect_eq nab) show "frontier (connected_component_set (- T) a) \ frontier (connected_component_set (- T) b)" by (simp add: ab) qed auto with ab \closed T\ show ?thesis by simp qed ultimately obtain C where "C \ components S" "T \ C" using exists_component_superset [of T S] by blast then show ?thesis by (meson Compl_anti_mono connected_component_of_subset nab that) qed lemma separation_by_component_closed: fixes S :: "'a :: euclidean_space set" assumes "closed S" "\ connected(- S)" obtains C where "C \ components S" "\ connected(- C)" proof - obtain x y where "closed S" "x \ S" "y \ S" and "\ connected_component (- S) x y" using assms by (auto simp: connected_iff_connected_component) then obtain C where "C \ components S" "\ connected_component(- C) x y" using separation_by_component_closed_pointwise by metis then show "thesis" apply (clarify elim!: componentsE) by (metis Compl_iff \C \ components S\ \x \ S\ \y \ S\ connected_component_eq connected_component_eq_eq connected_iff_connected_component that) qed lemma separation_by_Un_closed_pointwise: fixes S :: "'a :: euclidean_space set" assumes ST: "closed S" "closed T" "S \ T = {}" and conS: "connected_component (- S) a b" and conT: "connected_component (- T) a b" shows "connected_component (- (S \ T)) a b" proof (rule ccontr) have "a \ S" "b \ S" "a \ T" "b \ T" using conS conT connected_component_in by auto assume "\ connected_component (- (S \ T)) a b" then obtain C where "C \ components (S \ T)" and C: "\ connected_component(- C) a b" using separation_by_component_closed_pointwise assms by blast then have "C \ S \ C \ T" proof - have "connected C" "C \ S \ T" using \C \ components (S \ T)\ in_components_subset by (blast elim: componentsE)+ moreover then have "C \ T = {} \ C \ S = {}" by (metis Int_empty_right ST inf.commute connected_closed) ultimately show ?thesis by blast qed then show False by (meson Compl_anti_mono C conS conT connected_component_of_subset) qed lemma separation_by_Un_closed: fixes S :: "'a :: euclidean_space set" assumes ST: "closed S" "closed T" "S \ T = {}" and conS: "connected(- S)" and conT: "connected(- T)" shows "connected(- (S \ T))" using assms separation_by_Un_closed_pointwise by (fastforce simp add: connected_iff_connected_component) lemma open_unicoherent_UNIV: fixes S :: "'a :: euclidean_space set" assumes "open S" "open T" "connected S" "connected T" "S \ T = UNIV" shows "connected(S \ T)" proof - have "connected(- (-S \ -T))" by (metis closed_Compl compl_sup compl_top_eq double_compl separation_by_Un_closed assms) then show ?thesis by simp qed lemma separation_by_component_open_aux: fixes S :: "'a :: euclidean_space set" assumes ST: "closed S" "closed T" "S \ T = {}" and "S \ {}" "T \ {}" obtains C where "C \ components(-(S \ T))" "C \ {}" "frontier C \ S \ {}" "frontier C \ T \ {}" proof (rule ccontr) let ?S = "S \ \{C \ components(- (S \ T)). frontier C \ S}" let ?T = "T \ \{C \ components(- (S \ T)). frontier C \ T}" assume "\ thesis" with that have *: "frontier C \ S = {} \ frontier C \ T = {}" if C: "C \ components (- (S \ T))" "C \ {}" for C using C by blast have "\A B::'a set. closed A \ closed B \ UNIV \ A \ B \ A \ B = {} \ A \ {} \ B \ {}" proof (intro exI conjI) have "frontier (\{C \ components (- S \ - T). frontier C \ S}) \ S" apply (rule subset_trans [OF frontier_Union_subset_closure]) by (metis (no_types, lifting) SUP_least \closed S\ closure_minimal mem_Collect_eq) then have "frontier ?S \ S" by (simp add: frontier_subset_eq assms subset_trans [OF frontier_Un_subset]) then show "closed ?S" using frontier_subset_eq by fastforce have "frontier (\{C \ components (- S \ - T). frontier C \ T}) \ T" apply (rule subset_trans [OF frontier_Union_subset_closure]) by (metis (no_types, lifting) SUP_least \closed T\ closure_minimal mem_Collect_eq) then have "frontier ?T \ T" by (simp add: frontier_subset_eq assms subset_trans [OF frontier_Un_subset]) then show "closed ?T" using frontier_subset_eq by fastforce have "UNIV \ (S \ T) \ \(components(- (S \ T)))" using Union_components by blast also have "... \ ?S \ ?T" proof - have "C \ components (-(S \ T)) \ frontier C \ S \ C \ components (-(S \ T)) \ frontier C \ T" if "C \ components (- (S \ T))" "C \ {}" for C using * [OF that] that by clarify (metis (no_types, lifting) UnE \closed S\ \closed T\ closed_Un disjoint_iff_not_equal frontier_of_components_closed_complement subsetCE) then show ?thesis by blast qed finally show "UNIV \ ?S \ ?T" . have "\{C \ components (- (S \ T)). frontier C \ S} \ \{C \ components (- (S \ T)). frontier C \ T} \ - (S \ T)" using in_components_subset by fastforce moreover have "\{C \ components (- (S \ T)). frontier C \ S} \ \{C \ components (- (S \ T)). frontier C \ T} = {}" proof - have "C \ C' = {}" if "C \ components (- (S \ T))" "frontier C \ S" "C' \ components (- (S \ T))" "frontier C' \ T" for C C' proof - have NUN: "- S \ - T \ UNIV" using \T \ {}\ by blast have "C \ C'" proof assume "C = C'" with that have "frontier C' \ S \ T" by simp also have "... = {}" using \S \ T = {}\ by blast finally have "C' = {} \ C' = UNIV" using frontier_eq_empty by auto then show False using \C = C'\ NUN that by (force simp: dest: in_components_nonempty in_components_subset) qed with that show ?thesis by (simp add: components_nonoverlap [of _ "-(S \ T)"]) qed then show ?thesis by blast qed ultimately show "?S \ ?T = {}" using ST by blast show "?S \ {}" "?T \ {}" using \S \ {}\ \T \ {}\ by blast+ qed then show False by (metis Compl_disjoint connected_UNIV compl_bot_eq compl_unique connected_closedD inf_sup_absorb sup_compl_top_left1 top.extremum_uniqueI) qed proposition separation_by_component_open: fixes S :: "'a :: euclidean_space set" assumes "open S" and non: "\ connected(- S)" obtains C where "C \ components S" "\ connected(- C)" proof - obtain T U where "closed T" "closed U" and TU: "T \ U = - S" "T \ U = {}" "T \ {}" "U \ {}" using assms by (auto simp: connected_closed_set closed_def) then obtain C where C: "C \ components(-(T \ U))" "C \ {}" and "frontier C \ T \ {}" "frontier C \ U \ {}" using separation_by_component_open_aux [OF \closed T\ \closed U\ \T \ U = {}\] by force show "thesis" proof show "C \ components S" using C(1) TU(1) by auto show "\ connected (- C)" proof assume "connected (- C)" then have "connected (frontier C)" using connected_frontier_simple [of C] \C \ components S\ in_components_connected by blast then show False unfolding connected_closed by (metis C(1) TU(2) \closed T\ \closed U\ \frontier C \ T \ {}\ \frontier C \ U \ {}\ closed_Un frontier_of_components_closed_complement inf_bot_right inf_commute) qed qed qed lemma separation_by_Un_open: fixes S :: "'a :: euclidean_space set" assumes "open S" "open T" "S \ T = {}" and cS: "connected(-S)" and cT: "connected(-T)" shows "connected(- (S \ T))" using assms unicoherent_UNIV unfolding unicoherent_def by force lemma nonseparation_by_component_eq: fixes S :: "'a :: euclidean_space set" assumes "open S \ closed S" shows "((\C \ components S. connected(-C)) \ connected(- S))" (is "?lhs = ?rhs") proof assume ?lhs with assms show ?rhs by (meson separation_by_component_closed separation_by_component_open) next assume ?rhs with assms show ?lhs using component_complement_connected by force qed text\Another interesting equivalent of an inessential mapping into C-{0}\ proposition inessential_eq_extensible: fixes f :: "'a::euclidean_space \ complex" assumes "closed S" shows "(\a. homotopic_with_canon (\h. True) S (-{0}) f (\t. a)) \ (\g. continuous_on UNIV g \ (\x \ S. g x = f x) \ (\x. g x \ 0))" (is "?lhs = ?rhs") proof assume ?lhs then obtain a where a: "homotopic_with_canon (\h. True) S (-{0}) f (\t. a)" .. show ?rhs proof (cases "S = {}") case True with a show ?thesis by force next case False have anr: "ANR (-{0::complex})" by (simp add: ANR_delete open_Compl open_imp_ANR) obtain g where contg: "continuous_on UNIV g" and gim: "g ` UNIV \ -{0}" and gf: "\x. x \ S \ g x = f x" proof (rule Borsuk_homotopy_extension_homotopic [OF _ _ continuous_on_const _ homotopic_with_symD [OF a]]) show "closedin (top_of_set UNIV) S" using assms by auto show "range (\t. a) \ - {0}" using a homotopic_with_imp_subset2 False by blast qed (use anr that in \force+\) then show ?thesis by force qed next assume ?rhs then obtain g where contg: "continuous_on UNIV g" and gf: "\x. x \ S \ g x = f x" and non0: "\x. g x \ 0" by metis obtain h k::"'a\'a" where hk: "homeomorphism (ball 0 1) UNIV h k" using homeomorphic_ball01_UNIV homeomorphic_def by blast then have "continuous_on (ball 0 1) (g \ h)" by (meson contg continuous_on_compose continuous_on_subset homeomorphism_cont1 top_greatest) then obtain j where contj: "continuous_on (ball 0 1) j" and j: "\z. z \ ball 0 1 \ exp(j z) = (g \ h) z" by (metis (mono_tags, hide_lams) continuous_logarithm_on_ball comp_apply non0) have [simp]: "\x. x \ S \ h (k x) = x" using hk homeomorphism_apply2 by blast have "\\. continuous_on S \\ (\x\S. f x = exp (\ x))" proof (intro exI conjI ballI) show "continuous_on S (j \ k)" proof (rule continuous_on_compose) show "continuous_on S k" by (meson continuous_on_subset hk homeomorphism_cont2 top_greatest) show "continuous_on (k ` S) j" apply (rule continuous_on_subset [OF contj]) using homeomorphism_image2 [OF hk] continuous_on_subset [OF contj] by blast qed show "f x = exp ((j \ k) x)" if "x \ S" for x proof - have "f x = (g \ h) (k x)" by (simp add: gf that) also have "... = exp (j (k x))" by (metis rangeI homeomorphism_image2 [OF hk] j) finally show ?thesis by simp qed qed then show ?lhs by (simp add: inessential_eq_continuous_logarithm) qed lemma inessential_on_clopen_Union: fixes \ :: "'a::euclidean_space set set" assumes T: "path_connected T" and "\S. S \ \ \ closedin (top_of_set (\\)) S" and "\S. S \ \ \ openin (top_of_set (\\)) S" and hom: "\S. S \ \ \ \a. homotopic_with_canon (\x. True) S T f (\x. a)" obtains a where "homotopic_with_canon (\x. True) (\\) T f (\x. a)" proof (cases "\\ = {}") case True with that show ?thesis by force next case False then obtain C where "C \ \" "C \ {}" by blast then obtain a where clo: "closedin (top_of_set (\\)) C" and ope: "openin (top_of_set (\\)) C" and "homotopic_with_canon (\x. True) C T f (\x. a)" using assms by blast with \C \ {}\ have "f ` C \ T" "a \ T" using homotopic_with_imp_subset1 homotopic_with_imp_subset2 by blast+ have "homotopic_with_canon (\x. True) (\\) T f (\x. a)" proof (rule homotopic_on_clopen_Union) show "\S. S \ \ \ closedin (top_of_set (\\)) S" "\S. S \ \ \ openin (top_of_set (\\)) S" by (simp_all add: assms) show "homotopic_with_canon (\x. True) S T f (\x. a)" if "S \ \" for S proof (cases "S = {}") case True then show ?thesis by auto next case False then obtain b where "b \ S" by blast obtain c where c: "homotopic_with_canon (\x. True) S T f (\x. c)" using \S \ \\ hom by blast then have "c \ T" using \b \ S\ homotopic_with_imp_subset2 by blast then have "homotopic_with_canon (\x. True) S T (\x. a) (\x. c)" using T \a \ T\ homotopic_constant_maps path_connected_component by (simp add: homotopic_constant_maps path_connected_component) then show ?thesis using c homotopic_with_symD homotopic_with_trans by blast qed qed then show ?thesis .. qed proposition Janiszewski_dual: fixes S :: "complex set" assumes "compact S" "compact T" "connected S" "connected T" "connected(- (S \ T))" shows "connected(S \ T)" proof - have ST: "compact (S \ T)" by (simp add: assms compact_Un) with Borsukian_imp_unicoherent [of "S \ T"] ST assms show ?thesis by (auto simp: closed_subset compact_imp_closed Borsukian_separation_compact unicoherent_def) qed end diff --git a/src/HOL/Analysis/Measure_Space.thy b/src/HOL/Analysis/Measure_Space.thy --- a/src/HOL/Analysis/Measure_Space.thy +++ b/src/HOL/Analysis/Measure_Space.thy @@ -1,3883 +1,3885 @@ (* Title: HOL/Analysis/Measure_Space.thy Author: Lawrence C Paulson Author: Johannes Hölzl, TU München Author: Armin Heller, TU München *) section \Measure Spaces\ theory Measure_Space imports Measurable "HOL-Library.Extended_Nonnegative_Real" begin subsection\<^marker>\tag unimportant\ "Relate extended reals and the indicator function" lemma suminf_cmult_indicator: fixes f :: "nat \ ennreal" assumes "disjoint_family A" "x \ A i" shows "(\n. f n * indicator (A n) x) = f i" proof - have **: "\n. f n * indicator (A n) x = (if n = i then f n else 0 :: ennreal)" using \x \ A i\ assms unfolding disjoint_family_on_def indicator_def by auto then have "\n. (\jn. n \ UNIV \ (if i < n then f i else 0) \ y" from this[of "Suc i"] show "f i \ y" by auto qed (insert assms, simp) ultimately show ?thesis using assms by (subst suminf_eq_SUP) (auto simp: indicator_def) qed lemma suminf_indicator: assumes "disjoint_family A" shows "(\n. indicator (A n) x :: ennreal) = indicator (\i. A i) x" proof cases assume *: "x \ (\i. A i)" then obtain i where "x \ A i" by auto from suminf_cmult_indicator[OF assms(1), OF \x \ A i\, of "\k. 1"] show ?thesis using * by simp qed simp lemma sum_indicator_disjoint_family: fixes f :: "'d \ 'e::semiring_1" assumes d: "disjoint_family_on A P" and "x \ A j" and "finite P" and "j \ P" shows "(\i\P. f i * indicator (A i) x) = f j" proof - have "P \ {i. x \ A i} = {j}" using d \x \ A j\ \j \ P\ unfolding disjoint_family_on_def by auto thus ?thesis unfolding indicator_def by (simp add: if_distrib sum.If_cases[OF \finite P\]) qed text \ The type for emeasure spaces is already defined in \<^theory>\HOL-Analysis.Sigma_Algebra\, as it is also used to represent sigma algebras (with an arbitrary emeasure). \ subsection\<^marker>\tag unimportant\ "Extend binary sets" lemma LIMSEQ_binaryset: assumes f: "f {} = 0" shows "(\n. \i f A + f B" proof - have "(\n. \i < Suc (Suc n). f (binaryset A B i)) = (\n. f A + f B)" proof fix n show "(\i < Suc (Suc n). f (binaryset A B i)) = f A + f B" by (induct n) (auto simp add: binaryset_def f) qed moreover have "... \ f A + f B" by (rule tendsto_const) ultimately have "(\n. \i< Suc (Suc n). f (binaryset A B i)) \ f A + f B" by metis hence "(\n. \i< n+2. f (binaryset A B i)) \ f A + f B" by simp thus ?thesis by (rule LIMSEQ_offset [where k=2]) qed lemma binaryset_sums: assumes f: "f {} = 0" shows "(\n. f (binaryset A B n)) sums (f A + f B)" by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f] atLeast0LessThan) lemma suminf_binaryset_eq: fixes f :: "'a set \ 'b::{comm_monoid_add, t2_space}" shows "f {} = 0 \ (\n. f (binaryset A B n)) = f A + f B" by (metis binaryset_sums sums_unique) subsection\<^marker>\tag unimportant\ \Properties of a premeasure \<^term>\\\\ text \ The definitions for \<^const>\positive\ and \<^const>\countably_additive\ should be here, by they are necessary to define \<^typ>\'a measure\ in \<^theory>\HOL-Analysis.Sigma_Algebra\. \ definition subadditive where "subadditive M f \ (\x\M. \y\M. x \ y = {} \ f (x \ y) \ f x + f y)" lemma subadditiveD: "subadditive M f \ x \ y = {} \ x \ M \ y \ M \ f (x \ y) \ f x + f y" by (auto simp add: subadditive_def) definition countably_subadditive where "countably_subadditive M f \ (\A. range A \ M \ disjoint_family A \ (\i. A i) \ M \ (f (\i. A i) \ (\i. f (A i))))" lemma (in ring_of_sets) countably_subadditive_subadditive: fixes f :: "'a set \ ennreal" assumes f: "positive M f" and cs: "countably_subadditive M f" shows "subadditive M f" proof (auto simp add: subadditive_def) fix x y assume x: "x \ M" and y: "y \ M" and "x \ y = {}" hence "disjoint_family (binaryset x y)" by (auto simp add: disjoint_family_on_def binaryset_def) hence "range (binaryset x y) \ M \ (\i. binaryset x y i) \ M \ f (\i. binaryset x y i) \ (\ n. f (binaryset x y n))" using cs by (auto simp add: countably_subadditive_def) hence "{x,y,{}} \ M \ x \ y \ M \ f (x \ y) \ (\ n. f (binaryset x y n))" by (simp add: range_binaryset_eq UN_binaryset_eq) thus "f (x \ y) \ f x + f y" using f x y by (auto simp add: Un o_def suminf_binaryset_eq positive_def) qed definition additive where "additive M \ \ (\x\M. \y\M. x \ y = {} \ \ (x \ y) = \ x + \ y)" definition increasing where "increasing M \ \ (\x\M. \y\M. x \ y \ \ x \ \ y)" lemma positiveD1: "positive M f \ f {} = 0" by (auto simp: positive_def) lemma positiveD_empty: "positive M f \ f {} = 0" by (auto simp add: positive_def) lemma additiveD: "additive M f \ x \ y = {} \ x \ M \ y \ M \ f (x \ y) = f x + f y" by (auto simp add: additive_def) lemma increasingD: "increasing M f \ x \ y \ x\M \ y\M \ f x \ f y" by (auto simp add: increasing_def) lemma countably_additiveI[case_names countably]: "(\A. range A \ M \ disjoint_family A \ (\i. A i) \ M \ (\i. f (A i)) = f (\i. A i)) \ countably_additive M f" by (simp add: countably_additive_def) lemma (in ring_of_sets) disjointed_additive: assumes f: "positive M f" "additive M f" and A: "range A \ M" "incseq A" shows "(\i\n. f (disjointed A i)) = f (A n)" proof (induct n) case (Suc n) then have "(\i\Suc n. f (disjointed A i)) = f (A n) + f (disjointed A (Suc n))" by simp also have "\ = f (A n \ disjointed A (Suc n))" using A by (subst f(2)[THEN additiveD]) (auto simp: disjointed_mono) also have "A n \ disjointed A (Suc n) = A (Suc n)" using \incseq A\ by (auto dest: incseq_SucD simp: disjointed_mono) finally show ?case . qed simp lemma (in ring_of_sets) additive_sum: fixes A:: "'i \ 'a set" assumes f: "positive M f" and ad: "additive M f" and "finite S" and A: "A`S \ M" and disj: "disjoint_family_on A S" shows "(\i\S. f (A i)) = f (\i\S. A i)" using \finite S\ disj A proof induct case empty show ?case using f by (simp add: positive_def) next case (insert s S) then have "A s \ (\i\S. A i) = {}" by (auto simp add: disjoint_family_on_def neq_iff) moreover have "A s \ M" using insert by blast moreover have "(\i\S. A i) \ M" using insert \finite S\ by auto ultimately have "f (A s \ (\i\S. A i)) = f (A s) + f(\i\S. A i)" using ad UNION_in_sets A by (auto simp add: additive_def) with insert show ?case using ad disjoint_family_on_mono[of S "insert s S" A] by (auto simp add: additive_def subset_insertI) qed lemma (in ring_of_sets) additive_increasing: fixes f :: "'a set \ ennreal" assumes posf: "positive M f" and addf: "additive M f" shows "increasing M f" proof (auto simp add: increasing_def) fix x y assume xy: "x \ M" "y \ M" "x \ y" then have "y - x \ M" by auto then have "f x + 0 \ f x + f (y-x)" by (intro add_left_mono zero_le) also have "... = f (x \ (y-x))" using addf by (auto simp add: additive_def) (metis Diff_disjoint Un_Diff_cancel Diff xy(1,2)) also have "... = f y" by (metis Un_Diff_cancel Un_absorb1 xy(3)) finally show "f x \ f y" by simp qed lemma (in ring_of_sets) subadditive: fixes f :: "'a set \ ennreal" assumes f: "positive M f" "additive M f" and A: "A`S \ M" and S: "finite S" shows "f (\i\S. A i) \ (\i\S. f (A i))" using S A proof (induct S) case empty thus ?case using f by (auto simp: positive_def) next case (insert x F) hence in_M: "A x \ M" "(\i\F. A i) \ M" "(\i\F. A i) - A x \ M" using A by force+ have subs: "(\i\F. A i) - A x \ (\i\F. A i)" by auto have "(\i\(insert x F). A i) = A x \ ((\i\F. A i) - A x)" by auto hence "f (\i\(insert x F). A i) = f (A x \ ((\i\F. A i) - A x))" by simp also have "\ = f (A x) + f ((\i\F. A i) - A x)" using f(2) by (rule additiveD) (insert in_M, auto) also have "\ \ f (A x) + f (\i\F. A i)" using additive_increasing[OF f] in_M subs by (auto simp: increasing_def intro: add_left_mono) also have "\ \ f (A x) + (\i\F. f (A i))" using insert by (auto intro: add_left_mono) finally show "f (\i\(insert x F). A i) \ (\i\(insert x F). f (A i))" using insert by simp qed lemma (in ring_of_sets) countably_additive_additive: fixes f :: "'a set \ ennreal" assumes posf: "positive M f" and ca: "countably_additive M f" shows "additive M f" proof (auto simp add: additive_def) fix x y assume x: "x \ M" and y: "y \ M" and "x \ y = {}" hence "disjoint_family (binaryset x y)" by (auto simp add: disjoint_family_on_def binaryset_def) hence "range (binaryset x y) \ M \ (\i. binaryset x y i) \ M \ f (\i. binaryset x y i) = (\ n. f (binaryset x y n))" using ca by (simp add: countably_additive_def) hence "{x,y,{}} \ M \ x \ y \ M \ f (x \ y) = (\n. f (binaryset x y n))" by (simp add: range_binaryset_eq UN_binaryset_eq) thus "f (x \ y) = f x + f y" using posf x y by (auto simp add: Un suminf_binaryset_eq positive_def) qed lemma (in algebra) increasing_additive_bound: fixes A:: "nat \ 'a set" and f :: "'a set \ ennreal" assumes f: "positive M f" and ad: "additive M f" and inc: "increasing M f" and A: "range A \ M" and disj: "disjoint_family A" shows "(\i. f (A i)) \ f \" proof (safe intro!: suminf_le_const) fix N note disj_N = disjoint_family_on_mono[OF _ disj, of "{..ii\{.. f \" using space_closed A by (intro increasingD[OF inc] finite_UN) auto finally show "(\i f \" by simp qed (insert f A, auto simp: positive_def) lemma (in ring_of_sets) countably_additiveI_finite: fixes \ :: "'a set \ ennreal" assumes "finite \" "positive M \" "additive M \" shows "countably_additive M \" proof (rule countably_additiveI) fix F :: "nat \ 'a set" assume F: "range F \ M" "(\i. F i) \ M" and disj: "disjoint_family F" have "\i\{i. F i \ {}}. \x. x \ F i" by auto from bchoice[OF this] obtain f where f: "\i. F i \ {} \ f i \ F i" by auto have inj_f: "inj_on f {i. F i \ {}}" proof (rule inj_onI, simp) fix i j a b assume *: "f i = f j" "F i \ {}" "F j \ {}" then have "f i \ F i" "f j \ F j" using f by force+ with disj * show "i = j" by (auto simp: disjoint_family_on_def) qed have "finite (\i. F i)" by (metis F(2) assms(1) infinite_super sets_into_space) have F_subset: "{i. \ (F i) \ 0} \ {i. F i \ {}}" by (auto simp: positiveD_empty[OF \positive M \\]) moreover have fin_not_empty: "finite {i. F i \ {}}" proof (rule finite_imageD) from f have "f`{i. F i \ {}} \ (\i. F i)" by auto then show "finite (f`{i. F i \ {}})" by (rule finite_subset) fact qed fact ultimately have fin_not_0: "finite {i. \ (F i) \ 0}" by (rule finite_subset) have disj_not_empty: "disjoint_family_on F {i. F i \ {}}" using disj by (auto simp: disjoint_family_on_def) from fin_not_0 have "(\i. \ (F i)) = (\i | \ (F i) \ 0. \ (F i))" by (rule suminf_finite) auto also have "\ = (\i | F i \ {}. \ (F i))" using fin_not_empty F_subset by (rule sum.mono_neutral_left) auto also have "\ = \ (\i\{i. F i \ {}}. F i)" using \positive M \\ \additive M \\ fin_not_empty disj_not_empty F by (intro additive_sum) auto also have "\ = \ (\i. F i)" by (rule arg_cong[where f=\]) auto finally show "(\i. \ (F i)) = \ (\i. F i)" . qed lemma (in ring_of_sets) countably_additive_iff_continuous_from_below: fixes f :: "'a set \ ennreal" assumes f: "positive M f" "additive M f" shows "countably_additive M f \ (\A. range A \ M \ incseq A \ (\i. A i) \ M \ (\i. f (A i)) \ f (\i. A i))" unfolding countably_additive_def proof safe assume count_sum: "\A. range A \ M \ disjoint_family A \ \(A ` UNIV) \ M \ (\i. f (A i)) = f (\(A ` UNIV))" fix A :: "nat \ 'a set" assume A: "range A \ M" "incseq A" "(\i. A i) \ M" then have dA: "range (disjointed A) \ M" by (auto simp: range_disjointed_sets) with count_sum[THEN spec, of "disjointed A"] A(3) have f_UN: "(\i. f (disjointed A i)) = f (\i. A i)" by (auto simp: UN_disjointed_eq disjoint_family_disjointed) moreover have "(\n. (\i (\i. f (disjointed A i))" using f(1)[unfolded positive_def] dA by (auto intro!: summable_LIMSEQ) from LIMSEQ_Suc[OF this] have "(\n. (\i\n. f (disjointed A i))) \ (\i. f (disjointed A i))" unfolding lessThan_Suc_atMost . moreover have "\n. (\i\n. f (disjointed A i)) = f (A n)" using disjointed_additive[OF f A(1,2)] . ultimately show "(\i. f (A i)) \ f (\i. A i)" by simp next assume cont: "\A. range A \ M \ incseq A \ (\i. A i) \ M \ (\i. f (A i)) \ f (\i. A i)" fix A :: "nat \ 'a set" assume A: "range A \ M" "disjoint_family A" "(\i. A i) \ M" have *: "(\n. (\ii. A i)" by auto have "(\n. f (\i f (\i. A i)" proof (unfold *[symmetric], intro cont[rule_format]) show "range (\i. \i M" "(\i. \i M" using A * by auto qed (force intro!: incseq_SucI) moreover have "\n. f (\iii. f (A i)) sums f (\i. A i)" unfolding sums_def by simp from sums_unique[OF this] show "(\i. f (A i)) = f (\i. A i)" by simp qed lemma (in ring_of_sets) continuous_from_above_iff_empty_continuous: fixes f :: "'a set \ ennreal" assumes f: "positive M f" "additive M f" shows "(\A. range A \ M \ decseq A \ (\i. A i) \ M \ (\i. f (A i) \ \) \ (\i. f (A i)) \ f (\i. A i)) \ (\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i) \ \) \ (\i. f (A i)) \ 0)" proof safe assume cont: "(\A. range A \ M \ decseq A \ (\i. A i) \ M \ (\i. f (A i) \ \) \ (\i. f (A i)) \ f (\i. A i))" fix A :: "nat \ 'a set" assume A: "range A \ M" "decseq A" "(\i. A i) = {}" "\i. f (A i) \ \" with cont[THEN spec, of A] show "(\i. f (A i)) \ 0" using \positive M f\[unfolded positive_def] by auto next assume cont: "\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i) \ \) \ (\i. f (A i)) \ 0" fix A :: "nat \ 'a set" assume A: "range A \ M" "decseq A" "(\i. A i) \ M" "\i. f (A i) \ \" have f_mono: "\a b. a \ M \ b \ M \ a \ b \ f a \ f b" using additive_increasing[OF f] unfolding increasing_def by simp have decseq_fA: "decseq (\i. f (A i))" using A by (auto simp: decseq_def intro!: f_mono) have decseq: "decseq (\i. A i - (\i. A i))" using A by (auto simp: decseq_def) then have decseq_f: "decseq (\i. f (A i - (\i. A i)))" using A unfolding decseq_def by (auto intro!: f_mono Diff) have "f (\x. A x) \ f (A 0)" using A by (auto intro!: f_mono) then have f_Int_fin: "f (\x. A x) \ \" using A by (auto simp: top_unique) { fix i have "f (A i - (\i. A i)) \ f (A i)" using A by (auto intro!: f_mono) then have "f (A i - (\i. A i)) \ \" using A by (auto simp: top_unique) } note f_fin = this have "(\i. f (A i - (\i. A i))) \ 0" proof (intro cont[rule_format, OF _ decseq _ f_fin]) show "range (\i. A i - (\i. A i)) \ M" "(\i. A i - (\i. A i)) = {}" using A by auto qed from INF_Lim[OF decseq_f this] have "(INF n. f (A n - (\i. A i))) = 0" . moreover have "(INF n. f (\i. A i)) = f (\i. A i)" by auto ultimately have "(INF n. f (A n - (\i. A i)) + f (\i. A i)) = 0 + f (\i. A i)" using A(4) f_fin f_Int_fin by (subst INF_ennreal_add_const) (auto simp: decseq_f) moreover { fix n have "f (A n - (\i. A i)) + f (\i. A i) = f ((A n - (\i. A i)) \ (\i. A i))" using A by (subst f(2)[THEN additiveD]) auto also have "(A n - (\i. A i)) \ (\i. A i) = A n" by auto finally have "f (A n - (\i. A i)) + f (\i. A i) = f (A n)" . } ultimately have "(INF n. f (A n)) = f (\i. A i)" by simp with LIMSEQ_INF[OF decseq_fA] show "(\i. f (A i)) \ f (\i. A i)" by simp qed lemma (in ring_of_sets) empty_continuous_imp_continuous_from_below: fixes f :: "'a set \ ennreal" assumes f: "positive M f" "additive M f" "\A\M. f A \ \" assumes cont: "\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i)) \ 0" assumes A: "range A \ M" "incseq A" "(\i. A i) \ M" shows "(\i. f (A i)) \ f (\i. A i)" proof - from A have "(\i. f ((\i. A i) - A i)) \ 0" by (intro cont[rule_format]) (auto simp: decseq_def incseq_def) moreover { fix i have "f ((\i. A i) - A i \ A i) = f ((\i. A i) - A i) + f (A i)" using A by (intro f(2)[THEN additiveD]) auto also have "((\i. A i) - A i) \ A i = (\i. A i)" by auto finally have "f ((\i. A i) - A i) = f (\i. A i) - f (A i)" using f(3)[rule_format, of "A i"] A by (auto simp: ennreal_add_diff_cancel subset_eq) } moreover have "\\<^sub>F i in sequentially. f (A i) \ f (\i. A i)" using increasingD[OF additive_increasing[OF f(1, 2)], of "A _" "\i. A i"] A by (auto intro!: always_eventually simp: subset_eq) ultimately show "(\i. f (A i)) \ f (\i. A i)" by (auto intro: ennreal_tendsto_const_minus) qed lemma (in ring_of_sets) empty_continuous_imp_countably_additive: fixes f :: "'a set \ ennreal" assumes f: "positive M f" "additive M f" and fin: "\A\M. f A \ \" assumes cont: "\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i)) \ 0" shows "countably_additive M f" using countably_additive_iff_continuous_from_below[OF f] using empty_continuous_imp_continuous_from_below[OF f fin] cont by blast subsection\<^marker>\tag unimportant\ \Properties of \<^const>\emeasure\\ lemma emeasure_positive: "positive (sets M) (emeasure M)" by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def) lemma emeasure_empty[simp, intro]: "emeasure M {} = 0" using emeasure_positive[of M] by (simp add: positive_def) lemma emeasure_single_in_space: "emeasure M {x} \ 0 \ x \ space M" using emeasure_notin_sets[of "{x}" M] by (auto dest: sets.sets_into_space zero_less_iff_neq_zero[THEN iffD2]) lemma emeasure_countably_additive: "countably_additive (sets M) (emeasure M)" by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def) lemma suminf_emeasure: "range A \ sets M \ disjoint_family A \ (\i. emeasure M (A i)) = emeasure M (\i. A i)" using sets.countable_UN[of A UNIV M] emeasure_countably_additive[of M] by (simp add: countably_additive_def) lemma sums_emeasure: "disjoint_family F \ (\i. F i \ sets M) \ (\i. emeasure M (F i)) sums emeasure M (\i. F i)" unfolding sums_iff by (intro conjI suminf_emeasure) auto lemma emeasure_additive: "additive (sets M) (emeasure M)" by (metis sets.countably_additive_additive emeasure_positive emeasure_countably_additive) lemma plus_emeasure: "a \ sets M \ b \ sets M \ a \ b = {} \ emeasure M a + emeasure M b = emeasure M (a \ b)" using additiveD[OF emeasure_additive] .. lemma emeasure_Un: "A \ sets M \ B \ sets M \ emeasure M (A \ B) = emeasure M A + emeasure M (B - A)" using plus_emeasure[of A M "B - A"] by auto lemma emeasure_Un_Int: assumes "A \ sets M" "B \ sets M" shows "emeasure M A + emeasure M B = emeasure M (A \ B) + emeasure M (A \ B)" proof - have "A = (A-B) \ (A \ B)" by auto then have "emeasure M A = emeasure M (A-B) + emeasure M (A \ B)" by (metis Diff_Diff_Int Diff_disjoint assms plus_emeasure sets.Diff) moreover have "A \ B = (A-B) \ B" by auto then have "emeasure M (A \ B) = emeasure M (A-B) + emeasure M B" by (metis Diff_disjoint Int_commute assms plus_emeasure sets.Diff) ultimately show ?thesis by (metis add.assoc add.commute) qed lemma sum_emeasure: "F`I \ sets M \ disjoint_family_on F I \ finite I \ (\i\I. emeasure M (F i)) = emeasure M (\i\I. F i)" by (metis sets.additive_sum emeasure_positive emeasure_additive) lemma emeasure_mono: "a \ b \ b \ sets M \ emeasure M a \ emeasure M b" by (metis zero_le sets.additive_increasing emeasure_additive emeasure_notin_sets emeasure_positive increasingD) lemma emeasure_space: "emeasure M A \ emeasure M (space M)" by (metis emeasure_mono emeasure_notin_sets sets.sets_into_space sets.top zero_le) lemma emeasure_Diff: assumes finite: "emeasure M B \ \" and [measurable]: "A \ sets M" "B \ sets M" and "B \ A" shows "emeasure M (A - B) = emeasure M A - emeasure M B" proof - have "(A - B) \ B = A" using \B \ A\ by auto then have "emeasure M A = emeasure M ((A - B) \ B)" by simp also have "\ = emeasure M (A - B) + emeasure M B" by (subst plus_emeasure[symmetric]) auto finally show "emeasure M (A - B) = emeasure M A - emeasure M B" using finite by simp qed lemma emeasure_compl: "s \ sets M \ emeasure M s \ \ \ emeasure M (space M - s) = emeasure M (space M) - emeasure M s" by (rule emeasure_Diff) (auto dest: sets.sets_into_space) lemma Lim_emeasure_incseq: "range A \ sets M \ incseq A \ (\i. (emeasure M (A i))) \ emeasure M (\i. A i)" using emeasure_countably_additive by (auto simp add: sets.countably_additive_iff_continuous_from_below emeasure_positive emeasure_additive) lemma incseq_emeasure: assumes "range B \ sets M" "incseq B" shows "incseq (\i. emeasure M (B i))" using assms by (auto simp: incseq_def intro!: emeasure_mono) lemma SUP_emeasure_incseq: assumes A: "range A \ sets M" "incseq A" shows "(SUP n. emeasure M (A n)) = emeasure M (\i. A i)" using LIMSEQ_SUP[OF incseq_emeasure, OF A] Lim_emeasure_incseq[OF A] by (simp add: LIMSEQ_unique) lemma decseq_emeasure: assumes "range B \ sets M" "decseq B" shows "decseq (\i. emeasure M (B i))" using assms by (auto simp: decseq_def intro!: emeasure_mono) lemma INF_emeasure_decseq: assumes A: "range A \ sets M" and "decseq A" and finite: "\i. emeasure M (A i) \ \" shows "(INF n. emeasure M (A n)) = emeasure M (\i. A i)" proof - have le_MI: "emeasure M (\i. A i) \ emeasure M (A 0)" using A by (auto intro!: emeasure_mono) hence *: "emeasure M (\i. A i) \ \" using finite[of 0] by (auto simp: top_unique) have "emeasure M (A 0) - (INF n. emeasure M (A n)) = (SUP n. emeasure M (A 0) - emeasure M (A n))" by (simp add: ennreal_INF_const_minus) also have "\ = (SUP n. emeasure M (A 0 - A n))" using A finite \decseq A\[unfolded decseq_def] by (subst emeasure_Diff) auto also have "\ = emeasure M (\i. A 0 - A i)" proof (rule SUP_emeasure_incseq) show "range (\n. A 0 - A n) \ sets M" using A by auto show "incseq (\n. A 0 - A n)" using \decseq A\ by (auto simp add: incseq_def decseq_def) qed also have "\ = emeasure M (A 0) - emeasure M (\i. A i)" using A finite * by (simp, subst emeasure_Diff) auto finally show ?thesis by (rule ennreal_minus_cancel[rotated 3]) (insert finite A, auto intro: INF_lower emeasure_mono) qed lemma INF_emeasure_decseq': assumes A: "\i. A i \ sets M" and "decseq A" and finite: "\i. emeasure M (A i) \ \" shows "(INF n. emeasure M (A n)) = emeasure M (\i. A i)" proof - from finite obtain i where i: "emeasure M (A i) < \" by (auto simp: less_top) have fin: "i \ j \ emeasure M (A j) < \" for j by (rule le_less_trans[OF emeasure_mono i]) (auto intro!: decseqD[OF \decseq A\] A) have "(INF n. emeasure M (A n)) = (INF n. emeasure M (A (n + i)))" proof (rule INF_eq) show "\j\UNIV. emeasure M (A (j + i)) \ emeasure M (A i')" for i' by (intro bexI[of _ i'] emeasure_mono decseqD[OF \decseq A\] A) auto qed auto also have "\ = emeasure M (INF n. (A (n + i)))" using A \decseq A\ fin by (intro INF_emeasure_decseq) (auto simp: decseq_def less_top) also have "(INF n. (A (n + i))) = (INF n. A n)" by (meson INF_eq UNIV_I assms(2) decseqD le_add1) finally show ?thesis . qed lemma emeasure_INT_decseq_subset: fixes F :: "nat \ 'a set" assumes I: "I \ {}" and F: "\i j. i \ I \ j \ I \ i \ j \ F j \ F i" assumes F_sets[measurable]: "\i. i \ I \ F i \ sets M" and fin: "\i. i \ I \ emeasure M (F i) \ \" shows "emeasure M (\i\I. F i) = (INF i\I. emeasure M (F i))" proof cases assume "finite I" have "(\i\I. F i) = F (Max I)" using I \finite I\ by (intro antisym INF_lower INF_greatest F) auto moreover have "(INF i\I. emeasure M (F i)) = emeasure M (F (Max I))" using I \finite I\ by (intro antisym INF_lower INF_greatest F emeasure_mono) auto ultimately show ?thesis by simp next assume "infinite I" define L where "L n = (LEAST i. i \ I \ i \ n)" for n have L: "L n \ I \ n \ L n" for n unfolding L_def proof (rule LeastI_ex) show "\x. x \ I \ n \ x" using \infinite I\ finite_subset[of I "{..< n}"] by (rule_tac ccontr) (auto simp: not_le) qed have L_eq[simp]: "i \ I \ L i = i" for i unfolding L_def by (intro Least_equality) auto have L_mono: "i \ j \ L i \ L j" for i j using L[of j] unfolding L_def by (intro Least_le) (auto simp: L_def) have "emeasure M (\i. F (L i)) = (INF i. emeasure M (F (L i)))" proof (intro INF_emeasure_decseq[symmetric]) show "decseq (\i. F (L i))" using L by (intro antimonoI F L_mono) auto qed (insert L fin, auto) also have "\ = (INF i\I. emeasure M (F i))" proof (intro antisym INF_greatest) show "i \ I \ (INF i. emeasure M (F (L i))) \ emeasure M (F i)" for i by (intro INF_lower2[of i]) auto qed (insert L, auto intro: INF_lower) also have "(\i. F (L i)) = (\i\I. F i)" proof (intro antisym INF_greatest) show "i \ I \ (\i. F (L i)) \ F i" for i by (intro INF_lower2[of i]) auto qed (insert L, auto) finally show ?thesis . qed lemma Lim_emeasure_decseq: assumes A: "range A \ sets M" "decseq A" and fin: "\i. emeasure M (A i) \ \" shows "(\i. emeasure M (A i)) \ emeasure M (\i. A i)" using LIMSEQ_INF[OF decseq_emeasure, OF A] using INF_emeasure_decseq[OF A fin] by simp lemma emeasure_lfp'[consumes 1, case_names cont measurable]: assumes "P M" assumes cont: "sup_continuous F" assumes *: "\M A. P M \ (\N. P N \ Measurable.pred N A) \ Measurable.pred M (F A)" shows "emeasure M {x\space M. lfp F x} = (SUP i. emeasure M {x\space M. (F ^^ i) (\x. False) x})" proof - have "emeasure M {x\space M. lfp F x} = emeasure M (\i. {x\space M. (F ^^ i) (\x. False) x})" using sup_continuous_lfp[OF cont] by (auto simp add: bot_fun_def intro!: arg_cong2[where f=emeasure]) moreover { fix i from \P M\ have "{x\space M. (F ^^ i) (\x. False) x} \ sets M" by (induct i arbitrary: M) (auto simp add: pred_def[symmetric] intro: *) } moreover have "incseq (\i. {x\space M. (F ^^ i) (\x. False) x})" proof (rule incseq_SucI) fix i have "(F ^^ i) (\x. False) \ (F ^^ (Suc i)) (\x. False)" proof (induct i) case 0 show ?case by (simp add: le_fun_def) next case Suc thus ?case using monoD[OF sup_continuous_mono[OF cont] Suc] by auto qed then show "{x \ space M. (F ^^ i) (\x. False) x} \ {x \ space M. (F ^^ Suc i) (\x. False) x}" by auto qed ultimately show ?thesis by (subst SUP_emeasure_incseq) auto qed lemma emeasure_lfp: assumes [simp]: "\s. sets (M s) = sets N" assumes cont: "sup_continuous F" "sup_continuous f" assumes meas: "\P. Measurable.pred N P \ Measurable.pred N (F P)" assumes iter: "\P s. Measurable.pred N P \ P \ lfp F \ emeasure (M s) {x\space N. F P x} = f (\s. emeasure (M s) {x\space N. P x}) s" shows "emeasure (M s) {x\space N. lfp F x} = lfp f s" proof (subst lfp_transfer_bounded[where \="\F s. emeasure (M s) {x\space N. F x}" and g=f and f=F and P="Measurable.pred N", symmetric]) fix C assume "incseq C" "\i. Measurable.pred N (C i)" then show "(\s. emeasure (M s) {x \ space N. (SUP i. C i) x}) = (SUP i. (\s. emeasure (M s) {x \ space N. C i x}))" unfolding SUP_apply[abs_def] by (subst SUP_emeasure_incseq) (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure]) qed (auto simp add: iter le_fun_def SUP_apply[abs_def] intro!: meas cont) lemma emeasure_subadditive_finite: "finite I \ A ` I \ sets M \ emeasure M (\i\I. A i) \ (\i\I. emeasure M (A i))" by (rule sets.subadditive[OF emeasure_positive emeasure_additive]) auto lemma emeasure_subadditive: "A \ sets M \ B \ sets M \ emeasure M (A \ B) \ emeasure M A + emeasure M B" using emeasure_subadditive_finite[of "{True, False}" "\True \ A | False \ B" M] by simp lemma emeasure_subadditive_countably: assumes "range f \ sets M" shows "emeasure M (\i. f i) \ (\i. emeasure M (f i))" proof - have "emeasure M (\i. f i) = emeasure M (\i. disjointed f i)" unfolding UN_disjointed_eq .. also have "\ = (\i. emeasure M (disjointed f i))" using sets.range_disjointed_sets[OF assms] suminf_emeasure[of "disjointed f"] by (simp add: disjoint_family_disjointed comp_def) also have "\ \ (\i. emeasure M (f i))" using sets.range_disjointed_sets[OF assms] assms by (auto intro!: suminf_le emeasure_mono disjointed_subset) finally show ?thesis . qed lemma emeasure_insert: assumes sets: "{x} \ sets M" "A \ sets M" and "x \ A" shows "emeasure M (insert x A) = emeasure M {x} + emeasure M A" proof - have "{x} \ A = {}" using \x \ A\ by auto from plus_emeasure[OF sets this] show ?thesis by simp qed lemma emeasure_insert_ne: "A \ {} \ {x} \ sets M \ A \ sets M \ x \ A \ emeasure M (insert x A) = emeasure M {x} + emeasure M A" by (rule emeasure_insert) lemma emeasure_eq_sum_singleton: assumes "finite S" "\x. x \ S \ {x} \ sets M" shows "emeasure M S = (\x\S. emeasure M {x})" using sum_emeasure[of "\x. {x}" S M] assms by (auto simp: disjoint_family_on_def subset_eq) lemma sum_emeasure_cover: assumes "finite S" and "A \ sets M" and br_in_M: "B ` S \ sets M" assumes A: "A \ (\i\S. B i)" assumes disj: "disjoint_family_on B S" shows "emeasure M A = (\i\S. emeasure M (A \ (B i)))" proof - have "(\i\S. emeasure M (A \ (B i))) = emeasure M (\i\S. A \ (B i))" proof (rule sum_emeasure) show "disjoint_family_on (\i. A \ B i) S" using \disjoint_family_on B S\ unfolding disjoint_family_on_def by auto qed (insert assms, auto) also have "(\i\S. A \ (B i)) = A" using A by auto finally show ?thesis by simp qed lemma emeasure_eq_0: "N \ sets M \ emeasure M N = 0 \ K \ N \ emeasure M K = 0" by (metis emeasure_mono order_eq_iff zero_le) lemma emeasure_UN_eq_0: assumes "\i::nat. emeasure M (N i) = 0" and "range N \ sets M" shows "emeasure M (\i. N i) = 0" proof - have "emeasure M (\i. N i) \ 0" using emeasure_subadditive_countably[OF assms(2)] assms(1) by simp then show ?thesis by (auto intro: antisym zero_le) qed lemma measure_eqI_finite: assumes [simp]: "sets M = Pow A" "sets N = Pow A" and "finite A" assumes eq: "\a. a \ A \ emeasure M {a} = emeasure N {a}" shows "M = N" proof (rule measure_eqI) fix X assume "X \ sets M" then have X: "X \ A" by auto then have "emeasure M X = (\a\X. emeasure M {a})" using \finite A\ by (subst emeasure_eq_sum_singleton) (auto dest: finite_subset) also have "\ = (\a\X. emeasure N {a})" using X eq by (auto intro!: sum.cong) also have "\ = emeasure N X" using X \finite A\ by (subst emeasure_eq_sum_singleton) (auto dest: finite_subset) finally show "emeasure M X = emeasure N X" . qed simp lemma measure_eqI_generator_eq: fixes M N :: "'a measure" and E :: "'a set set" and A :: "nat \ 'a set" assumes "Int_stable E" "E \ Pow \" and eq: "\X. X \ E \ emeasure M X = emeasure N X" and M: "sets M = sigma_sets \ E" and N: "sets N = sigma_sets \ E" and A: "range A \ E" "(\i. A i) = \" "\i. emeasure M (A i) \ \" shows "M = N" proof - let ?\ = "emeasure M" and ?\ = "emeasure N" interpret S: sigma_algebra \ "sigma_sets \ E" by (rule sigma_algebra_sigma_sets) fact have "space M = \" using sets.top[of M] sets.space_closed[of M] S.top S.space_closed \sets M = sigma_sets \ E\ by blast { fix F D assume "F \ E" and "?\ F \ \" then have [intro]: "F \ sigma_sets \ E" by auto have "?\ F \ \" using \?\ F \ \\ \F \ E\ eq by simp assume "D \ sets M" with \Int_stable E\ \E \ Pow \\ have "emeasure M (F \ D) = emeasure N (F \ D)" unfolding M proof (induct rule: sigma_sets_induct_disjoint) case (basic A) then have "F \ A \ E" using \Int_stable E\ \F \ E\ by (auto simp: Int_stable_def) then show ?case using eq by auto next case empty then show ?case by simp next case (compl A) then have **: "F \ (\ - A) = F - (F \ A)" and [intro]: "F \ A \ sigma_sets \ E" using \F \ E\ S.sets_into_space by (auto simp: M) have "?\ (F \ A) \ ?\ F" by (auto intro!: emeasure_mono simp: M N) then have "?\ (F \ A) \ \" using \?\ F \ \\ by (auto simp: top_unique) have "?\ (F \ A) \ ?\ F" by (auto intro!: emeasure_mono simp: M N) then have "?\ (F \ A) \ \" using \?\ F \ \\ by (auto simp: top_unique) then have "?\ (F \ (\ - A)) = ?\ F - ?\ (F \ A)" unfolding ** using \F \ A \ sigma_sets \ E\ by (auto intro!: emeasure_Diff simp: M N) also have "\ = ?\ F - ?\ (F \ A)" using eq \F \ E\ compl by simp also have "\ = ?\ (F \ (\ - A))" unfolding ** using \F \ A \ sigma_sets \ E\ \?\ (F \ A) \ \\ by (auto intro!: emeasure_Diff[symmetric] simp: M N) finally show ?case using \space M = \\ by auto next case (union A) then have "?\ (\x. F \ A x) = ?\ (\x. F \ A x)" by (subst (1 2) suminf_emeasure[symmetric]) (auto simp: disjoint_family_on_def subset_eq M N) with A show ?case by auto qed } note * = this show "M = N" proof (rule measure_eqI) show "sets M = sets N" using M N by simp have [simp, intro]: "\i. A i \ sets M" using A(1) by (auto simp: subset_eq M) fix F assume "F \ sets M" let ?D = "disjointed (\i. F \ A i)" from \space M = \\ have F_eq: "F = (\i. ?D i)" using \F \ sets M\[THEN sets.sets_into_space] A(2)[symmetric] by (auto simp: UN_disjointed_eq) have [simp, intro]: "\i. ?D i \ sets M" using sets.range_disjointed_sets[of "\i. F \ A i" M] \F \ sets M\ by (auto simp: subset_eq) have "disjoint_family ?D" by (auto simp: disjoint_family_disjointed) moreover have "(\i. emeasure M (?D i)) = (\i. emeasure N (?D i))" proof (intro arg_cong[where f=suminf] ext) fix i have "A i \ ?D i = ?D i" by (auto simp: disjointed_def) then show "emeasure M (?D i) = emeasure N (?D i)" using *[of "A i" "?D i", OF _ A(3)] A(1) by auto qed ultimately show "emeasure M F = emeasure N F" by (simp add: image_subset_iff \sets M = sets N\[symmetric] F_eq[symmetric] suminf_emeasure) qed qed lemma space_empty: "space M = {} \ M = count_space {}" by (rule measure_eqI) (simp_all add: space_empty_iff) lemma measure_eqI_generator_eq_countable: fixes M N :: "'a measure" and E :: "'a set set" and A :: "'a set set" assumes E: "Int_stable E" "E \ Pow \" "\X. X \ E \ emeasure M X = emeasure N X" and sets: "sets M = sigma_sets \ E" "sets N = sigma_sets \ E" and A: "A \ E" "(\A) = \" "countable A" "\a. a \ A \ emeasure M a \ \" shows "M = N" proof cases assume "\ = {}" have *: "sigma_sets \ E = sets (sigma \ E)" using E(2) by simp have "space M = \" "space N = \" using sets E(2) unfolding * by (auto dest: sets_eq_imp_space_eq simp del: sets_measure_of) then show "M = N" unfolding \\ = {}\ by (auto dest: space_empty) next assume "\ \ {}" with \\A = \\ have "A \ {}" by auto from this \countable A\ have rng: "range (from_nat_into A) = A" by (rule range_from_nat_into) show "M = N" proof (rule measure_eqI_generator_eq[OF E sets]) show "range (from_nat_into A) \ E" unfolding rng using \A \ E\ . show "(\i. from_nat_into A i) = \" unfolding rng using \\A = \\ . show "emeasure M (from_nat_into A i) \ \" for i using rng by (intro A) auto qed qed lemma measure_of_of_measure: "measure_of (space M) (sets M) (emeasure M) = M" proof (intro measure_eqI emeasure_measure_of_sigma) show "sigma_algebra (space M) (sets M)" .. show "positive (sets M) (emeasure M)" by (simp add: positive_def) show "countably_additive (sets M) (emeasure M)" by (simp add: emeasure_countably_additive) qed simp_all subsection \\\\-null sets\ definition\<^marker>\tag important\ null_sets :: "'a measure \ 'a set set" where "null_sets M = {N\sets M. emeasure M N = 0}" lemma null_setsD1[dest]: "A \ null_sets M \ emeasure M A = 0" by (simp add: null_sets_def) lemma null_setsD2[dest]: "A \ null_sets M \ A \ sets M" unfolding null_sets_def by simp lemma null_setsI[intro]: "emeasure M A = 0 \ A \ sets M \ A \ null_sets M" unfolding null_sets_def by simp interpretation null_sets: ring_of_sets "space M" "null_sets M" for M proof (rule ring_of_setsI) show "null_sets M \ Pow (space M)" using sets.sets_into_space by auto show "{} \ null_sets M" by auto fix A B assume null_sets: "A \ null_sets M" "B \ null_sets M" then have sets: "A \ sets M" "B \ sets M" by auto then have *: "emeasure M (A \ B) \ emeasure M A + emeasure M B" "emeasure M (A - B) \ emeasure M A" by (auto intro!: emeasure_subadditive emeasure_mono) then have "emeasure M B = 0" "emeasure M A = 0" using null_sets by auto with sets * show "A - B \ null_sets M" "A \ B \ null_sets M" by (auto intro!: antisym zero_le) qed lemma UN_from_nat_into: assumes I: "countable I" "I \ {}" shows "(\i\I. N i) = (\i. N (from_nat_into I i))" proof - have "(\i\I. N i) = \(N ` range (from_nat_into I))" using I by simp also have "\ = (\i. (N \ from_nat_into I) i)" by simp finally show ?thesis by simp qed lemma null_sets_UN': assumes "countable I" assumes "\i. i \ I \ N i \ null_sets M" shows "(\i\I. N i) \ null_sets M" proof cases assume "I = {}" then show ?thesis by simp next assume "I \ {}" show ?thesis proof (intro conjI CollectI null_setsI) show "(\i\I. N i) \ sets M" using assms by (intro sets.countable_UN') auto have "emeasure M (\i\I. N i) \ (\n. emeasure M (N (from_nat_into I n)))" unfolding UN_from_nat_into[OF \countable I\ \I \ {}\] using assms \I \ {}\ by (intro emeasure_subadditive_countably) (auto intro: from_nat_into) also have "(\n. emeasure M (N (from_nat_into I n))) = (\_. 0)" using assms \I \ {}\ by (auto intro: from_nat_into) finally show "emeasure M (\i\I. N i) = 0" by (intro antisym zero_le) simp qed qed lemma null_sets_UN[intro]: "(\i::'i::countable. N i \ null_sets M) \ (\i. N i) \ null_sets M" by (rule null_sets_UN') auto lemma null_set_Int1: assumes "B \ null_sets M" "A \ sets M" shows "A \ B \ null_sets M" proof (intro CollectI conjI null_setsI) show "emeasure M (A \ B) = 0" using assms by (intro emeasure_eq_0[of B _ "A \ B"]) auto qed (insert assms, auto) lemma null_set_Int2: assumes "B \ null_sets M" "A \ sets M" shows "B \ A \ null_sets M" using assms by (subst Int_commute) (rule null_set_Int1) lemma emeasure_Diff_null_set: assumes "B \ null_sets M" "A \ sets M" shows "emeasure M (A - B) = emeasure M A" proof - have *: "A - B = (A - (A \ B))" by auto have "A \ B \ null_sets M" using assms by (rule null_set_Int1) then show ?thesis unfolding * using assms by (subst emeasure_Diff) auto qed lemma null_set_Diff: assumes "B \ null_sets M" "A \ sets M" shows "B - A \ null_sets M" proof (intro CollectI conjI null_setsI) show "emeasure M (B - A) = 0" using assms by (intro emeasure_eq_0[of B _ "B - A"]) auto qed (insert assms, auto) lemma emeasure_Un_null_set: assumes "A \ sets M" "B \ null_sets M" shows "emeasure M (A \ B) = emeasure M A" proof - have *: "A \ B = A \ (B - A)" by auto have "B - A \ null_sets M" using assms(2,1) by (rule null_set_Diff) then show ?thesis unfolding * using assms by (subst plus_emeasure[symmetric]) auto qed lemma emeasure_Un': assumes "A \ sets M" "B \ sets M" "A \ B \ null_sets M" shows "emeasure M (A \ B) = emeasure M A + emeasure M B" proof - have "A \ B = A \ (B - A \ B)" by blast also have "emeasure M \ = emeasure M A + emeasure M (B - A \ B)" using assms by (subst plus_emeasure) auto also have "emeasure M (B - A \ B) = emeasure M B" using assms by (intro emeasure_Diff_null_set) auto finally show ?thesis . qed subsection \The almost everywhere filter (i.e.\ quantifier)\ definition\<^marker>\tag important\ ae_filter :: "'a measure \ 'a filter" where "ae_filter M = (INF N\null_sets M. principal (space M - N))" abbreviation almost_everywhere :: "'a measure \ ('a \ bool) \ bool" where "almost_everywhere M P \ eventually P (ae_filter M)" syntax "_almost_everywhere" :: "pttrn \ 'a \ bool \ bool" ("AE _ in _. _" [0,0,10] 10) translations "AE x in M. P" \ "CONST almost_everywhere M (\x. P)" abbreviation "set_almost_everywhere A M P \ AE x in M. x \ A \ P x" syntax "_set_almost_everywhere" :: "pttrn \ 'a set \ 'a \ bool \ bool" ("AE _\_ in _./ _" [0,0,0,10] 10) translations "AE x\A in M. P" \ "CONST set_almost_everywhere A M (\x. P)" lemma eventually_ae_filter: "eventually P (ae_filter M) \ (\N\null_sets M. {x \ space M. \ P x} \ N)" unfolding ae_filter_def by (subst eventually_INF_base) (auto simp: eventually_principal subset_eq) lemma AE_I': "N \ null_sets M \ {x\space M. \ P x} \ N \ (AE x in M. P x)" unfolding eventually_ae_filter by auto lemma AE_iff_null: assumes "{x\space M. \ P x} \ sets M" (is "?P \ sets M") shows "(AE x in M. P x) \ {x\space M. \ P x} \ null_sets M" proof assume "AE x in M. P x" then obtain N where N: "N \ sets M" "?P \ N" "emeasure M N = 0" unfolding eventually_ae_filter by auto have "emeasure M ?P \ emeasure M N" using assms N(1,2) by (auto intro: emeasure_mono) then have "emeasure M ?P = 0" unfolding \emeasure M N = 0\ by auto then show "?P \ null_sets M" using assms by auto next assume "?P \ null_sets M" with assms show "AE x in M. P x" by (auto intro: AE_I') qed lemma AE_iff_null_sets: "N \ sets M \ N \ null_sets M \ (AE x in M. x \ N)" using Int_absorb1[OF sets.sets_into_space, of N M] by (subst AE_iff_null) (auto simp: Int_def[symmetric]) lemma AE_not_in: "N \ null_sets M \ AE x in M. x \ N" by (metis AE_iff_null_sets null_setsD2) lemma AE_iff_measurable: "N \ sets M \ {x\space M. \ P x} = N \ (AE x in M. P x) \ emeasure M N = 0" using AE_iff_null[of _ P] by auto lemma AE_E[consumes 1]: assumes "AE x in M. P x" obtains N where "{x \ space M. \ P x} \ N" "emeasure M N = 0" "N \ sets M" using assms unfolding eventually_ae_filter by auto lemma AE_E2: assumes "AE x in M. P x" "{x\space M. P x} \ sets M" shows "emeasure M {x\space M. \ P x} = 0" (is "emeasure M ?P = 0") proof - have "{x\space M. \ P x} = space M - {x\space M. P x}" by auto with AE_iff_null[of M P] assms show ?thesis by auto qed lemma AE_E3: assumes "AE x in M. P x" obtains N where "\x. x \ space M - N \ P x" "N \ null_sets M" using assms unfolding eventually_ae_filter by auto lemma AE_I: assumes "{x \ space M. \ P x} \ N" "emeasure M N = 0" "N \ sets M" shows "AE x in M. P x" using assms unfolding eventually_ae_filter by auto lemma AE_mp[elim!]: assumes AE_P: "AE x in M. P x" and AE_imp: "AE x in M. P x \ Q x" shows "AE x in M. Q x" proof - from AE_P obtain A where P: "{x\space M. \ P x} \ A" and A: "A \ sets M" "emeasure M A = 0" by (auto elim!: AE_E) from AE_imp obtain B where imp: "{x\space M. P x \ \ Q x} \ B" and B: "B \ sets M" "emeasure M B = 0" by (auto elim!: AE_E) show ?thesis proof (intro AE_I) have "emeasure M (A \ B) \ 0" using emeasure_subadditive[of A M B] A B by auto then show "A \ B \ sets M" "emeasure M (A \ B) = 0" using A B by auto show "{x\space M. \ Q x} \ A \ B" using P imp by auto qed qed text \The next lemma is convenient to combine with a lemma whose conclusion is of the form \AE x in M. P x = Q x\: for such a lemma, there is no \[symmetric]\ variant, but using \AE_symmetric[OF...]\ will replace it.\ (* depricated replace by laws about eventually *) lemma shows AE_iffI: "AE x in M. P x \ AE x in M. P x \ Q x \ AE x in M. Q x" and AE_disjI1: "AE x in M. P x \ AE x in M. P x \ Q x" and AE_disjI2: "AE x in M. Q x \ AE x in M. P x \ Q x" and AE_conjI: "AE x in M. P x \ AE x in M. Q x \ AE x in M. P x \ Q x" and AE_conj_iff[simp]: "(AE x in M. P x \ Q x) \ (AE x in M. P x) \ (AE x in M. Q x)" by auto lemma AE_symmetric: assumes "AE x in M. P x = Q x" shows "AE x in M. Q x = P x" using assms by auto lemma AE_impI: "(P \ AE x in M. Q x) \ AE x in M. P \ Q x" by fastforce lemma AE_measure: assumes AE: "AE x in M. P x" and sets: "{x\space M. P x} \ sets M" (is "?P \ sets M") shows "emeasure M {x\space M. P x} = emeasure M (space M)" proof - from AE_E[OF AE] guess N . note N = this with sets have "emeasure M (space M) \ emeasure M (?P \ N)" by (intro emeasure_mono) auto also have "\ \ emeasure M ?P + emeasure M N" using sets N by (intro emeasure_subadditive) auto also have "\ = emeasure M ?P" using N by simp finally show "emeasure M ?P = emeasure M (space M)" using emeasure_space[of M "?P"] by auto qed lemma AE_space: "AE x in M. x \ space M" by (rule AE_I[where N="{}"]) auto lemma AE_I2[simp, intro]: "(\x. x \ space M \ P x) \ AE x in M. P x" using AE_space by force lemma AE_Ball_mp: "\x\space M. P x \ AE x in M. P x \ Q x \ AE x in M. Q x" by auto lemma AE_cong[cong]: "(\x. x \ space M \ P x \ Q x) \ (AE x in M. P x) \ (AE x in M. Q x)" by auto lemma AE_cong_simp: "M = N \ (\x. x \ space N =simp=> P x = Q x) \ (AE x in M. P x) \ (AE x in N. Q x)" by (auto simp: simp_implies_def) lemma AE_all_countable: "(AE x in M. \i. P i x) \ (\i::'i::countable. AE x in M. P i x)" proof assume "\i. AE x in M. P i x" from this[unfolded eventually_ae_filter Bex_def, THEN choice] obtain N where N: "\i. N i \ null_sets M" "\i. {x\space M. \ P i x} \ N i" by auto have "{x\space M. \ (\i. P i x)} \ (\i. {x\space M. \ P i x})" by auto also have "\ \ (\i. N i)" using N by auto finally have "{x\space M. \ (\i. P i x)} \ (\i. N i)" . moreover from N have "(\i. N i) \ null_sets M" by (intro null_sets_UN) auto ultimately show "AE x in M. \i. P i x" unfolding eventually_ae_filter by auto qed auto lemma AE_ball_countable: assumes [intro]: "countable X" shows "(AE x in M. \y\X. P x y) \ (\y\X. AE x in M. P x y)" proof assume "\y\X. AE x in M. P x y" from this[unfolded eventually_ae_filter Bex_def, THEN bchoice] obtain N where N: "\y. y \ X \ N y \ null_sets M" "\y. y \ X \ {x\space M. \ P x y} \ N y" by auto have "{x\space M. \ (\y\X. P x y)} \ (\y\X. {x\space M. \ P x y})" by auto also have "\ \ (\y\X. N y)" using N by auto finally have "{x\space M. \ (\y\X. P x y)} \ (\y\X. N y)" . moreover from N have "(\y\X. N y) \ null_sets M" by (intro null_sets_UN') auto ultimately show "AE x in M. \y\X. P x y" unfolding eventually_ae_filter by auto qed auto lemma AE_ball_countable': "(\N. N \ I \ AE x in M. P N x) \ countable I \ AE x in M. \N \ I. P N x" unfolding AE_ball_countable by simp lemma AE_pairwise: "countable F \ pairwise (\A B. AE x in M. R x A B) F \ (AE x in M. pairwise (R x) F)" unfolding pairwise_alt by (simp add: AE_ball_countable) lemma AE_discrete_difference: assumes X: "countable X" assumes null: "\x. x \ X \ emeasure M {x} = 0" assumes sets: "\x. x \ X \ {x} \ sets M" shows "AE x in M. x \ X" proof - have "(\x\X. {x}) \ null_sets M" using assms by (intro null_sets_UN') auto from AE_not_in[OF this] show "AE x in M. x \ X" by auto qed lemma AE_finite_all: assumes f: "finite S" shows "(AE x in M. \i\S. P i x) \ (\i\S. AE x in M. P i x)" using f by induct auto lemma AE_finite_allI: assumes "finite S" shows "(\s. s \ S \ AE x in M. Q s x) \ AE x in M. \s\S. Q s x" using AE_finite_all[OF \finite S\] by auto lemma emeasure_mono_AE: assumes imp: "AE x in M. x \ A \ x \ B" and B: "B \ sets M" shows "emeasure M A \ emeasure M B" proof cases assume A: "A \ sets M" from imp obtain N where N: "{x\space M. \ (x \ A \ x \ B)} \ N" "N \ null_sets M" by (auto simp: eventually_ae_filter) have "emeasure M A = emeasure M (A - N)" using N A by (subst emeasure_Diff_null_set) auto also have "emeasure M (A - N) \ emeasure M (B - N)" using N A B sets.sets_into_space by (auto intro!: emeasure_mono) also have "emeasure M (B - N) = emeasure M B" using N B by (subst emeasure_Diff_null_set) auto finally show ?thesis . qed (simp add: emeasure_notin_sets) lemma emeasure_eq_AE: assumes iff: "AE x in M. x \ A \ x \ B" assumes A: "A \ sets M" and B: "B \ sets M" shows "emeasure M A = emeasure M B" using assms by (safe intro!: antisym emeasure_mono_AE) auto lemma emeasure_Collect_eq_AE: "AE x in M. P x \ Q x \ Measurable.pred M Q \ Measurable.pred M P \ emeasure M {x\space M. P x} = emeasure M {x\space M. Q x}" by (intro emeasure_eq_AE) auto lemma emeasure_eq_0_AE: "AE x in M. \ P x \ emeasure M {x\space M. P x} = 0" using AE_iff_measurable[OF _ refl, of M "\x. \ P x"] by (cases "{x\space M. P x} \ sets M") (simp_all add: emeasure_notin_sets) lemma emeasure_0_AE: assumes "emeasure M (space M) = 0" shows "AE x in M. P x" using eventually_ae_filter assms by blast lemma emeasure_add_AE: assumes [measurable]: "A \ sets M" "B \ sets M" "C \ sets M" assumes 1: "AE x in M. x \ C \ x \ A \ x \ B" assumes 2: "AE x in M. \ (x \ A \ x \ B)" shows "emeasure M C = emeasure M A + emeasure M B" proof - have "emeasure M C = emeasure M (A \ B)" by (rule emeasure_eq_AE) (insert 1, auto) also have "\ = emeasure M A + emeasure M (B - A)" by (subst plus_emeasure) auto also have "emeasure M (B - A) = emeasure M B" by (rule emeasure_eq_AE) (insert 2, auto) finally show ?thesis . qed subsection \\\\-finite Measures\ locale\<^marker>\tag important\ sigma_finite_measure = fixes M :: "'a measure" assumes sigma_finite_countable: "\A::'a set set. countable A \ A \ sets M \ (\A) = space M \ (\a\A. emeasure M a \ \)" lemma (in sigma_finite_measure) sigma_finite: obtains A :: "nat \ 'a set" where "range A \ sets M" "(\i. A i) = space M" "\i. emeasure M (A i) \ \" proof - obtain A :: "'a set set" where [simp]: "countable A" and A: "A \ sets M" "(\A) = space M" "\a. a \ A \ emeasure M a \ \" using sigma_finite_countable by metis show thesis proof cases assume "A = {}" with \(\A) = space M\ show thesis by (intro that[of "\_. {}"]) auto next assume "A \ {}" show thesis proof show "range (from_nat_into A) \ sets M" using \A \ {}\ A by auto have "(\i. from_nat_into A i) = \A" using range_from_nat_into[OF \A \ {}\ \countable A\] by auto with A show "(\i. from_nat_into A i) = space M" by auto qed (intro A from_nat_into \A \ {}\) qed qed lemma (in sigma_finite_measure) sigma_finite_disjoint: obtains A :: "nat \ 'a set" where "range A \ sets M" "(\i. A i) = space M" "\i. emeasure M (A i) \ \" "disjoint_family A" proof - obtain A :: "nat \ 'a set" where range: "range A \ sets M" and space: "(\i. A i) = space M" and measure: "\i. emeasure M (A i) \ \" using sigma_finite by blast show thesis proof (rule that[of "disjointed A"]) show "range (disjointed A) \ sets M" by (rule sets.range_disjointed_sets[OF range]) show "(\i. disjointed A i) = space M" and "disjoint_family (disjointed A)" using disjoint_family_disjointed UN_disjointed_eq[of A] space range by auto show "emeasure M (disjointed A i) \ \" for i proof - have "emeasure M (disjointed A i) \ emeasure M (A i)" using range disjointed_subset[of A i] by (auto intro!: emeasure_mono) then show ?thesis using measure[of i] by (auto simp: top_unique) qed qed qed lemma (in sigma_finite_measure) sigma_finite_incseq: obtains A :: "nat \ 'a set" where "range A \ sets M" "(\i. A i) = space M" "\i. emeasure M (A i) \ \" "incseq A" proof - obtain F :: "nat \ 'a set" where F: "range F \ sets M" "(\i. F i) = space M" "\i. emeasure M (F i) \ \" using sigma_finite by blast show thesis proof (rule that[of "\n. \i\n. F i"]) show "range (\n. \i\n. F i) \ sets M" using F by (force simp: incseq_def) show "(\n. \i\n. F i) = space M" proof - from F have "\x. x \ space M \ \i. x \ F i" by auto with F show ?thesis by fastforce qed show "emeasure M (\i\n. F i) \ \" for n proof - have "emeasure M (\i\n. F i) \ (\i\n. emeasure M (F i))" using F by (auto intro!: emeasure_subadditive_finite) also have "\ < \" using F by (auto simp: sum_Pinfty less_top) finally show ?thesis by simp qed show "incseq (\n. \i\n. F i)" by (force simp: incseq_def) qed qed lemma (in sigma_finite_measure) approx_PInf_emeasure_with_finite: fixes C::real assumes W_meas: "W \ sets M" and W_inf: "emeasure M W = \" obtains Z where "Z \ sets M" "Z \ W" "emeasure M Z < \" "emeasure M Z > C" proof - obtain A :: "nat \ 'a set" where A: "range A \ sets M" "(\i. A i) = space M" "\i. emeasure M (A i) \ \" "incseq A" using sigma_finite_incseq by blast define B where "B = (\i. W \ A i)" have B_meas: "\i. B i \ sets M" using W_meas \range A \ sets M\ B_def by blast have b: "\i. B i \ W" using B_def by blast { fix i have "emeasure M (B i) \ emeasure M (A i)" using A by (intro emeasure_mono) (auto simp: B_def) also have "emeasure M (A i) < \" using \\i. emeasure M (A i) \ \\ by (simp add: less_top) finally have "emeasure M (B i) < \" . } note c = this have "W = (\i. B i)" using B_def \(\i. A i) = space M\ W_meas by auto moreover have "incseq B" using B_def \incseq A\ by (simp add: incseq_def subset_eq) ultimately have "(\i. emeasure M (B i)) \ emeasure M W" using W_meas B_meas by (simp add: B_meas Lim_emeasure_incseq image_subset_iff) then have "(\i. emeasure M (B i)) \ \" using W_inf by simp from order_tendstoD(1)[OF this, of C] obtain i where d: "emeasure M (B i) > C" by (auto simp: eventually_sequentially) have "B i \ sets M" "B i \ W" "emeasure M (B i) < \" "emeasure M (B i) > C" using B_meas b c d by auto then show ?thesis using that by blast qed subsection \Measure space induced by distribution of \<^const>\measurable\-functions\ definition\<^marker>\tag important\ distr :: "'a measure \ 'b measure \ ('a \ 'b) \ 'b measure" where "distr M N f = measure_of (space N) (sets N) (\A. emeasure M (f -` A \ space M))" lemma shows sets_distr[simp, measurable_cong]: "sets (distr M N f) = sets N" and space_distr[simp]: "space (distr M N f) = space N" by (auto simp: distr_def) lemma shows measurable_distr_eq1[simp]: "measurable (distr Mf Nf f) Mf' = measurable Nf Mf'" and measurable_distr_eq2[simp]: "measurable Mg' (distr Mg Ng g) = measurable Mg' Ng" by (auto simp: measurable_def) lemma distr_cong: "M = K \ sets N = sets L \ (\x. x \ space M \ f x = g x) \ distr M N f = distr K L g" using sets_eq_imp_space_eq[of N L] by (simp add: distr_def Int_def cong: rev_conj_cong) lemma emeasure_distr: fixes f :: "'a \ 'b" assumes f: "f \ measurable M N" and A: "A \ sets N" shows "emeasure (distr M N f) A = emeasure M (f -` A \ space M)" (is "_ = ?\ A") unfolding distr_def proof (rule emeasure_measure_of_sigma) show "positive (sets N) ?\" by (auto simp: positive_def) show "countably_additive (sets N) ?\" proof (intro countably_additiveI) fix A :: "nat \ 'b set" assume "range A \ sets N" "disjoint_family A" then have A: "\i. A i \ sets N" "(\i. A i) \ sets N" by auto then have *: "range (\i. f -` (A i) \ space M) \ sets M" using f by (auto simp: measurable_def) moreover have "(\i. f -` A i \ space M) \ sets M" using * by blast moreover have **: "disjoint_family (\i. f -` A i \ space M)" using \disjoint_family A\ by (auto simp: disjoint_family_on_def) ultimately show "(\i. ?\ (A i)) = ?\ (\i. A i)" using suminf_emeasure[OF _ **] A f by (auto simp: comp_def vimage_UN) qed show "sigma_algebra (space N) (sets N)" .. qed fact lemma emeasure_Collect_distr: assumes X[measurable]: "X \ measurable M N" "Measurable.pred N P" shows "emeasure (distr M N X) {x\space N. P x} = emeasure M {x\space M. P (X x)}" by (subst emeasure_distr) (auto intro!: arg_cong2[where f=emeasure] X(1)[THEN measurable_space]) lemma emeasure_lfp2[consumes 1, case_names cont f measurable]: assumes "P M" assumes cont: "sup_continuous F" assumes f: "\M. P M \ f \ measurable M' M" assumes *: "\M A. P M \ (\N. P N \ Measurable.pred N A) \ Measurable.pred M (F A)" shows "emeasure M' {x\space M'. lfp F (f x)} = (SUP i. emeasure M' {x\space M'. (F ^^ i) (\x. False) (f x)})" proof (subst (1 2) emeasure_Collect_distr[symmetric, where X=f]) show "f \ measurable M' M" "f \ measurable M' M" using f[OF \P M\] by auto { fix i show "Measurable.pred M ((F ^^ i) (\x. False))" using \P M\ by (induction i arbitrary: M) (auto intro!: *) } show "Measurable.pred M (lfp F)" using \P M\ cont * by (rule measurable_lfp_coinduct[of P]) have "emeasure (distr M' M f) {x \ space (distr M' M f). lfp F x} = (SUP i. emeasure (distr M' M f) {x \ space (distr M' M f). (F ^^ i) (\x. False) x})" using \P M\ proof (coinduction arbitrary: M rule: emeasure_lfp') case (measurable A N) then have "\N. P N \ Measurable.pred (distr M' N f) A" by metis then have "\N. P N \ Measurable.pred N A" by simp with \P N\[THEN *] show ?case by auto qed fact then show "emeasure (distr M' M f) {x \ space M. lfp F x} = (SUP i. emeasure (distr M' M f) {x \ space M. (F ^^ i) (\x. False) x})" by simp qed lemma distr_id[simp]: "distr N N (\x. x) = N" by (rule measure_eqI) (auto simp: emeasure_distr) lemma distr_id2: "sets M = sets N \ distr N M (\x. x) = N" by (rule measure_eqI) (auto simp: emeasure_distr) lemma measure_distr: "f \ measurable M N \ S \ sets N \ measure (distr M N f) S = measure M (f -` S \ space M)" by (simp add: emeasure_distr measure_def) lemma distr_cong_AE: assumes 1: "M = K" "sets N = sets L" and 2: "(AE x in M. f x = g x)" and "f \ measurable M N" and "g \ measurable K L" shows "distr M N f = distr K L g" proof (rule measure_eqI) fix A assume "A \ sets (distr M N f)" with assms show "emeasure (distr M N f) A = emeasure (distr K L g) A" by (auto simp add: emeasure_distr intro!: emeasure_eq_AE measurable_sets) qed (insert 1, simp) lemma AE_distrD: assumes f: "f \ measurable M M'" and AE: "AE x in distr M M' f. P x" shows "AE x in M. P (f x)" proof - from AE[THEN AE_E] guess N . with f show ?thesis unfolding eventually_ae_filter by (intro bexI[of _ "f -` N \ space M"]) (auto simp: emeasure_distr measurable_def) qed lemma AE_distr_iff: assumes f[measurable]: "f \ measurable M N" and P[measurable]: "{x \ space N. P x} \ sets N" shows "(AE x in distr M N f. P x) \ (AE x in M. P (f x))" proof (subst (1 2) AE_iff_measurable[OF _ refl]) have "f -` {x\space N. \ P x} \ space M = {x \ space M. \ P (f x)}" using f[THEN measurable_space] by auto then show "(emeasure (distr M N f) {x \ space (distr M N f). \ P x} = 0) = (emeasure M {x \ space M. \ P (f x)} = 0)" by (simp add: emeasure_distr) qed auto lemma null_sets_distr_iff: "f \ measurable M N \ A \ null_sets (distr M N f) \ f -` A \ space M \ null_sets M \ A \ sets N" by (auto simp add: null_sets_def emeasure_distr) proposition distr_distr: "g \ measurable N L \ f \ measurable M N \ distr (distr M N f) L g = distr M L (g \ f)" by (auto simp add: emeasure_distr measurable_space intro!: arg_cong[where f="emeasure M"] measure_eqI) subsection\<^marker>\tag unimportant\ \Real measure values\ lemma ring_of_finite_sets: "ring_of_sets (space M) {A\sets M. emeasure M A \ top}" proof (rule ring_of_setsI) show "a \ {A \ sets M. emeasure M A \ top} \ b \ {A \ sets M. emeasure M A \ top} \ a \ b \ {A \ sets M. emeasure M A \ top}" for a b using emeasure_subadditive[of a M b] by (auto simp: top_unique) show "a \ {A \ sets M. emeasure M A \ top} \ b \ {A \ sets M. emeasure M A \ top} \ a - b \ {A \ sets M. emeasure M A \ top}" for a b using emeasure_mono[of "a - b" a M] by (auto simp: top_unique) qed (auto dest: sets.sets_into_space) lemma measure_nonneg[simp]: "0 \ measure M A" unfolding measure_def by auto lemma measure_nonneg' [simp]: "\ measure M A < 0" using measure_nonneg not_le by blast lemma zero_less_measure_iff: "0 < measure M A \ measure M A \ 0" using measure_nonneg[of M A] by (auto simp add: le_less) lemma measure_le_0_iff: "measure M X \ 0 \ measure M X = 0" using measure_nonneg[of M X] by linarith lemma measure_empty[simp]: "measure M {} = 0" unfolding measure_def by (simp add: zero_ennreal.rep_eq) lemma emeasure_eq_ennreal_measure: "emeasure M A \ top \ emeasure M A = ennreal (measure M A)" by (cases "emeasure M A" rule: ennreal_cases) (auto simp: measure_def) lemma measure_zero_top: "emeasure M A = top \ measure M A = 0" by (simp add: measure_def) lemma measure_eq_emeasure_eq_ennreal: "0 \ x \ emeasure M A = ennreal x \ measure M A = x" using emeasure_eq_ennreal_measure[of M A] by (cases "A \ M") (auto simp: measure_notin_sets emeasure_notin_sets) lemma enn2real_plus:"a < top \ b < top \ enn2real (a + b) = enn2real a + enn2real b" by (simp add: enn2real_def plus_ennreal.rep_eq real_of_ereal_add less_top del: real_of_ereal_enn2ereal) lemma enn2real_sum:"(\i. i \ I \ f i < top) \ enn2real (sum f I) = sum (enn2real \ f) I" by (induction I rule: infinite_finite_induct) (auto simp: enn2real_plus) lemma measure_eq_AE: assumes iff: "AE x in M. x \ A \ x \ B" assumes A: "A \ sets M" and B: "B \ sets M" shows "measure M A = measure M B" using assms emeasure_eq_AE[OF assms] by (simp add: measure_def) lemma measure_Union: "emeasure M A \ \ \ emeasure M B \ \ \ A \ sets M \ B \ sets M \ A \ B = {} \ measure M (A \ B) = measure M A + measure M B" by (simp add: measure_def plus_emeasure[symmetric] enn2real_plus less_top) lemma disjoint_family_on_insert: "i \ I \ disjoint_family_on A (insert i I) \ A i \ (\i\I. A i) = {} \ disjoint_family_on A I" by (fastforce simp: disjoint_family_on_def) lemma measure_finite_Union: "finite S \ A`S \ sets M \ disjoint_family_on A S \ (\i. i \ S \ emeasure M (A i) \ \) \ measure M (\i\S. A i) = (\i\S. measure M (A i))" by (induction S rule: finite_induct) (auto simp: disjoint_family_on_insert measure_Union sum_emeasure[symmetric] sets.countable_UN'[OF countable_finite]) lemma measure_Diff: assumes finite: "emeasure M A \ \" and measurable: "A \ sets M" "B \ sets M" "B \ A" shows "measure M (A - B) = measure M A - measure M B" proof - have "emeasure M (A - B) \ emeasure M A" "emeasure M B \ emeasure M A" using measurable by (auto intro!: emeasure_mono) hence "measure M ((A - B) \ B) = measure M (A - B) + measure M B" using measurable finite by (rule_tac measure_Union) (auto simp: top_unique) thus ?thesis using \B \ A\ by (auto simp: Un_absorb2) qed lemma measure_UNION: assumes measurable: "range A \ sets M" "disjoint_family A" assumes finite: "emeasure M (\i. A i) \ \" shows "(\i. measure M (A i)) sums (measure M (\i. A i))" proof - have "(\i. emeasure M (A i)) sums (emeasure M (\i. A i))" unfolding suminf_emeasure[OF measurable, symmetric] by (simp add: summable_sums) moreover { fix i have "emeasure M (A i) \ emeasure M (\i. A i)" using measurable by (auto intro!: emeasure_mono) then have "emeasure M (A i) = ennreal ((measure M (A i)))" using finite by (intro emeasure_eq_ennreal_measure) (auto simp: top_unique) } ultimately show ?thesis using finite by (subst (asm) (2) emeasure_eq_ennreal_measure) simp_all qed lemma measure_subadditive: assumes measurable: "A \ sets M" "B \ sets M" and fin: "emeasure M A \ \" "emeasure M B \ \" shows "measure M (A \ B) \ measure M A + measure M B" proof - have "emeasure M (A \ B) \ \" using emeasure_subadditive[OF measurable] fin by (auto simp: top_unique) then show "(measure M (A \ B)) \ (measure M A) + (measure M B)" using emeasure_subadditive[OF measurable] fin apply simp apply (subst (asm) (2 3 4) emeasure_eq_ennreal_measure) apply (auto simp flip: ennreal_plus) done qed lemma measure_subadditive_finite: assumes A: "finite I" "A`I \ sets M" and fin: "\i. i \ I \ emeasure M (A i) \ \" shows "measure M (\i\I. A i) \ (\i\I. measure M (A i))" proof - { have "emeasure M (\i\I. A i) \ (\i\I. emeasure M (A i))" using emeasure_subadditive_finite[OF A] . also have "\ < \" using fin by (simp add: less_top A) finally have "emeasure M (\i\I. A i) \ top" by simp } note * = this show ?thesis using emeasure_subadditive_finite[OF A] fin unfolding emeasure_eq_ennreal_measure[OF *] by (simp_all add: sum_nonneg emeasure_eq_ennreal_measure) qed lemma measure_subadditive_countably: assumes A: "range A \ sets M" and fin: "(\i. emeasure M (A i)) \ \" shows "measure M (\i. A i) \ (\i. measure M (A i))" proof - from fin have **: "\i. emeasure M (A i) \ top" using ennreal_suminf_lessD[of "\i. emeasure M (A i)"] by (simp add: less_top) { have "emeasure M (\i. A i) \ (\i. emeasure M (A i))" using emeasure_subadditive_countably[OF A] . also have "\ < \" using fin by (simp add: less_top) finally have "emeasure M (\i. A i) \ top" by simp } then have "ennreal (measure M (\i. A i)) = emeasure M (\i. A i)" by (rule emeasure_eq_ennreal_measure[symmetric]) also have "\ \ (\i. emeasure M (A i))" using emeasure_subadditive_countably[OF A] . also have "\ = ennreal (\i. measure M (A i))" using fin unfolding emeasure_eq_ennreal_measure[OF **] by (subst suminf_ennreal) (auto simp: **) finally show ?thesis apply (rule ennreal_le_iff[THEN iffD1, rotated]) apply (intro suminf_nonneg allI measure_nonneg summable_suminf_not_top) using fin apply (simp add: emeasure_eq_ennreal_measure[OF **]) done qed lemma measure_Un_null_set: "A \ sets M \ B \ null_sets M \ measure M (A \ B) = measure M A" by (simp add: measure_def emeasure_Un_null_set) lemma measure_Diff_null_set: "A \ sets M \ B \ null_sets M \ measure M (A - B) = measure M A" by (simp add: measure_def emeasure_Diff_null_set) lemma measure_eq_sum_singleton: "finite S \ (\x. x \ S \ {x} \ sets M) \ (\x. x \ S \ emeasure M {x} \ \) \ measure M S = (\x\S. measure M {x})" using emeasure_eq_sum_singleton[of S M] by (intro measure_eq_emeasure_eq_ennreal) (auto simp: sum_nonneg emeasure_eq_ennreal_measure) lemma Lim_measure_incseq: assumes A: "range A \ sets M" "incseq A" and fin: "emeasure M (\i. A i) \ \" shows "(\i. measure M (A i)) \ measure M (\i. A i)" proof (rule tendsto_ennrealD) have "ennreal (measure M (\i. A i)) = emeasure M (\i. A i)" using fin by (auto simp: emeasure_eq_ennreal_measure) moreover have "ennreal (measure M (A i)) = emeasure M (A i)" for i using assms emeasure_mono[of "A _" "\i. A i" M] by (intro emeasure_eq_ennreal_measure[symmetric]) (auto simp: less_top UN_upper intro: le_less_trans) ultimately show "(\x. ennreal (measure M (A x))) \ ennreal (measure M (\i. A i))" using A by (auto intro!: Lim_emeasure_incseq) qed auto lemma Lim_measure_decseq: assumes A: "range A \ sets M" "decseq A" and fin: "\i. emeasure M (A i) \ \" shows "(\n. measure M (A n)) \ measure M (\i. A i)" proof (rule tendsto_ennrealD) have "ennreal (measure M (\i. A i)) = emeasure M (\i. A i)" using fin[of 0] A emeasure_mono[of "\i. A i" "A 0" M] by (auto intro!: emeasure_eq_ennreal_measure[symmetric] simp: INT_lower less_top intro: le_less_trans) moreover have "ennreal (measure M (A i)) = emeasure M (A i)" for i using A fin[of i] by (intro emeasure_eq_ennreal_measure[symmetric]) auto ultimately show "(\x. ennreal (measure M (A x))) \ ennreal (measure M (\i. A i))" using fin A by (auto intro!: Lim_emeasure_decseq) qed auto subsection \Set of measurable sets with finite measure\ definition\<^marker>\tag important\ fmeasurable :: "'a measure \ 'a set set" where "fmeasurable M = {A\sets M. emeasure M A < \}" lemma fmeasurableD[dest, measurable_dest]: "A \ fmeasurable M \ A \ sets M" by (auto simp: fmeasurable_def) lemma fmeasurableD2: "A \ fmeasurable M \ emeasure M A \ top" by (auto simp: fmeasurable_def) lemma fmeasurableI: "A \ sets M \ emeasure M A < \ \ A \ fmeasurable M" by (auto simp: fmeasurable_def) lemma fmeasurableI_null_sets: "A \ null_sets M \ A \ fmeasurable M" by (auto simp: fmeasurable_def) lemma fmeasurableI2: "A \ fmeasurable M \ B \ A \ B \ sets M \ B \ fmeasurable M" using emeasure_mono[of B A M] by (auto simp: fmeasurable_def) lemma measure_mono_fmeasurable: "A \ B \ A \ sets M \ B \ fmeasurable M \ measure M A \ measure M B" by (auto simp: measure_def fmeasurable_def intro!: emeasure_mono enn2real_mono) lemma emeasure_eq_measure2: "A \ fmeasurable M \ emeasure M A = measure M A" by (simp add: emeasure_eq_ennreal_measure fmeasurable_def less_top) interpretation fmeasurable: ring_of_sets "space M" "fmeasurable M" proof (rule ring_of_setsI) show "fmeasurable M \ Pow (space M)" "{} \ fmeasurable M" by (auto simp: fmeasurable_def dest: sets.sets_into_space) fix a b assume *: "a \ fmeasurable M" "b \ fmeasurable M" then have "emeasure M (a \ b) \ emeasure M a + emeasure M b" by (intro emeasure_subadditive) auto also have "\ < top" using * by (auto simp: fmeasurable_def) finally show "a \ b \ fmeasurable M" using * by (auto intro: fmeasurableI) show "a - b \ fmeasurable M" using emeasure_mono[of "a - b" a M] * by (auto simp: fmeasurable_def) qed subsection\<^marker>\tag unimportant\\Measurable sets formed by unions and intersections\ lemma fmeasurable_Diff: "A \ fmeasurable M \ B \ sets M \ A - B \ fmeasurable M" using fmeasurableI2[of A M "A - B"] by auto lemma fmeasurable_Int_fmeasurable: "\S \ fmeasurable M; T \ sets M\ \ (S \ T) \ fmeasurable M" by (meson fmeasurableD fmeasurableI2 inf_le1 sets.Int) lemma fmeasurable_UN: assumes "countable I" "\i. i \ I \ F i \ A" "\i. i \ I \ F i \ sets M" "A \ fmeasurable M" shows "(\i\I. F i) \ fmeasurable M" proof (rule fmeasurableI2) show "A \ fmeasurable M" "(\i\I. F i) \ A" using assms by auto show "(\i\I. F i) \ sets M" using assms by (intro sets.countable_UN') auto qed lemma fmeasurable_INT: assumes "countable I" "i \ I" "\i. i \ I \ F i \ sets M" "F i \ fmeasurable M" shows "(\i\I. F i) \ fmeasurable M" proof (rule fmeasurableI2) show "F i \ fmeasurable M" "(\i\I. F i) \ F i" using assms by auto show "(\i\I. F i) \ sets M" using assms by (intro sets.countable_INT') auto qed lemma measurable_measure_Diff: assumes "A \ fmeasurable M" "B \ sets M" "B \ A" shows "measure M (A - B) = measure M A - measure M B" by (simp add: assms fmeasurableD fmeasurableD2 measure_Diff) lemma measurable_Un_null_set: assumes "B \ null_sets M" shows "(A \ B \ fmeasurable M \ A \ sets M) \ A \ fmeasurable M" using assms by (fastforce simp add: fmeasurable.Un fmeasurableI_null_sets intro: fmeasurableI2) lemma measurable_Diff_null_set: assumes "B \ null_sets M" shows "(A - B) \ fmeasurable M \ A \ sets M \ A \ fmeasurable M" using assms by (metis Un_Diff_cancel2 fmeasurable.Diff fmeasurableD fmeasurableI_null_sets measurable_Un_null_set) lemma fmeasurable_Diff_D: assumes m: "T - S \ fmeasurable M" "S \ fmeasurable M" and sub: "S \ T" shows "T \ fmeasurable M" proof - have "T = S \ (T - S)" using assms by blast then show ?thesis by (metis m fmeasurable.Un) qed lemma measure_Un2: "A \ fmeasurable M \ B \ fmeasurable M \ measure M (A \ B) = measure M A + measure M (B - A)" using measure_Union[of M A "B - A"] by (auto simp: fmeasurableD2 fmeasurable.Diff) lemma measure_Un3: assumes "A \ fmeasurable M" "B \ fmeasurable M" shows "measure M (A \ B) = measure M A + measure M B - measure M (A \ B)" proof - have "measure M (A \ B) = measure M A + measure M (B - A)" using assms by (rule measure_Un2) also have "B - A = B - (A \ B)" by auto also have "measure M (B - (A \ B)) = measure M B - measure M (A \ B)" using assms by (intro measure_Diff) (auto simp: fmeasurable_def) finally show ?thesis by simp qed lemma measure_Un_AE: "AE x in M. x \ A \ x \ B \ A \ fmeasurable M \ B \ fmeasurable M \ measure M (A \ B) = measure M A + measure M B" by (subst measure_Un2) (auto intro!: measure_eq_AE) lemma measure_UNION_AE: assumes I: "finite I" shows "(\i. i \ I \ F i \ fmeasurable M) \ pairwise (\i j. AE x in M. x \ F i \ x \ F j) I \ measure M (\i\I. F i) = (\i\I. measure M (F i))" unfolding AE_pairwise[OF countable_finite, OF I] using I proof (induction I rule: finite_induct) case (insert x I) have "measure M (F x \ \(F ` I)) = measure M (F x) + measure M (\(F ` I))" by (rule measure_Un_AE) (use insert in \auto simp: pairwise_insert\) with insert show ?case by (simp add: pairwise_insert ) qed simp lemma measure_UNION': "finite I \ (\i. i \ I \ F i \ fmeasurable M) \ pairwise (\i j. disjnt (F i) (F j)) I \ measure M (\i\I. F i) = (\i\I. measure M (F i))" by (intro measure_UNION_AE) (auto simp: disjnt_def elim!: pairwise_mono intro!: always_eventually) lemma measure_Union_AE: "finite F \ (\S. S \ F \ S \ fmeasurable M) \ pairwise (\S T. AE x in M. x \ S \ x \ T) F \ measure M (\F) = (\S\F. measure M S)" using measure_UNION_AE[of F "\x. x" M] by simp lemma measure_Union': "finite F \ (\S. S \ F \ S \ fmeasurable M) \ pairwise disjnt F \ measure M (\F) = (\S\F. measure M S)" using measure_UNION'[of F "\x. x" M] by simp lemma measure_Un_le: assumes "A \ sets M" "B \ sets M" shows "measure M (A \ B) \ measure M A + measure M B" proof cases assume "A \ fmeasurable M \ B \ fmeasurable M" with measure_subadditive[of A M B] assms show ?thesis by (auto simp: fmeasurableD2) next assume "\ (A \ fmeasurable M \ B \ fmeasurable M)" then have "A \ B \ fmeasurable M" using fmeasurableI2[of "A \ B" M A] fmeasurableI2[of "A \ B" M B] assms by auto with assms show ?thesis by (auto simp: fmeasurable_def measure_def less_top[symmetric]) qed lemma measure_UNION_le: "finite I \ (\i. i \ I \ F i \ sets M) \ measure M (\i\I. F i) \ (\i\I. measure M (F i))" proof (induction I rule: finite_induct) case (insert i I) - then have "measure M (\i\insert i I. F i) \ measure M (F i) + measure M (\i\I. F i)" - by (auto intro!: measure_Un_le) + then have "measure M (\i\insert i I. F i) = measure M (F i \ \ (F ` I))" + by simp + also from insert have "measure M (F i \ \ (F ` I)) \ measure M (F i) + measure M (\ (F ` I))" + by (intro measure_Un_le sets.finite_Union) auto also have "measure M (\i\I. F i) \ (\i\I. measure M (F i))" using insert by auto finally show ?case using insert by simp qed simp lemma measure_Union_le: "finite F \ (\S. S \ F \ S \ sets M) \ measure M (\F) \ (\S\F. measure M S)" using measure_UNION_le[of F "\x. x" M] by simp text\Version for indexed union over a countable set\ lemma assumes "countable I" and I: "\i. i \ I \ A i \ fmeasurable M" and bound: "\I'. I' \ I \ finite I' \ measure M (\i\I'. A i) \ B" shows fmeasurable_UN_bound: "(\i\I. A i) \ fmeasurable M" (is ?fm) and measure_UN_bound: "measure M (\i\I. A i) \ B" (is ?m) proof - have "B \ 0" using bound by force have "?fm \ ?m" proof cases assume "I = {}" with \B \ 0\ show ?thesis by simp next assume "I \ {}" have "(\i\I. A i) = (\i. (\n\i. A (from_nat_into I n)))" by (subst range_from_nat_into[symmetric, OF \I \ {}\ \countable I\]) auto then have "emeasure M (\i\I. A i) = emeasure M (\i. (\n\i. A (from_nat_into I n)))" by simp also have "\ = (SUP i. emeasure M (\n\i. A (from_nat_into I n)))" using I \I \ {}\[THEN from_nat_into] by (intro SUP_emeasure_incseq[symmetric]) (fastforce simp: incseq_Suc_iff)+ also have "\ \ B" proof (intro SUP_least) fix i :: nat have "emeasure M (\n\i. A (from_nat_into I n)) = measure M (\n\i. A (from_nat_into I n))" using I \I \ {}\[THEN from_nat_into] by (intro emeasure_eq_measure2 fmeasurable.finite_UN) auto also have "\ = measure M (\n\from_nat_into I ` {..i}. A n)" by simp also have "\ \ B" by (intro ennreal_leI bound) (auto intro: from_nat_into[OF \I \ {}\]) finally show "emeasure M (\n\i. A (from_nat_into I n)) \ ennreal B" . qed finally have *: "emeasure M (\i\I. A i) \ B" . then have ?fm using I \countable I\ by (intro fmeasurableI conjI) (auto simp: less_top[symmetric] top_unique) with * \0\B\ show ?thesis by (simp add: emeasure_eq_measure2) qed then show ?fm ?m by auto qed text\Version for big union of a countable set\ lemma assumes "countable \" and meas: "\D. D \ \ \ D \ fmeasurable M" and bound: "\\. \\ \ \; finite \\ \ measure M (\\) \ B" shows fmeasurable_Union_bound: "\\ \ fmeasurable M" (is ?fm) and measure_Union_bound: "measure M (\\) \ B" (is ?m) proof - have "B \ 0" using bound by force have "?fm \ ?m" proof (cases "\ = {}") case True with \B \ 0\ show ?thesis by auto next case False then obtain D :: "nat \ 'a set" where D: "\ = range D" using \countable \\ uncountable_def by force have 1: "\i. D i \ fmeasurable M" by (simp add: D meas) have 2: "\I'. finite I' \ measure M (\x\I'. D x) \ B" by (simp add: D bound image_subset_iff) show ?thesis unfolding D by (intro conjI fmeasurable_UN_bound [OF _ 1 2] measure_UN_bound [OF _ 1 2]) auto qed then show ?fm ?m by auto qed text\Version for indexed union over the type of naturals\ lemma fixes S :: "nat \ 'a set" assumes S: "\i. S i \ fmeasurable M" and B: "\n. measure M (\i\n. S i) \ B" shows fmeasurable_countable_Union: "(\i. S i) \ fmeasurable M" and measure_countable_Union_le: "measure M (\i. S i) \ B" proof - have mB: "measure M (\i\I. S i) \ B" if "finite I" for I proof - have "(\i\I. S i) \ (\i\Max I. S i)" using Max_ge that by force then have "measure M (\i\I. S i) \ measure M (\i \ Max I. S i)" by (rule measure_mono_fmeasurable) (use S in \blast+\) then show ?thesis using B order_trans by blast qed show "(\i. S i) \ fmeasurable M" by (auto intro: fmeasurable_UN_bound [OF _ S mB]) show "measure M (\n. S n) \ B" by (auto intro: measure_UN_bound [OF _ S mB]) qed lemma measure_diff_le_measure_setdiff: assumes "S \ fmeasurable M" "T \ fmeasurable M" shows "measure M S - measure M T \ measure M (S - T)" proof - have "measure M S \ measure M ((S - T) \ T)" by (simp add: assms fmeasurable.Un fmeasurableD measure_mono_fmeasurable) also have "\ \ measure M (S - T) + measure M T" using assms by (blast intro: measure_Un_le) finally show ?thesis by (simp add: algebra_simps) qed lemma suminf_exist_split2: fixes f :: "nat \ 'a::real_normed_vector" assumes "summable f" shows "(\n. (\k. f(k+n))) \ 0" by (subst lim_sequentially, auto simp add: dist_norm suminf_exist_split[OF _ assms]) lemma emeasure_union_summable: assumes [measurable]: "\n. A n \ sets M" and "\n. emeasure M (A n) < \" "summable (\n. measure M (A n))" shows "emeasure M (\n. A n) < \" "emeasure M (\n. A n) \ (\n. measure M (A n))" proof - define B where "B = (\N. (\n\{.. sets M" for N unfolding B_def by auto have "(\N. emeasure M (B N)) \ emeasure M (\N. B N)" apply (rule Lim_emeasure_incseq) unfolding B_def by (auto simp add: SUP_subset_mono incseq_def) moreover have "emeasure M (B N) \ ennreal (\n. measure M (A n))" for N proof - have *: "(\n\{.. (\n. measure M (A n))" using assms(3) measure_nonneg sum_le_suminf by blast have "emeasure M (B N) \ (\n\{..n\{..n\{.. ennreal (\n. measure M (A n))" using * by (auto simp: ennreal_leI) finally show ?thesis by simp qed ultimately have "emeasure M (\N. B N) \ ennreal (\n. measure M (A n))" by (simp add: Lim_bounded) then show "emeasure M (\n. A n) \ (\n. measure M (A n))" unfolding B_def by (metis UN_UN_flatten UN_lessThan_UNIV) then show "emeasure M (\n. A n) < \" by (auto simp: less_top[symmetric] top_unique) qed lemma borel_cantelli_limsup1: assumes [measurable]: "\n. A n \ sets M" and "\n. emeasure M (A n) < \" "summable (\n. measure M (A n))" shows "limsup A \ null_sets M" proof - have "emeasure M (limsup A) \ 0" proof (rule LIMSEQ_le_const) have "(\n. (\k. measure M (A (k+n)))) \ 0" by (rule suminf_exist_split2[OF assms(3)]) then show "(\n. ennreal (\k. measure M (A (k+n)))) \ 0" unfolding ennreal_0[symmetric] by (intro tendsto_ennrealI) have "emeasure M (limsup A) \ (\k. measure M (A (k+n)))" for n proof - have I: "(\k\{n..}. A k) = (\k. A (k+n))" by (auto, metis le_add_diff_inverse2, fastforce) have "emeasure M (limsup A) \ emeasure M (\k\{n..}. A k)" by (rule emeasure_mono, auto simp add: limsup_INF_SUP) also have "... = emeasure M (\k. A (k+n))" using I by auto also have "... \ (\k. measure M (A (k+n)))" apply (rule emeasure_union_summable) using assms summable_ignore_initial_segment[OF assms(3), of n] by auto finally show ?thesis by simp qed then show "\N. \n\N. emeasure M (limsup A) \ (\k. measure M (A (k+n)))" by auto qed then show ?thesis using assms(1) measurable_limsup by auto qed lemma borel_cantelli_AE1: assumes [measurable]: "\n. A n \ sets M" and "\n. emeasure M (A n) < \" "summable (\n. measure M (A n))" shows "AE x in M. eventually (\n. x \ space M - A n) sequentially" proof - have "AE x in M. x \ limsup A" using borel_cantelli_limsup1[OF assms] unfolding eventually_ae_filter by auto moreover { fix x assume "x \ limsup A" then obtain N where "x \ (\n\{N..}. A n)" unfolding limsup_INF_SUP by blast then have "eventually (\n. x \ A n) sequentially" using eventually_sequentially by auto } ultimately show ?thesis by auto qed subsection \Measure spaces with \<^term>\emeasure M (space M) < \\\ locale\<^marker>\tag important\ finite_measure = sigma_finite_measure M for M + assumes finite_emeasure_space: "emeasure M (space M) \ top" lemma finite_measureI[Pure.intro!]: "emeasure M (space M) \ \ \ finite_measure M" proof qed (auto intro!: exI[of _ "{space M}"]) lemma (in finite_measure) emeasure_finite[simp, intro]: "emeasure M A \ top" using finite_emeasure_space emeasure_space[of M A] by (auto simp: top_unique) lemma (in finite_measure) fmeasurable_eq_sets: "fmeasurable M = sets M" by (auto simp: fmeasurable_def less_top[symmetric]) lemma (in finite_measure) emeasure_eq_measure: "emeasure M A = ennreal (measure M A)" by (intro emeasure_eq_ennreal_measure) simp lemma (in finite_measure) emeasure_real: "\r. 0 \ r \ emeasure M A = ennreal r" using emeasure_finite[of A] by (cases "emeasure M A" rule: ennreal_cases) auto lemma (in finite_measure) bounded_measure: "measure M A \ measure M (space M)" using emeasure_space[of M A] emeasure_real[of A] emeasure_real[of "space M"] by (auto simp: measure_def) lemma (in finite_measure) finite_measure_Diff: assumes sets: "A \ sets M" "B \ sets M" and "B \ A" shows "measure M (A - B) = measure M A - measure M B" using measure_Diff[OF _ assms] by simp lemma (in finite_measure) finite_measure_Union: assumes sets: "A \ sets M" "B \ sets M" and "A \ B = {}" shows "measure M (A \ B) = measure M A + measure M B" using measure_Union[OF _ _ assms] by simp lemma (in finite_measure) finite_measure_finite_Union: assumes measurable: "finite S" "A`S \ sets M" "disjoint_family_on A S" shows "measure M (\i\S. A i) = (\i\S. measure M (A i))" using measure_finite_Union[OF assms] by simp lemma (in finite_measure) finite_measure_UNION: assumes A: "range A \ sets M" "disjoint_family A" shows "(\i. measure M (A i)) sums (measure M (\i. A i))" using measure_UNION[OF A] by simp lemma (in finite_measure) finite_measure_mono: assumes "A \ B" "B \ sets M" shows "measure M A \ measure M B" using emeasure_mono[OF assms] emeasure_real[of A] emeasure_real[of B] by (auto simp: measure_def) lemma (in finite_measure) finite_measure_subadditive: assumes m: "A \ sets M" "B \ sets M" shows "measure M (A \ B) \ measure M A + measure M B" using measure_subadditive[OF m] by simp lemma (in finite_measure) finite_measure_subadditive_finite: assumes "finite I" "A`I \ sets M" shows "measure M (\i\I. A i) \ (\i\I. measure M (A i))" using measure_subadditive_finite[OF assms] by simp lemma (in finite_measure) finite_measure_subadditive_countably: "range A \ sets M \ summable (\i. measure M (A i)) \ measure M (\i. A i) \ (\i. measure M (A i))" by (rule measure_subadditive_countably) (simp_all add: ennreal_suminf_neq_top emeasure_eq_measure) lemma (in finite_measure) finite_measure_eq_sum_singleton: assumes "finite S" and *: "\x. x \ S \ {x} \ sets M" shows "measure M S = (\x\S. measure M {x})" using measure_eq_sum_singleton[OF assms] by simp lemma (in finite_measure) finite_Lim_measure_incseq: assumes A: "range A \ sets M" "incseq A" shows "(\i. measure M (A i)) \ measure M (\i. A i)" using Lim_measure_incseq[OF A] by simp lemma (in finite_measure) finite_Lim_measure_decseq: assumes A: "range A \ sets M" "decseq A" shows "(\n. measure M (A n)) \ measure M (\i. A i)" using Lim_measure_decseq[OF A] by simp lemma (in finite_measure) finite_measure_compl: assumes S: "S \ sets M" shows "measure M (space M - S) = measure M (space M) - measure M S" using measure_Diff[OF _ sets.top S sets.sets_into_space] S by simp lemma (in finite_measure) finite_measure_mono_AE: assumes imp: "AE x in M. x \ A \ x \ B" and B: "B \ sets M" shows "measure M A \ measure M B" using assms emeasure_mono_AE[OF imp B] by (simp add: emeasure_eq_measure) lemma (in finite_measure) finite_measure_eq_AE: assumes iff: "AE x in M. x \ A \ x \ B" assumes A: "A \ sets M" and B: "B \ sets M" shows "measure M A = measure M B" using assms emeasure_eq_AE[OF assms] by (simp add: emeasure_eq_measure) lemma (in finite_measure) measure_increasing: "increasing M (measure M)" by (auto intro!: finite_measure_mono simp: increasing_def) lemma (in finite_measure) measure_zero_union: assumes "s \ sets M" "t \ sets M" "measure M t = 0" shows "measure M (s \ t) = measure M s" using assms proof - have "measure M (s \ t) \ measure M s" using finite_measure_subadditive[of s t] assms by auto moreover have "measure M (s \ t) \ measure M s" using assms by (blast intro: finite_measure_mono) ultimately show ?thesis by simp qed lemma (in finite_measure) measure_eq_compl: assumes "s \ sets M" "t \ sets M" assumes "measure M (space M - s) = measure M (space M - t)" shows "measure M s = measure M t" using assms finite_measure_compl by auto lemma (in finite_measure) measure_eq_bigunion_image: assumes "range f \ sets M" "range g \ sets M" assumes "disjoint_family f" "disjoint_family g" assumes "\ n :: nat. measure M (f n) = measure M (g n)" shows "measure M (\i. f i) = measure M (\i. g i)" using assms proof - have a: "(\ i. measure M (f i)) sums (measure M (\i. f i))" by (rule finite_measure_UNION[OF assms(1,3)]) have b: "(\ i. measure M (g i)) sums (measure M (\i. g i))" by (rule finite_measure_UNION[OF assms(2,4)]) show ?thesis using sums_unique[OF b] sums_unique[OF a] assms by simp qed lemma (in finite_measure) measure_countably_zero: assumes "range c \ sets M" assumes "\ i. measure M (c i) = 0" shows "measure M (\i :: nat. c i) = 0" proof (rule antisym) show "measure M (\i :: nat. c i) \ 0" using finite_measure_subadditive_countably[OF assms(1)] by (simp add: assms(2)) qed simp lemma (in finite_measure) measure_space_inter: assumes events:"s \ sets M" "t \ sets M" assumes "measure M t = measure M (space M)" shows "measure M (s \ t) = measure M s" proof - have "measure M ((space M - s) \ (space M - t)) = measure M (space M - s)" using events assms finite_measure_compl[of "t"] by (auto intro!: measure_zero_union) also have "(space M - s) \ (space M - t) = space M - (s \ t)" by blast finally show "measure M (s \ t) = measure M s" using events by (auto intro!: measure_eq_compl[of "s \ t" s]) qed lemma (in finite_measure) measure_equiprobable_finite_unions: assumes s: "finite s" "\x. x \ s \ {x} \ sets M" assumes "\ x y. \x \ s; y \ s\ \ measure M {x} = measure M {y}" shows "measure M s = real (card s) * measure M {SOME x. x \ s}" proof cases assume "s \ {}" then have "\ x. x \ s" by blast from someI_ex[OF this] assms have prob_some: "\ x. x \ s \ measure M {x} = measure M {SOME y. y \ s}" by blast have "measure M s = (\ x \ s. measure M {x})" using finite_measure_eq_sum_singleton[OF s] by simp also have "\ = (\ x \ s. measure M {SOME y. y \ s})" using prob_some by auto also have "\ = real (card s) * measure M {(SOME x. x \ s)}" using sum_constant assms by simp finally show ?thesis by simp qed simp lemma (in finite_measure) measure_real_sum_image_fn: assumes "e \ sets M" assumes "\ x. x \ s \ e \ f x \ sets M" assumes "finite s" assumes disjoint: "\ x y. \x \ s ; y \ s ; x \ y\ \ f x \ f y = {}" assumes upper: "space M \ (\i \ s. f i)" shows "measure M e = (\ x \ s. measure M (e \ f x))" proof - have "e \ (\i\s. f i)" using \e \ sets M\ sets.sets_into_space upper by blast then have e: "e = (\i \ s. e \ f i)" by auto hence "measure M e = measure M (\i \ s. e \ f i)" by simp also have "\ = (\ x \ s. measure M (e \ f x))" proof (rule finite_measure_finite_Union) show "finite s" by fact show "(\i. e \ f i)`s \ sets M" using assms(2) by auto show "disjoint_family_on (\i. e \ f i) s" using disjoint by (auto simp: disjoint_family_on_def) qed finally show ?thesis . qed lemma (in finite_measure) measure_exclude: assumes "A \ sets M" "B \ sets M" assumes "measure M A = measure M (space M)" "A \ B = {}" shows "measure M B = 0" using measure_space_inter[of B A] assms by (auto simp: ac_simps) lemma (in finite_measure) finite_measure_distr: assumes f: "f \ measurable M M'" shows "finite_measure (distr M M' f)" proof (rule finite_measureI) have "f -` space M' \ space M = space M" using f by (auto dest: measurable_space) with f show "emeasure (distr M M' f) (space (distr M M' f)) \ \" by (auto simp: emeasure_distr) qed lemma emeasure_gfp[consumes 1, case_names cont measurable]: assumes sets[simp]: "\s. sets (M s) = sets N" assumes "\s. finite_measure (M s)" assumes cont: "inf_continuous F" "inf_continuous f" assumes meas: "\P. Measurable.pred N P \ Measurable.pred N (F P)" assumes iter: "\P s. Measurable.pred N P \ emeasure (M s) {x\space N. F P x} = f (\s. emeasure (M s) {x\space N. P x}) s" assumes bound: "\P. f P \ f (\s. emeasure (M s) (space (M s)))" shows "emeasure (M s) {x\space N. gfp F x} = gfp f s" proof (subst gfp_transfer_bounded[where \="\F s. emeasure (M s) {x\space N. F x}" and g=f and f=F and P="Measurable.pred N", symmetric]) interpret finite_measure "M s" for s by fact fix C assume "decseq C" "\i. Measurable.pred N (C i)" then show "(\s. emeasure (M s) {x \ space N. (INF i. C i) x}) = (INF i. (\s. emeasure (M s) {x \ space N. C i x}))" unfolding INF_apply[abs_def] by (subst INF_emeasure_decseq) (auto simp: antimono_def fun_eq_iff intro!: arg_cong2[where f=emeasure]) next show "f x \ (\s. emeasure (M s) {x \ space N. F top x})" for x using bound[of x] sets_eq_imp_space_eq[OF sets] by (simp add: iter) qed (auto simp add: iter le_fun_def INF_apply[abs_def] intro!: meas cont) subsection\<^marker>\tag unimportant\ \Counting space\ lemma strict_monoI_Suc: assumes ord [simp]: "(\n. f n < f (Suc n))" shows "strict_mono f" unfolding strict_mono_def proof safe fix n m :: nat assume "n < m" then show "f n < f m" by (induct m) (auto simp: less_Suc_eq intro: less_trans ord) qed lemma emeasure_count_space: assumes "X \ A" shows "emeasure (count_space A) X = (if finite X then of_nat (card X) else \)" (is "_ = ?M X") unfolding count_space_def proof (rule emeasure_measure_of_sigma) show "X \ Pow A" using \X \ A\ by auto show "sigma_algebra A (Pow A)" by (rule sigma_algebra_Pow) show positive: "positive (Pow A) ?M" by (auto simp: positive_def) have additive: "additive (Pow A) ?M" by (auto simp: card_Un_disjoint additive_def) interpret ring_of_sets A "Pow A" by (rule ring_of_setsI) auto show "countably_additive (Pow A) ?M" unfolding countably_additive_iff_continuous_from_below[OF positive additive] proof safe fix F :: "nat \ 'a set" assume "incseq F" show "(\i. ?M (F i)) \ ?M (\i. F i)" proof cases assume "\i. \j\i. F i = F j" then guess i .. note i = this { fix j from i \incseq F\ have "F j \ F i" by (cases "i \ j") (auto simp: incseq_def) } then have eq: "(\i. F i) = F i" by auto with i show ?thesis by (auto intro!: Lim_transform_eventually[OF tendsto_const] eventually_sequentiallyI[where c=i]) next assume "\ (\i. \j\i. F i = F j)" then obtain f where f: "\i. i \ f i" "\i. F i \ F (f i)" by metis then have "\i. F i \ F (f i)" using \incseq F\ by (auto simp: incseq_def) with f have *: "\i. F i \ F (f i)" by auto have "incseq (\i. ?M (F i))" using \incseq F\ unfolding incseq_def by (auto simp: card_mono dest: finite_subset) then have "(\i. ?M (F i)) \ (SUP n. ?M (F n))" by (rule LIMSEQ_SUP) moreover have "(SUP n. ?M (F n)) = top" proof (rule ennreal_SUP_eq_top) fix n :: nat show "\k::nat\UNIV. of_nat n \ ?M (F k)" proof (induct n) case (Suc n) then guess k .. note k = this moreover have "finite (F k) \ finite (F (f k)) \ card (F k) < card (F (f k))" using \F k \ F (f k)\ by (simp add: psubset_card_mono) moreover have "finite (F (f k)) \ finite (F k)" using \k \ f k\ \incseq F\ by (auto simp: incseq_def dest: finite_subset) ultimately show ?case by (auto intro!: exI[of _ "f k"] simp del: of_nat_Suc) qed auto qed moreover have "inj (\n. F ((f ^^ n) 0))" by (intro strict_mono_imp_inj_on strict_monoI_Suc) (simp add: *) then have 1: "infinite (range (\i. F ((f ^^ i) 0)))" by (rule range_inj_infinite) have "infinite (Pow (\i. F i))" by (rule infinite_super[OF _ 1]) auto then have "infinite (\i. F i)" by auto ultimately show ?thesis by (simp only:) simp qed qed qed lemma distr_bij_count_space: assumes f: "bij_betw f A B" shows "distr (count_space A) (count_space B) f = count_space B" proof (rule measure_eqI) have f': "f \ measurable (count_space A) (count_space B)" using f unfolding Pi_def bij_betw_def by auto fix X assume "X \ sets (distr (count_space A) (count_space B) f)" then have X: "X \ sets (count_space B)" by auto moreover from X have "f -` X \ A = the_inv_into A f ` X" using f by (auto simp: bij_betw_def subset_image_iff image_iff the_inv_into_f_f intro: the_inv_into_f_f[symmetric]) moreover have "inj_on (the_inv_into A f) B" using X f by (auto simp: bij_betw_def inj_on_the_inv_into) with X have "inj_on (the_inv_into A f) X" by (auto intro: subset_inj_on) ultimately show "emeasure (distr (count_space A) (count_space B) f) X = emeasure (count_space B) X" using f unfolding emeasure_distr[OF f' X] by (subst (1 2) emeasure_count_space) (auto simp: card_image dest: finite_imageD) qed simp lemma emeasure_count_space_finite[simp]: "X \ A \ finite X \ emeasure (count_space A) X = of_nat (card X)" using emeasure_count_space[of X A] by simp lemma emeasure_count_space_infinite[simp]: "X \ A \ infinite X \ emeasure (count_space A) X = \" using emeasure_count_space[of X A] by simp lemma measure_count_space: "measure (count_space A) X = (if X \ A then of_nat (card X) else 0)" by (cases "finite X") (auto simp: measure_notin_sets ennreal_of_nat_eq_real_of_nat measure_zero_top measure_eq_emeasure_eq_ennreal) lemma emeasure_count_space_eq_0: "emeasure (count_space A) X = 0 \ (X \ A \ X = {})" proof cases assume X: "X \ A" then show ?thesis proof (intro iffI impI) assume "emeasure (count_space A) X = 0" with X show "X = {}" by (subst (asm) emeasure_count_space) (auto split: if_split_asm) qed simp qed (simp add: emeasure_notin_sets) lemma null_sets_count_space: "null_sets (count_space A) = { {} }" unfolding null_sets_def by (auto simp add: emeasure_count_space_eq_0) lemma AE_count_space: "(AE x in count_space A. P x) \ (\x\A. P x)" unfolding eventually_ae_filter by (auto simp add: null_sets_count_space) lemma sigma_finite_measure_count_space_countable: assumes A: "countable A" shows "sigma_finite_measure (count_space A)" proof qed (insert A, auto intro!: exI[of _ "(\a. {a}) ` A"]) lemma sigma_finite_measure_count_space: fixes A :: "'a::countable set" shows "sigma_finite_measure (count_space A)" by (rule sigma_finite_measure_count_space_countable) auto lemma finite_measure_count_space: assumes [simp]: "finite A" shows "finite_measure (count_space A)" by rule simp lemma sigma_finite_measure_count_space_finite: assumes A: "finite A" shows "sigma_finite_measure (count_space A)" proof - interpret finite_measure "count_space A" using A by (rule finite_measure_count_space) show "sigma_finite_measure (count_space A)" .. qed subsection\<^marker>\tag unimportant\ \Measure restricted to space\ lemma emeasure_restrict_space: assumes "\ \ space M \ sets M" "A \ \" shows "emeasure (restrict_space M \) A = emeasure M A" proof (cases "A \ sets M") case True show ?thesis proof (rule emeasure_measure_of[OF restrict_space_def]) show "(\) \ ` sets M \ Pow (\ \ space M)" "A \ sets (restrict_space M \)" using \A \ \\ \A \ sets M\ sets.space_closed by (auto simp: sets_restrict_space) show "positive (sets (restrict_space M \)) (emeasure M)" by (auto simp: positive_def) show "countably_additive (sets (restrict_space M \)) (emeasure M)" proof (rule countably_additiveI) fix A :: "nat \ _" assume "range A \ sets (restrict_space M \)" "disjoint_family A" with assms have "\i. A i \ sets M" "\i. A i \ space M" "disjoint_family A" by (fastforce simp: sets_restrict_space_iff[OF assms(1)] image_subset_iff dest: sets.sets_into_space)+ then show "(\i. emeasure M (A i)) = emeasure M (\i. A i)" by (subst suminf_emeasure) (auto simp: disjoint_family_subset) qed qed next case False with assms have "A \ sets (restrict_space M \)" by (simp add: sets_restrict_space_iff) with False show ?thesis by (simp add: emeasure_notin_sets) qed lemma measure_restrict_space: assumes "\ \ space M \ sets M" "A \ \" shows "measure (restrict_space M \) A = measure M A" using emeasure_restrict_space[OF assms] by (simp add: measure_def) lemma AE_restrict_space_iff: assumes "\ \ space M \ sets M" shows "(AE x in restrict_space M \. P x) \ (AE x in M. x \ \ \ P x)" proof - have ex_cong: "\P Q f. (\x. P x \ Q x) \ (\x. Q x \ P (f x)) \ (\x. P x) \ (\x. Q x)" by auto { fix X assume X: "X \ sets M" "emeasure M X = 0" then have "emeasure M (\ \ space M \ X) \ emeasure M X" by (intro emeasure_mono) auto then have "emeasure M (\ \ space M \ X) = 0" using X by (auto intro!: antisym) } with assms show ?thesis unfolding eventually_ae_filter by (auto simp add: space_restrict_space null_sets_def sets_restrict_space_iff emeasure_restrict_space cong: conj_cong intro!: ex_cong[where f="\X. (\ \ space M) \ X"]) qed lemma restrict_restrict_space: assumes "A \ space M \ sets M" "B \ space M \ sets M" shows "restrict_space (restrict_space M A) B = restrict_space M (A \ B)" (is "?l = ?r") proof (rule measure_eqI[symmetric]) show "sets ?r = sets ?l" unfolding sets_restrict_space image_comp by (intro image_cong) auto next fix X assume "X \ sets (restrict_space M (A \ B))" then obtain Y where "Y \ sets M" "X = Y \ A \ B" by (auto simp: sets_restrict_space) with assms sets.Int[OF assms] show "emeasure ?r X = emeasure ?l X" by (subst (1 2) emeasure_restrict_space) (auto simp: space_restrict_space sets_restrict_space_iff emeasure_restrict_space ac_simps) qed lemma restrict_count_space: "restrict_space (count_space B) A = count_space (A \ B)" proof (rule measure_eqI) show "sets (restrict_space (count_space B) A) = sets (count_space (A \ B))" by (subst sets_restrict_space) auto moreover fix X assume "X \ sets (restrict_space (count_space B) A)" ultimately have "X \ A \ B" by auto then show "emeasure (restrict_space (count_space B) A) X = emeasure (count_space (A \ B)) X" by (cases "finite X") (auto simp add: emeasure_restrict_space) qed lemma sigma_finite_measure_restrict_space: assumes "sigma_finite_measure M" and A: "A \ sets M" shows "sigma_finite_measure (restrict_space M A)" proof - interpret sigma_finite_measure M by fact from sigma_finite_countable obtain C where C: "countable C" "C \ sets M" "(\C) = space M" "\a\C. emeasure M a \ \" by blast let ?C = "(\) A ` C" from C have "countable ?C" "?C \ sets (restrict_space M A)" "(\?C) = space (restrict_space M A)" by(auto simp add: sets_restrict_space space_restrict_space) moreover { fix a assume "a \ ?C" then obtain a' where "a = A \ a'" "a' \ C" .. then have "emeasure (restrict_space M A) a \ emeasure M a'" using A C by(auto simp add: emeasure_restrict_space intro: emeasure_mono) also have "\ < \" using C(4)[rule_format, of a'] \a' \ C\ by (simp add: less_top) finally have "emeasure (restrict_space M A) a \ \" by simp } ultimately show ?thesis by unfold_locales (rule exI conjI|assumption|blast)+ qed lemma finite_measure_restrict_space: assumes "finite_measure M" and A: "A \ sets M" shows "finite_measure (restrict_space M A)" proof - interpret finite_measure M by fact show ?thesis by(rule finite_measureI)(simp add: emeasure_restrict_space A space_restrict_space) qed lemma restrict_distr: assumes [measurable]: "f \ measurable M N" assumes [simp]: "\ \ space N \ sets N" and restrict: "f \ space M \ \" shows "restrict_space (distr M N f) \ = distr M (restrict_space N \) f" (is "?l = ?r") proof (rule measure_eqI) fix A assume "A \ sets (restrict_space (distr M N f) \)" with restrict show "emeasure ?l A = emeasure ?r A" by (subst emeasure_distr) (auto simp: sets_restrict_space_iff emeasure_restrict_space emeasure_distr intro!: measurable_restrict_space2) qed (simp add: sets_restrict_space) lemma measure_eqI_restrict_generator: assumes E: "Int_stable E" "E \ Pow \" "\X. X \ E \ emeasure M X = emeasure N X" assumes sets_eq: "sets M = sets N" and \: "\ \ sets M" assumes "sets (restrict_space M \) = sigma_sets \ E" assumes "sets (restrict_space N \) = sigma_sets \ E" assumes ae: "AE x in M. x \ \" "AE x in N. x \ \" assumes A: "countable A" "A \ {}" "A \ E" "\A = \" "\a. a \ A \ emeasure M a \ \" shows "M = N" proof (rule measure_eqI) fix X assume X: "X \ sets M" then have "emeasure M X = emeasure (restrict_space M \) (X \ \)" using ae \ by (auto simp add: emeasure_restrict_space intro!: emeasure_eq_AE) also have "restrict_space M \ = restrict_space N \" proof (rule measure_eqI_generator_eq) fix X assume "X \ E" then show "emeasure (restrict_space M \) X = emeasure (restrict_space N \) X" using E \ by (subst (1 2) emeasure_restrict_space) (auto simp: sets_eq sets_eq[THEN sets_eq_imp_space_eq]) next show "range (from_nat_into A) \ E" "(\i. from_nat_into A i) = \" using A by (auto cong del: SUP_cong_simp) next fix i have "emeasure (restrict_space M \) (from_nat_into A i) = emeasure M (from_nat_into A i)" using A \ by (subst emeasure_restrict_space) (auto simp: sets_eq sets_eq[THEN sets_eq_imp_space_eq] intro: from_nat_into) with A show "emeasure (restrict_space M \) (from_nat_into A i) \ \" by (auto intro: from_nat_into) qed fact+ also have "emeasure (restrict_space N \) (X \ \) = emeasure N X" using X ae \ by (auto simp add: emeasure_restrict_space sets_eq intro!: emeasure_eq_AE) finally show "emeasure M X = emeasure N X" . qed fact subsection\<^marker>\tag unimportant\ \Null measure\ definition null_measure :: "'a measure \ 'a measure" where "null_measure M = sigma (space M) (sets M)" lemma space_null_measure[simp]: "space (null_measure M) = space M" by (simp add: null_measure_def) lemma sets_null_measure[simp, measurable_cong]: "sets (null_measure M) = sets M" by (simp add: null_measure_def) lemma emeasure_null_measure[simp]: "emeasure (null_measure M) X = 0" by (cases "X \ sets M", rule emeasure_measure_of) (auto simp: positive_def countably_additive_def emeasure_notin_sets null_measure_def dest: sets.sets_into_space) lemma measure_null_measure[simp]: "measure (null_measure M) X = 0" by (intro measure_eq_emeasure_eq_ennreal) auto lemma null_measure_idem [simp]: "null_measure (null_measure M) = null_measure M" by(rule measure_eqI) simp_all subsection \Scaling a measure\ definition\<^marker>\tag important\ scale_measure :: "ennreal \ 'a measure \ 'a measure" where "scale_measure r M = measure_of (space M) (sets M) (\A. r * emeasure M A)" lemma space_scale_measure: "space (scale_measure r M) = space M" by (simp add: scale_measure_def) lemma sets_scale_measure [simp, measurable_cong]: "sets (scale_measure r M) = sets M" by (simp add: scale_measure_def) lemma emeasure_scale_measure [simp]: "emeasure (scale_measure r M) A = r * emeasure M A" (is "_ = ?\ A") proof(cases "A \ sets M") case True show ?thesis unfolding scale_measure_def proof(rule emeasure_measure_of_sigma) show "sigma_algebra (space M) (sets M)" .. show "positive (sets M) ?\" by (simp add: positive_def) show "countably_additive (sets M) ?\" proof (rule countably_additiveI) fix A :: "nat \ _" assume *: "range A \ sets M" "disjoint_family A" have "(\i. ?\ (A i)) = r * (\i. emeasure M (A i))" by simp also have "\ = ?\ (\i. A i)" using * by(simp add: suminf_emeasure) finally show "(\i. ?\ (A i)) = ?\ (\i. A i)" . qed qed(fact True) qed(simp add: emeasure_notin_sets) lemma scale_measure_1 [simp]: "scale_measure 1 M = M" by(rule measure_eqI) simp_all lemma scale_measure_0[simp]: "scale_measure 0 M = null_measure M" by(rule measure_eqI) simp_all lemma measure_scale_measure [simp]: "0 \ r \ measure (scale_measure r M) A = r * measure M A" using emeasure_scale_measure[of r M A] emeasure_eq_ennreal_measure[of M A] measure_eq_emeasure_eq_ennreal[of _ "scale_measure r M" A] by (cases "emeasure (scale_measure r M) A = top") (auto simp del: emeasure_scale_measure simp: ennreal_top_eq_mult_iff ennreal_mult_eq_top_iff measure_zero_top ennreal_mult[symmetric]) lemma scale_scale_measure [simp]: "scale_measure r (scale_measure r' M) = scale_measure (r * r') M" by (rule measure_eqI) (simp_all add: max_def mult.assoc) lemma scale_null_measure [simp]: "scale_measure r (null_measure M) = null_measure M" by (rule measure_eqI) simp_all subsection \Complete lattice structure on measures\ lemma (in finite_measure) finite_measure_Diff': "A \ sets M \ B \ sets M \ measure M (A - B) = measure M A - measure M (A \ B)" using finite_measure_Diff[of A "A \ B"] by (auto simp: Diff_Int) lemma (in finite_measure) finite_measure_Union': "A \ sets M \ B \ sets M \ measure M (A \ B) = measure M A + measure M (B - A)" using finite_measure_Union[of A "B - A"] by auto lemma finite_unsigned_Hahn_decomposition: assumes "finite_measure M" "finite_measure N" and [simp]: "sets N = sets M" shows "\Y\sets M. (\X\sets M. X \ Y \ N X \ M X) \ (\X\sets M. X \ Y = {} \ M X \ N X)" proof - interpret M: finite_measure M by fact interpret N: finite_measure N by fact define d where "d X = measure M X - measure N X" for X have [intro]: "bdd_above (d`sets M)" using sets.sets_into_space[of _ M] by (intro bdd_aboveI[where M="measure M (space M)"]) (auto simp: d_def field_simps subset_eq intro!: add_increasing M.finite_measure_mono) define \ where "\ = (SUP X\sets M. d X)" have le_\[intro]: "X \ sets M \ d X \ \" for X by (auto simp: \_def intro!: cSUP_upper) have "\f. \n. f n \ sets M \ d (f n) > \ - 1 / 2^n" proof (intro choice_iff[THEN iffD1] allI) fix n have "\X\sets M. \ - 1 / 2^n < d X" unfolding \_def by (intro less_cSUP_iff[THEN iffD1]) auto then show "\y. y \ sets M \ \ - 1 / 2 ^ n < d y" by auto qed then obtain E where [measurable]: "E n \ sets M" and E: "d (E n) > \ - 1 / 2^n" for n by auto define F where "F m n = (if m \ n then \i\{m..n}. E i else space M)" for m n have [measurable]: "m \ n \ F m n \ sets M" for m n by (auto simp: F_def) have 1: "\ - 2 / 2 ^ m + 1 / 2 ^ n \ d (F m n)" if "m \ n" for m n using that proof (induct rule: dec_induct) case base with E[of m] show ?case by (simp add: F_def field_simps) next case (step i) have F_Suc: "F m (Suc i) = F m i \ E (Suc i)" using \m \ i\ by (auto simp: F_def le_Suc_eq) have "\ + (\ - 2 / 2^m + 1 / 2 ^ Suc i) \ (\ - 1 / 2^Suc i) + (\ - 2 / 2^m + 1 / 2^i)" by (simp add: field_simps) also have "\ \ d (E (Suc i)) + d (F m i)" using E[of "Suc i"] by (intro add_mono step) auto also have "\ = d (E (Suc i)) + d (F m i - E (Suc i)) + d (F m (Suc i))" using \m \ i\ by (simp add: d_def field_simps F_Suc M.finite_measure_Diff' N.finite_measure_Diff') also have "\ = d (E (Suc i) \ F m i) + d (F m (Suc i))" using \m \ i\ by (simp add: d_def field_simps M.finite_measure_Union' N.finite_measure_Union') also have "\ \ \ + d (F m (Suc i))" using \m \ i\ by auto finally show ?case by auto qed define F' where "F' m = (\i\{m..}. E i)" for m have F'_eq: "F' m = (\i. F m (i + m))" for m by (fastforce simp: le_iff_add[of m] F'_def F_def) have [measurable]: "F' m \ sets M" for m by (auto simp: F'_def) have \_le: "\ - 0 \ d (\m. F' m)" proof (rule LIMSEQ_le) show "(\n. \ - 2 / 2 ^ n) \ \ - 0" by (intro tendsto_intros LIMSEQ_divide_realpow_zero) auto have "incseq F'" by (auto simp: incseq_def F'_def) then show "(\m. d (F' m)) \ d (\m. F' m)" unfolding d_def by (intro tendsto_diff M.finite_Lim_measure_incseq N.finite_Lim_measure_incseq) auto have "\ - 2 / 2 ^ m + 0 \ d (F' m)" for m proof (rule LIMSEQ_le) have *: "decseq (\n. F m (n + m))" by (auto simp: decseq_def F_def) show "(\n. d (F m n)) \ d (F' m)" unfolding d_def F'_eq by (rule LIMSEQ_offset[where k=m]) (auto intro!: tendsto_diff M.finite_Lim_measure_decseq N.finite_Lim_measure_decseq *) show "(\n. \ - 2 / 2 ^ m + 1 / 2 ^ n) \ \ - 2 / 2 ^ m + 0" by (intro tendsto_add LIMSEQ_divide_realpow_zero tendsto_const) auto show "\N. \n\N. \ - 2 / 2 ^ m + 1 / 2 ^ n \ d (F m n)" using 1[of m] by (intro exI[of _ m]) auto qed then show "\N. \n\N. \ - 2 / 2 ^ n \ d (F' n)" by auto qed show ?thesis proof (safe intro!: bexI[of _ "\m. F' m"]) fix X assume [measurable]: "X \ sets M" and X: "X \ (\m. F' m)" have "d (\m. F' m) - d X = d ((\m. F' m) - X)" using X by (auto simp: d_def M.finite_measure_Diff N.finite_measure_Diff) also have "\ \ \" by auto finally have "0 \ d X" using \_le by auto then show "emeasure N X \ emeasure M X" by (auto simp: d_def M.emeasure_eq_measure N.emeasure_eq_measure) next fix X assume [measurable]: "X \ sets M" and X: "X \ (\m. F' m) = {}" then have "d (\m. F' m) + d X = d (X \ (\m. F' m))" by (auto simp: d_def M.finite_measure_Union N.finite_measure_Union) also have "\ \ \" by auto finally have "d X \ 0" using \_le by auto then show "emeasure M X \ emeasure N X" by (auto simp: d_def M.emeasure_eq_measure N.emeasure_eq_measure) qed auto qed proposition unsigned_Hahn_decomposition: assumes [simp]: "sets N = sets M" and [measurable]: "A \ sets M" and [simp]: "emeasure M A \ top" "emeasure N A \ top" shows "\Y\sets M. Y \ A \ (\X\sets M. X \ Y \ N X \ M X) \ (\X\sets M. X \ A \ X \ Y = {} \ M X \ N X)" proof - have "\Y\sets (restrict_space M A). (\X\sets (restrict_space M A). X \ Y \ (restrict_space N A) X \ (restrict_space M A) X) \ (\X\sets (restrict_space M A). X \ Y = {} \ (restrict_space M A) X \ (restrict_space N A) X)" proof (rule finite_unsigned_Hahn_decomposition) show "finite_measure (restrict_space M A)" "finite_measure (restrict_space N A)" by (auto simp: space_restrict_space emeasure_restrict_space less_top intro!: finite_measureI) qed (simp add: sets_restrict_space) then guess Y .. then show ?thesis apply (intro bexI[of _ Y] conjI ballI conjI) apply (simp_all add: sets_restrict_space emeasure_restrict_space) apply safe subgoal for X Z by (erule ballE[of _ _ X]) (auto simp add: Int_absorb1) subgoal for X Z by (erule ballE[of _ _ X]) (auto simp add: Int_absorb1 ac_simps) apply auto done qed text\<^marker>\tag important\ \ Define a lexicographical order on \<^type>\measure\, in the order space, sets and measure. The parts of the lexicographical order are point-wise ordered. \ instantiation measure :: (type) order_bot begin inductive less_eq_measure :: "'a measure \ 'a measure \ bool" where "space M \ space N \ less_eq_measure M N" | "space M = space N \ sets M \ sets N \ less_eq_measure M N" | "space M = space N \ sets M = sets N \ emeasure M \ emeasure N \ less_eq_measure M N" lemma le_measure_iff: "M \ N \ (if space M = space N then if sets M = sets N then emeasure M \ emeasure N else sets M \ sets N else space M \ space N)" by (auto elim: less_eq_measure.cases intro: less_eq_measure.intros) definition\<^marker>\tag important\ less_measure :: "'a measure \ 'a measure \ bool" where "less_measure M N \ (M \ N \ \ N \ M)" definition\<^marker>\tag important\ bot_measure :: "'a measure" where "bot_measure = sigma {} {}" lemma shows space_bot[simp]: "space bot = {}" and sets_bot[simp]: "sets bot = {{}}" and emeasure_bot[simp]: "emeasure bot X = 0" by (auto simp: bot_measure_def sigma_sets_empty_eq emeasure_sigma) instance proof standard show "bot \ a" for a :: "'a measure" by (simp add: le_measure_iff bot_measure_def sigma_sets_empty_eq emeasure_sigma le_fun_def) qed (auto simp: le_measure_iff less_measure_def split: if_split_asm intro: measure_eqI) end proposition le_measure: "sets M = sets N \ M \ N \ (\A\sets M. emeasure M A \ emeasure N A)" apply - apply (auto simp: le_measure_iff le_fun_def dest: sets_eq_imp_space_eq) subgoal for X by (cases "X \ sets M") (auto simp: emeasure_notin_sets) done definition\<^marker>\tag important\ sup_measure' :: "'a measure \ 'a measure \ 'a measure" where "sup_measure' A B = measure_of (space A) (sets A) (\X. SUP Y\sets A. emeasure A (X \ Y) + emeasure B (X \ - Y))" lemma assumes [simp]: "sets B = sets A" shows space_sup_measure'[simp]: "space (sup_measure' A B) = space A" and sets_sup_measure'[simp]: "sets (sup_measure' A B) = sets A" using sets_eq_imp_space_eq[OF assms] by (simp_all add: sup_measure'_def) lemma emeasure_sup_measure': assumes sets_eq[simp]: "sets B = sets A" and [simp, intro]: "X \ sets A" shows "emeasure (sup_measure' A B) X = (SUP Y\sets A. emeasure A (X \ Y) + emeasure B (X \ - Y))" (is "_ = ?S X") proof - note sets_eq_imp_space_eq[OF sets_eq, simp] show ?thesis using sup_measure'_def proof (rule emeasure_measure_of) let ?d = "\X Y. emeasure A (X \ Y) + emeasure B (X \ - Y)" show "countably_additive (sets (sup_measure' A B)) (\X. SUP Y \ sets A. emeasure A (X \ Y) + emeasure B (X \ - Y))" proof (rule countably_additiveI, goal_cases) case (1 X) then have [measurable]: "\i. X i \ sets A" and "disjoint_family X" by auto have disjoint: "disjoint_family (\i. X i \ Y)" "disjoint_family (\i. X i - Y)" for Y by (auto intro: disjoint_family_on_bisimulation [OF \disjoint_family X\, simplified]) have "(\i. ?S (X i)) = (SUP Y\sets A. \i. ?d (X i) Y)" proof (rule ennreal_suminf_SUP_eq_directed) fix J :: "nat set" and a b assume "finite J" and [measurable]: "a \ sets A" "b \ sets A" have "\c\sets A. c \ X i \ (\a\sets A. ?d (X i) a \ ?d (X i) c)" for i proof cases assume "emeasure A (X i) = top \ emeasure B (X i) = top" then show ?thesis proof assume "emeasure A (X i) = top" then show ?thesis by (intro bexI[of _ "X i"]) auto next assume "emeasure B (X i) = top" then show ?thesis by (intro bexI[of _ "{}"]) auto qed next assume finite: "\ (emeasure A (X i) = top \ emeasure B (X i) = top)" then have "\Y\sets A. Y \ X i \ (\C\sets A. C \ Y \ B C \ A C) \ (\C\sets A. C \ X i \ C \ Y = {} \ A C \ B C)" using unsigned_Hahn_decomposition[of B A "X i"] by simp then obtain Y where [measurable]: "Y \ sets A" and [simp]: "Y \ X i" and B_le_A: "\C. C \ sets A \ C \ Y \ B C \ A C" and A_le_B: "\C. C \ sets A \ C \ X i \ C \ Y = {} \ A C \ B C" by auto show ?thesis proof (intro bexI[of _ Y] ballI conjI) fix a assume [measurable]: "a \ sets A" have *: "(X i \ a \ Y \ (X i \ a - Y)) = X i \ a" "(X i - a) \ Y \ (X i - a - Y) = X i \ - a" for a Y by auto then have "?d (X i) a = (A (X i \ a \ Y) + A (X i \ a \ - Y)) + (B (X i \ - a \ Y) + B (X i \ - a \ - Y))" by (subst (1 2) plus_emeasure) (auto simp: Diff_eq[symmetric]) also have "\ \ (A (X i \ a \ Y) + B (X i \ a \ - Y)) + (A (X i \ - a \ Y) + B (X i \ - a \ - Y))" by (intro add_mono order_refl B_le_A A_le_B) (auto simp: Diff_eq[symmetric]) also have "\ \ (A (X i \ Y \ a) + A (X i \ Y \ - a)) + (B (X i \ - Y \ a) + B (X i \ - Y \ - a))" by (simp add: ac_simps) also have "\ \ A (X i \ Y) + B (X i \ - Y)" by (subst (1 2) plus_emeasure) (auto simp: Diff_eq[symmetric] *) finally show "?d (X i) a \ ?d (X i) Y" . qed auto qed then obtain C where [measurable]: "C i \ sets A" and "C i \ X i" and C: "\a. a \ sets A \ ?d (X i) a \ ?d (X i) (C i)" for i by metis have *: "X i \ (\i. C i) = X i \ C i" for i proof safe fix x j assume "x \ X i" "x \ C j" moreover have "i = j \ X i \ X j = {}" using \disjoint_family X\ by (auto simp: disjoint_family_on_def) ultimately show "x \ C i" using \C i \ X i\ \C j \ X j\ by auto qed auto have **: "X i \ - (\i. C i) = X i \ - C i" for i proof safe fix x j assume "x \ X i" "x \ C i" "x \ C j" moreover have "i = j \ X i \ X j = {}" using \disjoint_family X\ by (auto simp: disjoint_family_on_def) ultimately show False using \C i \ X i\ \C j \ X j\ by auto qed auto show "\c\sets A. \i\J. ?d (X i) a \ ?d (X i) c \ ?d (X i) b \ ?d (X i) c" apply (intro bexI[of _ "\i. C i"]) unfolding * ** apply (intro C ballI conjI) apply auto done qed also have "\ = ?S (\i. X i)" apply (simp only: UN_extend_simps(4)) apply (rule arg_cong [of _ _ Sup]) apply (rule image_cong) apply (fact refl) using disjoint apply (auto simp add: suminf_add [symmetric] Diff_eq [symmetric] image_subset_iff suminf_emeasure simp del: UN_simps) done finally show "(\i. ?S (X i)) = ?S (\i. X i)" . qed qed (auto dest: sets.sets_into_space simp: positive_def intro!: SUP_const) qed lemma le_emeasure_sup_measure'1: assumes "sets B = sets A" "X \ sets A" shows "emeasure A X \ emeasure (sup_measure' A B) X" by (subst emeasure_sup_measure'[OF assms]) (auto intro!: SUP_upper2[of "X"] assms) lemma le_emeasure_sup_measure'2: assumes "sets B = sets A" "X \ sets A" shows "emeasure B X \ emeasure (sup_measure' A B) X" by (subst emeasure_sup_measure'[OF assms]) (auto intro!: SUP_upper2[of "{}"] assms) lemma emeasure_sup_measure'_le2: assumes [simp]: "sets B = sets C" "sets A = sets C" and [measurable]: "X \ sets C" assumes A: "\Y. Y \ X \ Y \ sets A \ emeasure A Y \ emeasure C Y" assumes B: "\Y. Y \ X \ Y \ sets A \ emeasure B Y \ emeasure C Y" shows "emeasure (sup_measure' A B) X \ emeasure C X" proof (subst emeasure_sup_measure') show "(SUP Y\sets A. emeasure A (X \ Y) + emeasure B (X \ - Y)) \ emeasure C X" unfolding \sets A = sets C\ proof (intro SUP_least) fix Y assume [measurable]: "Y \ sets C" have [simp]: "X \ Y \ (X - Y) = X" by auto have "emeasure A (X \ Y) + emeasure B (X \ - Y) \ emeasure C (X \ Y) + emeasure C (X \ - Y)" by (intro add_mono A B) (auto simp: Diff_eq[symmetric]) also have "\ = emeasure C X" by (subst plus_emeasure) (auto simp: Diff_eq[symmetric]) finally show "emeasure A (X \ Y) + emeasure B (X \ - Y) \ emeasure C X" . qed qed simp_all definition\<^marker>\tag important\ sup_lexord :: "'a \ 'a \ ('a \ 'b::order) \ 'a \ 'a \ 'a" where "sup_lexord A B k s c = (if k A = k B then c else if \ k A \ k B \ \ k B \ k A then s else if k B \ k A then A else B)" lemma sup_lexord: "(k A < k B \ P B) \ (k B < k A \ P A) \ (k A = k B \ P c) \ (\ k B \ k A \ \ k A \ k B \ P s) \ P (sup_lexord A B k s c)" by (auto simp: sup_lexord_def) lemmas le_sup_lexord = sup_lexord[where P="\a. c \ a" for c] lemma sup_lexord1: "k A = k B \ sup_lexord A B k s c = c" by (simp add: sup_lexord_def) lemma sup_lexord_commute: "sup_lexord A B k s c = sup_lexord B A k s c" by (auto simp: sup_lexord_def) lemma sigma_sets_le_sets_iff: "(sigma_sets (space x) \ \ sets x) = (\ \ sets x)" using sets.sigma_sets_subset[of \ x] by auto lemma sigma_le_iff: "\ \ Pow \ \ sigma \ \ \ x \ (\ \ space x \ (space x = \ \ \ \ sets x))" by (cases "\ = space x") (simp_all add: eq_commute[of _ "sets x"] le_measure_iff emeasure_sigma le_fun_def sigma_sets_superset_generator sigma_sets_le_sets_iff) instantiation measure :: (type) semilattice_sup begin definition\<^marker>\tag important\ sup_measure :: "'a measure \ 'a measure \ 'a measure" where "sup_measure A B = sup_lexord A B space (sigma (space A \ space B) {}) (sup_lexord A B sets (sigma (space A) (sets A \ sets B)) (sup_measure' A B))" instance proof fix x y z :: "'a measure" show "x \ sup x y" unfolding sup_measure_def proof (intro le_sup_lexord) assume "space x = space y" then have *: "sets x \ sets y \ Pow (space x)" using sets.space_closed by auto assume "\ sets y \ sets x" "\ sets x \ sets y" then have "sets x \ sets x \ sets y" by auto also have "\ \ sigma (space x) (sets x \ sets y)" by (subst sets_measure_of[OF *]) (rule sigma_sets_superset_generator) finally show "x \ sigma (space x) (sets x \ sets y)" by (simp add: space_measure_of[OF *] less_eq_measure.intros(2)) next assume "\ space y \ space x" "\ space x \ space y" then show "x \ sigma (space x \ space y) {}" by (intro less_eq_measure.intros) auto next assume "sets x = sets y" then show "x \ sup_measure' x y" by (simp add: le_measure le_emeasure_sup_measure'1) qed (auto intro: less_eq_measure.intros) show "y \ sup x y" unfolding sup_measure_def proof (intro le_sup_lexord) assume **: "space x = space y" then have *: "sets x \ sets y \ Pow (space y)" using sets.space_closed by auto assume "\ sets y \ sets x" "\ sets x \ sets y" then have "sets y \ sets x \ sets y" by auto also have "\ \ sigma (space y) (sets x \ sets y)" by (subst sets_measure_of[OF *]) (rule sigma_sets_superset_generator) finally show "y \ sigma (space x) (sets x \ sets y)" by (simp add: ** space_measure_of[OF *] less_eq_measure.intros(2)) next assume "\ space y \ space x" "\ space x \ space y" then show "y \ sigma (space x \ space y) {}" by (intro less_eq_measure.intros) auto next assume "sets x = sets y" then show "y \ sup_measure' x y" by (simp add: le_measure le_emeasure_sup_measure'2) qed (auto intro: less_eq_measure.intros) show "x \ y \ z \ y \ sup x z \ y" unfolding sup_measure_def proof (intro sup_lexord[where P="\x. x \ y"]) assume "x \ y" "z \ y" and [simp]: "space x = space z" "sets x = sets z" from \x \ y\ show "sup_measure' x z \ y" proof cases case 1 then show ?thesis by (intro less_eq_measure.intros(1)) simp next case 2 then show ?thesis by (intro less_eq_measure.intros(2)) simp_all next case 3 with \z \ y\ \x \ y\ show ?thesis by (auto simp add: le_measure intro!: emeasure_sup_measure'_le2) qed next assume **: "x \ y" "z \ y" "space x = space z" "\ sets z \ sets x" "\ sets x \ sets z" then have *: "sets x \ sets z \ Pow (space x)" using sets.space_closed by auto show "sigma (space x) (sets x \ sets z) \ y" unfolding sigma_le_iff[OF *] using ** by (auto simp: le_measure_iff split: if_split_asm) next assume "x \ y" "z \ y" "\ space z \ space x" "\ space x \ space z" then have "space x \ space y" "space z \ space y" by (auto simp: le_measure_iff split: if_split_asm) then show "sigma (space x \ space z) {} \ y" by (simp add: sigma_le_iff) qed qed end lemma space_empty_eq_bot: "space a = {} \ a = bot" using space_empty[of a] by (auto intro!: measure_eqI) lemma sets_eq_iff_bounded: "A \ B \ B \ C \ sets A = sets C \ sets B = sets A" by (auto dest: sets_eq_imp_space_eq simp add: le_measure_iff split: if_split_asm) lemma sets_sup: "sets A = sets M \ sets B = sets M \ sets (sup A B) = sets M" by (auto simp add: sup_measure_def sup_lexord_def dest: sets_eq_imp_space_eq) lemma le_measureD1: "A \ B \ space A \ space B" by (auto simp: le_measure_iff split: if_split_asm) lemma le_measureD2: "A \ B \ space A = space B \ sets A \ sets B" by (auto simp: le_measure_iff split: if_split_asm) lemma le_measureD3: "A \ B \ sets A = sets B \ emeasure A X \ emeasure B X" by (auto simp: le_measure_iff le_fun_def dest: sets_eq_imp_space_eq split: if_split_asm) lemma UN_space_closed: "\(sets ` S) \ Pow (\(space ` S))" using sets.space_closed by auto definition\<^marker>\tag important\ Sup_lexord :: "('a \ 'b::complete_lattice) \ ('a set \ 'a) \ ('a set \ 'a) \ 'a set \ 'a" where "Sup_lexord k c s A = (let U = (SUP a\A. k a) in if \a\A. k a = U then c {a\A. k a = U} else s A)" lemma Sup_lexord: "(\a S. a \ A \ k a = (SUP a\A. k a) \ S = {a'\A. k a' = k a} \ P (c S)) \ ((\a. a \ A \ k a \ (SUP a\A. k a)) \ P (s A)) \ P (Sup_lexord k c s A)" by (auto simp: Sup_lexord_def Let_def) lemma Sup_lexord1: assumes A: "A \ {}" "(\a. a \ A \ k a = (\a\A. k a))" "P (c A)" shows "P (Sup_lexord k c s A)" unfolding Sup_lexord_def Let_def proof (clarsimp, safe) show "\a\A. k a \ (\x\A. k x) \ P (s A)" by (metis assms(1,2) ex_in_conv) next fix a assume "a \ A" "k a = (\x\A. k x)" then have "{a \ A. k a = (\x\A. k x)} = {a \ A. k a = k a}" by (metis A(2)[symmetric]) then show "P (c {a \ A. k a = (\x\A. k x)})" by (simp add: A(3)) qed instantiation measure :: (type) complete_lattice begin interpretation sup_measure: comm_monoid_set sup "bot :: 'a measure" by standard (auto intro!: antisym) lemma sup_measure_F_mono': "finite J \ finite I \ sup_measure.F id I \ sup_measure.F id (I \ J)" proof (induction J rule: finite_induct) case empty then show ?case by simp next case (insert i J) show ?case proof cases assume "i \ I" with insert show ?thesis by (auto simp: insert_absorb) next assume "i \ I" have "sup_measure.F id I \ sup_measure.F id (I \ J)" by (intro insert) also have "\ \ sup_measure.F id (insert i (I \ J))" using insert \i \ I\ by (subst sup_measure.insert) auto finally show ?thesis by auto qed qed lemma sup_measure_F_mono: "finite I \ J \ I \ sup_measure.F id J \ sup_measure.F id I" using sup_measure_F_mono'[of I J] by (auto simp: finite_subset Un_absorb1) lemma sets_sup_measure_F: "finite I \ I \ {} \ (\i. i \ I \ sets i = sets M) \ sets (sup_measure.F id I) = sets M" by (induction I rule: finite_ne_induct) (simp_all add: sets_sup) definition\<^marker>\tag important\ Sup_measure' :: "'a measure set \ 'a measure" where "Sup_measure' M = measure_of (\a\M. space a) (\a\M. sets a) (\X. (SUP P\{P. finite P \ P \ M }. sup_measure.F id P X))" lemma space_Sup_measure'2: "space (Sup_measure' M) = (\m\M. space m)" unfolding Sup_measure'_def by (intro space_measure_of[OF UN_space_closed]) lemma sets_Sup_measure'2: "sets (Sup_measure' M) = sigma_sets (\m\M. space m) (\m\M. sets m)" unfolding Sup_measure'_def by (intro sets_measure_of[OF UN_space_closed]) lemma sets_Sup_measure': assumes sets_eq[simp]: "\m. m \ M \ sets m = sets A" and "M \ {}" shows "sets (Sup_measure' M) = sets A" using sets_eq[THEN sets_eq_imp_space_eq, simp] \M \ {}\ by (simp add: Sup_measure'_def) lemma space_Sup_measure': assumes sets_eq[simp]: "\m. m \ M \ sets m = sets A" and "M \ {}" shows "space (Sup_measure' M) = space A" using sets_eq[THEN sets_eq_imp_space_eq, simp] \M \ {}\ by (simp add: Sup_measure'_def ) lemma emeasure_Sup_measure': assumes sets_eq[simp]: "\m. m \ M \ sets m = sets A" and "X \ sets A" "M \ {}" shows "emeasure (Sup_measure' M) X = (SUP P\{P. finite P \ P \ M}. sup_measure.F id P X)" (is "_ = ?S X") using Sup_measure'_def proof (rule emeasure_measure_of) note sets_eq[THEN sets_eq_imp_space_eq, simp] have *: "sets (Sup_measure' M) = sets A" "space (Sup_measure' M) = space A" using \M \ {}\ by (simp_all add: Sup_measure'_def) let ?\ = "sup_measure.F id" show "countably_additive (sets (Sup_measure' M)) ?S" proof (rule countably_additiveI, goal_cases) case (1 F) then have **: "range F \ sets A" by (auto simp: *) show "(\i. ?S (F i)) = ?S (\i. F i)" proof (subst ennreal_suminf_SUP_eq_directed) fix i j and N :: "nat set" assume ij: "i \ {P. finite P \ P \ M}" "j \ {P. finite P \ P \ M}" have "(i \ {} \ sets (?\ i) = sets A) \ (j \ {} \ sets (?\ j) = sets A) \ (i \ {} \ j \ {} \ sets (?\ (i \ j)) = sets A)" using ij by (intro impI sets_sup_measure_F conjI) auto then have "?\ j (F n) \ ?\ (i \ j) (F n) \ ?\ i (F n) \ ?\ (i \ j) (F n)" for n using ij by (cases "i = {}"; cases "j = {}") (auto intro!: le_measureD3 sup_measure_F_mono simp: sets_sup_measure_F simp del: id_apply) with ij show "\k\{P. finite P \ P \ M}. \n\N. ?\ i (F n) \ ?\ k (F n) \ ?\ j (F n) \ ?\ k (F n)" by (safe intro!: bexI[of _ "i \ j"]) auto next show "(SUP P \ {P. finite P \ P \ M}. \n. ?\ P (F n)) = (SUP P \ {P. finite P \ P \ M}. ?\ P (\(F ` UNIV)))" proof (intro arg_cong [of _ _ Sup] image_cong refl) fix i assume i: "i \ {P. finite P \ P \ M}" show "(\n. ?\ i (F n)) = ?\ i (\(F ` UNIV))" proof cases assume "i \ {}" with i ** show ?thesis apply (intro suminf_emeasure \disjoint_family F\) apply (subst sets_sup_measure_F[OF _ _ sets_eq]) apply auto done qed simp qed qed qed show "positive (sets (Sup_measure' M)) ?S" by (auto simp: positive_def bot_ennreal[symmetric]) show "X \ sets (Sup_measure' M)" using assms * by auto qed (rule UN_space_closed) definition\<^marker>\tag important\ Sup_measure :: "'a measure set \ 'a measure" where "Sup_measure = Sup_lexord space (Sup_lexord sets Sup_measure' (\U. sigma (\u\U. space u) (\u\U. sets u))) (\U. sigma (\u\U. space u) {})" definition\<^marker>\tag important\ Inf_measure :: "'a measure set \ 'a measure" where "Inf_measure A = Sup {x. \a\A. x \ a}" definition\<^marker>\tag important\ inf_measure :: "'a measure \ 'a measure \ 'a measure" where "inf_measure a b = Inf {a, b}" definition\<^marker>\tag important\ top_measure :: "'a measure" where "top_measure = Inf {}" instance proof note UN_space_closed [simp] show upper: "x \ Sup A" if x: "x \ A" for x :: "'a measure" and A unfolding Sup_measure_def proof (intro Sup_lexord[where P="\y. x \ y"]) assume "\a. a \ A \ space a \ (\a\A. space a)" from this[OF \x \ A\] \x \ A\ show "x \ sigma (\a\A. space a) {}" by (intro less_eq_measure.intros) auto next fix a S assume "a \ A" and a: "space a = (\a\A. space a)" and S: "S = {a' \ A. space a' = space a}" and neq: "\aa. aa \ S \ sets aa \ (\a\S. sets a)" have sp_a: "space a = (\(space ` S))" using \a\A\ by (auto simp: S) show "x \ sigma (\(space ` S)) (\(sets ` S))" proof cases assume [simp]: "space x = space a" have "sets x \ (\a\S. sets a)" using \x\A\ neq[of x] by (auto simp: S) also have "\ \ sigma_sets (\x\S. space x) (\x\S. sets x)" by (rule sigma_sets_superset_generator) finally show ?thesis by (intro less_eq_measure.intros(2)) (simp_all add: sp_a) next assume "space x \ space a" moreover have "space x \ space a" unfolding a using \x\A\ by auto ultimately show ?thesis by (intro less_eq_measure.intros) (simp add: less_le sp_a) qed next fix a b S S' assume "a \ A" and a: "space a = (\a\A. space a)" and S: "S = {a' \ A. space a' = space a}" and "b \ S" and b: "sets b = (\a\S. sets a)" and S': "S' = {a' \ S. sets a' = sets b}" then have "S' \ {}" "space b = space a" by auto have sets_eq: "\x. x \ S' \ sets x = sets b" by (auto simp: S') note sets_eq[THEN sets_eq_imp_space_eq, simp] have *: "sets (Sup_measure' S') = sets b" "space (Sup_measure' S') = space b" using \S' \ {}\ by (simp_all add: Sup_measure'_def sets_eq) show "x \ Sup_measure' S'" proof cases assume "x \ S" with \b \ S\ have "space x = space b" by (simp add: S) show ?thesis proof cases assume "x \ S'" show "x \ Sup_measure' S'" proof (intro le_measure[THEN iffD2] ballI) show "sets x = sets (Sup_measure' S')" using \x\S'\ * by (simp add: S') fix X assume "X \ sets x" show "emeasure x X \ emeasure (Sup_measure' S') X" proof (subst emeasure_Sup_measure'[OF _ \X \ sets x\]) show "emeasure x X \ (SUP P \ {P. finite P \ P \ S'}. emeasure (sup_measure.F id P) X)" using \x\S'\ by (intro SUP_upper2[where i="{x}"]) auto qed (insert \x\S'\ S', auto) qed next assume "x \ S'" then have "sets x \ sets b" using \x\S\ by (auto simp: S') moreover have "sets x \ sets b" using \x\S\ unfolding b by auto ultimately show ?thesis using * \x \ S\ by (intro less_eq_measure.intros(2)) (simp_all add: * \space x = space b\ less_le) qed next assume "x \ S" with \x\A\ \x \ S\ \space b = space a\ show ?thesis by (intro less_eq_measure.intros) (simp_all add: * less_le a SUP_upper S) qed qed show least: "Sup A \ x" if x: "\z. z \ A \ z \ x" for x :: "'a measure" and A unfolding Sup_measure_def proof (intro Sup_lexord[where P="\y. y \ x"]) assume "\a. a \ A \ space a \ (\a\A. space a)" show "sigma (\(space ` A)) {} \ x" using x[THEN le_measureD1] by (subst sigma_le_iff) auto next fix a S assume "a \ A" "space a = (\a\A. space a)" and S: "S = {a' \ A. space a' = space a}" "\a. a \ S \ sets a \ (\a\S. sets a)" have "\(space ` S) \ space x" using S le_measureD1[OF x] by auto moreover have "\(space ` S) = space a" using \a\A\ S by auto then have "space x = \(space ` S) \ \(sets ` S) \ sets x" using \a \ A\ le_measureD2[OF x] by (auto simp: S) ultimately show "sigma (\(space ` S)) (\(sets ` S)) \ x" by (subst sigma_le_iff) simp_all next fix a b S S' assume "a \ A" and a: "space a = (\a\A. space a)" and S: "S = {a' \ A. space a' = space a}" and "b \ S" and b: "sets b = (\a\S. sets a)" and S': "S' = {a' \ S. sets a' = sets b}" then have "S' \ {}" "space b = space a" by auto have sets_eq: "\x. x \ S' \ sets x = sets b" by (auto simp: S') note sets_eq[THEN sets_eq_imp_space_eq, simp] have *: "sets (Sup_measure' S') = sets b" "space (Sup_measure' S') = space b" using \S' \ {}\ by (simp_all add: Sup_measure'_def sets_eq) show "Sup_measure' S' \ x" proof cases assume "space x = space a" show ?thesis proof cases assume **: "sets x = sets b" show ?thesis proof (intro le_measure[THEN iffD2] ballI) show ***: "sets (Sup_measure' S') = sets x" by (simp add: * **) fix X assume "X \ sets (Sup_measure' S')" show "emeasure (Sup_measure' S') X \ emeasure x X" unfolding *** proof (subst emeasure_Sup_measure'[OF _ \X \ sets (Sup_measure' S')\]) show "(SUP P \ {P. finite P \ P \ S'}. emeasure (sup_measure.F id P) X) \ emeasure x X" proof (safe intro!: SUP_least) fix P assume P: "finite P" "P \ S'" show "emeasure (sup_measure.F id P) X \ emeasure x X" proof cases assume "P = {}" then show ?thesis by auto next assume "P \ {}" from P have "finite P" "P \ A" unfolding S' S by (simp_all add: subset_eq) then have "sup_measure.F id P \ x" by (induction P) (auto simp: x) moreover have "sets (sup_measure.F id P) = sets x" using \finite P\ \P \ {}\ \P \ S'\ \sets x = sets b\ by (intro sets_sup_measure_F) (auto simp: S') ultimately show "emeasure (sup_measure.F id P) X \ emeasure x X" by (rule le_measureD3) qed qed show "m \ S' \ sets m = sets (Sup_measure' S')" for m unfolding * by (simp add: S') qed fact qed next assume "sets x \ sets b" moreover have "sets b \ sets x" unfolding b S using x[THEN le_measureD2] \space x = space a\ by auto ultimately show "Sup_measure' S' \ x" using \space x = space a\ \b \ S\ by (intro less_eq_measure.intros(2)) (simp_all add: * S) qed next assume "space x \ space a" then have "space a < space x" using le_measureD1[OF x[OF \a\A\]] by auto then show "Sup_measure' S' \ x" by (intro less_eq_measure.intros) (simp add: * \space b = space a\) qed qed show "Sup {} = (bot::'a measure)" "Inf {} = (top::'a measure)" by (auto intro!: antisym least simp: top_measure_def) show lower: "x \ A \ Inf A \ x" for x :: "'a measure" and A unfolding Inf_measure_def by (intro least) auto show greatest: "(\z. z \ A \ x \ z) \ x \ Inf A" for x :: "'a measure" and A unfolding Inf_measure_def by (intro upper) auto show "inf x y \ x" "inf x y \ y" "x \ y \ x \ z \ x \ inf y z" for x y z :: "'a measure" by (auto simp: inf_measure_def intro!: lower greatest) qed end lemma sets_SUP: assumes "\x. x \ I \ sets (M x) = sets N" shows "I \ {} \ sets (SUP i\I. M i) = sets N" unfolding Sup_measure_def using assms assms[THEN sets_eq_imp_space_eq] sets_Sup_measure'[where A=N and M="M`I"] by (intro Sup_lexord1[where P="\x. sets x = sets N"]) auto lemma emeasure_SUP: assumes sets: "\i. i \ I \ sets (M i) = sets N" "X \ sets N" "I \ {}" shows "emeasure (SUP i\I. M i) X = (SUP J\{J. J \ {} \ finite J \ J \ I}. emeasure (SUP i\J. M i) X)" proof - interpret sup_measure: comm_monoid_set sup "bot :: 'b measure" by standard (auto intro!: antisym) have eq: "finite J \ sup_measure.F id J = (SUP i\J. i)" for J :: "'b measure set" by (induction J rule: finite_induct) auto have 1: "J \ {} \ J \ I \ sets (SUP x\J. M x) = sets N" for J by (intro sets_SUP sets) (auto ) from \I \ {}\ obtain i where "i\I" by auto have "Sup_measure' (M`I) X = (SUP P\{P. finite P \ P \ M`I}. sup_measure.F id P X)" using sets by (intro emeasure_Sup_measure') auto also have "Sup_measure' (M`I) = (SUP i\I. M i)" unfolding Sup_measure_def using \I \ {}\ sets sets(1)[THEN sets_eq_imp_space_eq] by (intro Sup_lexord1[where P="\x. _ = x"]) auto also have "(SUP P\{P. finite P \ P \ M`I}. sup_measure.F id P X) = (SUP J\{J. J \ {} \ finite J \ J \ I}. (SUP i\J. M i) X)" proof (intro SUP_eq) fix J assume "J \ {P. finite P \ P \ M`I}" then obtain J' where J': "J' \ I" "finite J'" and J: "J = M`J'" and "finite J" using finite_subset_image[of J M I] by auto show "\j\{J. J \ {} \ finite J \ J \ I}. sup_measure.F id J X \ (SUP i\j. M i) X" proof cases assume "J' = {}" with \i \ I\ show ?thesis by (auto simp add: J) next assume "J' \ {}" with J J' show ?thesis by (intro bexI[of _ "J'"]) (auto simp add: eq simp del: id_apply) qed next fix J assume J: "J \ {P. P \ {} \ finite P \ P \ I}" show "\J'\{J. finite J \ J \ M`I}. (SUP i\J. M i) X \ sup_measure.F id J' X" using J by (intro bexI[of _ "M`J"]) (auto simp add: eq simp del: id_apply) qed finally show ?thesis . qed lemma emeasure_SUP_chain: assumes sets: "\i. i \ A \ sets (M i) = sets N" "X \ sets N" assumes ch: "Complete_Partial_Order.chain (\) (M ` A)" and "A \ {}" shows "emeasure (SUP i\A. M i) X = (SUP i\A. emeasure (M i) X)" proof (subst emeasure_SUP[OF sets \A \ {}\]) show "(SUP J\{J. J \ {} \ finite J \ J \ A}. emeasure (Sup (M ` J)) X) = (SUP i\A. emeasure (M i) X)" proof (rule SUP_eq) fix J assume "J \ {J. J \ {} \ finite J \ J \ A}" then have J: "Complete_Partial_Order.chain (\) (M ` J)" "finite J" "J \ {}" and "J \ A" using ch[THEN chain_subset, of "M`J"] by auto with in_chain_finite[OF J(1)] obtain j where "j \ J" "(SUP j\J. M j) = M j" by auto with \J \ A\ show "\j\A. emeasure (Sup (M ` J)) X \ emeasure (M j) X" by auto next fix j assume "j\A" then show "\i\{J. J \ {} \ finite J \ J \ A}. emeasure (M j) X \ emeasure (Sup (M ` i)) X" by (intro bexI[of _ "{j}"]) auto qed qed subsubsection\<^marker>\tag unimportant\ \Supremum of a set of \\\-algebras\ lemma space_Sup_eq_UN: "space (Sup M) = (\x\M. space x)" unfolding Sup_measure_def apply (intro Sup_lexord[where P="\x. space x = _"]) apply (subst space_Sup_measure'2) apply auto [] apply (subst space_measure_of[OF UN_space_closed]) apply auto done lemma sets_Sup_eq: assumes *: "\m. m \ M \ space m = X" and "M \ {}" shows "sets (Sup M) = sigma_sets X (\x\M. sets x)" unfolding Sup_measure_def apply (rule Sup_lexord1) apply fact apply (simp add: assms) apply (rule Sup_lexord) subgoal premises that for a S unfolding that(3) that(2)[symmetric] using that(1) apply (subst sets_Sup_measure'2) apply (intro arg_cong2[where f=sigma_sets]) apply (auto simp: *) done apply (subst sets_measure_of[OF UN_space_closed]) apply (simp add: assms) done lemma in_sets_Sup: "(\m. m \ M \ space m = X) \ m \ M \ A \ sets m \ A \ sets (Sup M)" by (subst sets_Sup_eq[where X=X]) auto lemma Sup_lexord_rel: assumes "\i. i \ I \ k (A i) = k (B i)" "R (c (A ` {a \ I. k (B a) = (SUP x\I. k (B x))})) (c (B ` {a \ I. k (B a) = (SUP x\I. k (B x))}))" "R (s (A`I)) (s (B`I))" shows "R (Sup_lexord k c s (A`I)) (Sup_lexord k c s (B`I))" proof - have "A ` {a \ I. k (B a) = (SUP x\I. k (B x))} = {a \ A ` I. k a = (SUP x\I. k (B x))}" using assms(1) by auto moreover have "B ` {a \ I. k (B a) = (SUP x\I. k (B x))} = {a \ B ` I. k a = (SUP x\I. k (B x))}" by auto ultimately show ?thesis using assms by (auto simp: Sup_lexord_def Let_def image_comp) qed lemma sets_SUP_cong: assumes eq: "\i. i \ I \ sets (M i) = sets (N i)" shows "sets (SUP i\I. M i) = sets (SUP i\I. N i)" unfolding Sup_measure_def using eq eq[THEN sets_eq_imp_space_eq] apply (intro Sup_lexord_rel[where R="\x y. sets x = sets y"]) apply simp apply simp apply (simp add: sets_Sup_measure'2) apply (intro arg_cong2[where f="\x y. sets (sigma x y)"]) apply auto done lemma sets_Sup_in_sets: assumes "M \ {}" assumes "\m. m \ M \ space m = space N" assumes "\m. m \ M \ sets m \ sets N" shows "sets (Sup M) \ sets N" proof - have *: "\(space ` M) = space N" using assms by auto show ?thesis unfolding * using assms by (subst sets_Sup_eq[of M "space N"]) (auto intro!: sets.sigma_sets_subset) qed lemma measurable_Sup1: assumes m: "m \ M" and f: "f \ measurable m N" and const_space: "\m n. m \ M \ n \ M \ space m = space n" shows "f \ measurable (Sup M) N" proof - have "space (Sup M) = space m" using m by (auto simp add: space_Sup_eq_UN dest: const_space) then show ?thesis using m f unfolding measurable_def by (auto intro: in_sets_Sup[OF const_space]) qed lemma measurable_Sup2: assumes M: "M \ {}" assumes f: "\m. m \ M \ f \ measurable N m" and const_space: "\m n. m \ M \ n \ M \ space m = space n" shows "f \ measurable N (Sup M)" proof - from M obtain m where "m \ M" by auto have space_eq: "\n. n \ M \ space n = space m" by (intro const_space \m \ M\) have "f \ measurable N (sigma (\m\M. space m) (\m\M. sets m))" proof (rule measurable_measure_of) show "f \ space N \ \(space ` M)" using measurable_space[OF f] M by auto qed (auto intro: measurable_sets f dest: sets.sets_into_space) also have "measurable N (sigma (\m\M. space m) (\m\M. sets m)) = measurable N (Sup M)" apply (intro measurable_cong_sets refl) apply (subst sets_Sup_eq[OF space_eq M]) apply simp apply (subst sets_measure_of[OF UN_space_closed]) apply (simp add: space_eq M) done finally show ?thesis . qed lemma measurable_SUP2: "I \ {} \ (\i. i \ I \ f \ measurable N (M i)) \ (\i j. i \ I \ j \ I \ space (M i) = space (M j)) \ f \ measurable N (SUP i\I. M i)" by (auto intro!: measurable_Sup2) lemma sets_Sup_sigma: assumes [simp]: "M \ {}" and M: "\m. m \ M \ m \ Pow \" shows "sets (SUP m\M. sigma \ m) = sets (sigma \ (\M))" proof - { fix a m assume "a \ sigma_sets \ m" "m \ M" then have "a \ sigma_sets \ (\M)" by induction (auto intro: sigma_sets.intros(2-)) } then show "sets (SUP m\M. sigma \ m) = sets (sigma \ (\M))" apply (subst sets_Sup_eq[where X="\"]) apply (auto simp add: M) [] apply auto [] apply (simp add: space_measure_of_conv M Union_least) apply (rule sigma_sets_eqI) apply auto done qed lemma Sup_sigma: assumes [simp]: "M \ {}" and M: "\m. m \ M \ m \ Pow \" shows "(SUP m\M. sigma \ m) = (sigma \ (\M))" proof (intro antisym SUP_least) have *: "\M \ Pow \" using M by auto show "sigma \ (\M) \ (SUP m\M. sigma \ m)" proof (intro less_eq_measure.intros(3)) show "space (sigma \ (\M)) = space (SUP m\M. sigma \ m)" "sets (sigma \ (\M)) = sets (SUP m\M. sigma \ m)" using sets_Sup_sigma[OF assms] sets_Sup_sigma[OF assms, THEN sets_eq_imp_space_eq] by auto qed (simp add: emeasure_sigma le_fun_def) fix m assume "m \ M" then show "sigma \ m \ sigma \ (\M)" by (subst sigma_le_iff) (auto simp add: M *) qed lemma SUP_sigma_sigma: "M \ {} \ (\m. m \ M \ f m \ Pow \) \ (SUP m\M. sigma \ (f m)) = sigma \ (\m\M. f m)" using Sup_sigma[of "f`M" \] by (auto simp: image_comp) lemma sets_vimage_Sup_eq: assumes *: "M \ {}" "f \ X \ Y" "\m. m \ M \ space m = Y" shows "sets (vimage_algebra X f (Sup M)) = sets (SUP m \ M. vimage_algebra X f m)" (is "?IS = ?SI") proof show "?IS \ ?SI" apply (intro sets_image_in_sets measurable_Sup2) apply (simp add: space_Sup_eq_UN *) apply (simp add: *) apply (intro measurable_Sup1) apply (rule imageI) apply assumption apply (rule measurable_vimage_algebra1) apply (auto simp: *) done show "?SI \ ?IS" apply (intro sets_Sup_in_sets) apply (auto simp: *) [] apply (auto simp: *) [] apply (elim imageE) apply simp apply (rule sets_image_in_sets) apply simp apply (simp add: measurable_def) apply (simp add: * space_Sup_eq_UN sets_vimage_algebra2) apply (auto intro: in_sets_Sup[OF *(3)]) done qed lemma restrict_space_eq_vimage_algebra': "sets (restrict_space M \) = sets (vimage_algebra (\ \ space M) (\x. x) M)" proof - have *: "{A \ (\ \ space M) |A. A \ sets M} = {A \ \ |A. A \ sets M}" using sets.sets_into_space[of _ M] by blast show ?thesis unfolding restrict_space_def by (subst sets_measure_of) (auto simp add: image_subset_iff sets_vimage_algebra * dest: sets.sets_into_space intro!: arg_cong2[where f=sigma_sets]) qed lemma sigma_le_sets: assumes [simp]: "A \ Pow X" shows "sets (sigma X A) \ sets N \ X \ sets N \ A \ sets N" proof have "X \ sigma_sets X A" "A \ sigma_sets X A" by (auto intro: sigma_sets_top) moreover assume "sets (sigma X A) \ sets N" ultimately show "X \ sets N \ A \ sets N" by auto next assume *: "X \ sets N \ A \ sets N" { fix Y assume "Y \ sigma_sets X A" from this * have "Y \ sets N" by induction auto } then show "sets (sigma X A) \ sets N" by auto qed lemma measurable_iff_sets: "f \ measurable M N \ (f \ space M \ space N \ sets (vimage_algebra (space M) f N) \ sets M)" proof - have *: "{f -` A \ space M |A. A \ sets N} \ Pow (space M)" by auto show ?thesis unfolding measurable_def by (auto simp add: vimage_algebra_def sigma_le_sets[OF *]) qed lemma sets_vimage_algebra_space: "X \ sets (vimage_algebra X f M)" using sets.top[of "vimage_algebra X f M"] by simp lemma measurable_mono: assumes N: "sets N' \ sets N" "space N = space N'" assumes M: "sets M \ sets M'" "space M = space M'" shows "measurable M N \ measurable M' N'" unfolding measurable_def proof safe fix f A assume "f \ space M \ space N" "A \ sets N'" moreover assume "\y\sets N. f -` y \ space M \ sets M" note this[THEN bspec, of A] ultimately show "f -` A \ space M' \ sets M'" using assms by auto qed (insert N M, auto) lemma measurable_Sup_measurable: assumes f: "f \ space N \ A" shows "f \ measurable N (Sup {M. space M = A \ f \ measurable N M})" proof (rule measurable_Sup2) show "{M. space M = A \ f \ measurable N M} \ {}" using f unfolding ex_in_conv[symmetric] by (intro exI[of _ "sigma A {}"]) (auto intro!: measurable_measure_of) qed auto lemma (in sigma_algebra) sigma_sets_subset': assumes a: "a \ M" "\' \ M" shows "sigma_sets \' a \ M" proof show "x \ M" if x: "x \ sigma_sets \' a" for x using x by (induct rule: sigma_sets.induct) (insert a, auto) qed lemma in_sets_SUP: "i \ I \ (\i. i \ I \ space (M i) = Y) \ X \ sets (M i) \ X \ sets (SUP i\I. M i)" by (intro in_sets_Sup[where X=Y]) auto lemma measurable_SUP1: "i \ I \ f \ measurable (M i) N \ (\m n. m \ I \ n \ I \ space (M m) = space (M n)) \ f \ measurable (SUP i\I. M i) N" by (auto intro: measurable_Sup1) lemma sets_image_in_sets': assumes X: "X \ sets N" assumes f: "\A. A \ sets M \ f -` A \ X \ sets N" shows "sets (vimage_algebra X f M) \ sets N" unfolding sets_vimage_algebra by (rule sets.sigma_sets_subset') (auto intro!: measurable_sets X f) lemma mono_vimage_algebra: "sets M \ sets N \ sets (vimage_algebra X f M) \ sets (vimage_algebra X f N)" using sets.top[of "sigma X {f -` A \ X |A. A \ sets N}"] unfolding vimage_algebra_def apply (subst (asm) space_measure_of) apply auto [] apply (subst sigma_le_sets) apply auto done lemma mono_restrict_space: "sets M \ sets N \ sets (restrict_space M X) \ sets (restrict_space N X)" unfolding sets_restrict_space by (rule image_mono) lemma sets_eq_bot: "sets M = {{}} \ M = bot" apply safe apply (intro measure_eqI) apply auto done lemma sets_eq_bot2: "{{}} = sets M \ M = bot" using sets_eq_bot[of M] by blast lemma (in finite_measure) countable_support: "countable {x. measure M {x} \ 0}" proof cases assume "measure M (space M) = 0" with bounded_measure measure_le_0_iff have "{x. measure M {x} \ 0} = {}" by auto then show ?thesis by simp next let ?M = "measure M (space M)" and ?m = "\x. measure M {x}" assume "?M \ 0" then have *: "{x. ?m x \ 0} = (\n. {x. ?M / Suc n < ?m x})" using reals_Archimedean[of "?m x / ?M" for x] by (auto simp: field_simps not_le[symmetric] divide_le_0_iff measure_le_0_iff) have **: "\n. finite {x. ?M / Suc n < ?m x}" proof (rule ccontr) fix n assume "infinite {x. ?M / Suc n < ?m x}" (is "infinite ?X") then obtain X where "finite X" "card X = Suc (Suc n)" "X \ ?X" by (metis infinite_arbitrarily_large) from this(3) have *: "\x. x \ X \ ?M / Suc n \ ?m x" by auto { fix x assume "x \ X" from \?M \ 0\ *[OF this] have "?m x \ 0" by (auto simp: field_simps measure_le_0_iff) then have "{x} \ sets M" by (auto dest: measure_notin_sets) } note singleton_sets = this have "?M < (\x\X. ?M / Suc n)" using \?M \ 0\ by (simp add: \card X = Suc (Suc n)\ field_simps less_le) also have "\ \ (\x\X. ?m x)" by (rule sum_mono) fact also have "\ = measure M (\x\X. {x})" using singleton_sets \finite X\ by (intro finite_measure_finite_Union[symmetric]) (auto simp: disjoint_family_on_def) finally have "?M < measure M (\x\X. {x})" . moreover have "measure M (\x\X. {x}) \ ?M" using singleton_sets[THEN sets.sets_into_space] by (intro finite_measure_mono) auto ultimately show False by simp qed show ?thesis unfolding * by (intro countable_UN countableI_type countable_finite[OF **]) qed end diff --git a/src/HOL/Analysis/Polytope.thy b/src/HOL/Analysis/Polytope.thy --- a/src/HOL/Analysis/Polytope.thy +++ b/src/HOL/Analysis/Polytope.thy @@ -1,3999 +1,3998 @@ section \Faces, Extreme Points, Polytopes, Polyhedra etc\ text\Ported from HOL Light by L C Paulson\ theory Polytope imports Cartesian_Euclidean_Space Path_Connected begin subsection \Faces of a (usually convex) set\ definition\<^marker>\tag important\ face_of :: "['a::real_vector set, 'a set] \ bool" (infixr "(face'_of)" 50) where "T face_of S \ T \ S \ convex T \ (\a \ S. \b \ S. \x \ T. x \ open_segment a b \ a \ T \ b \ T)" lemma face_ofD: "\T face_of S; x \ open_segment a b; a \ S; b \ S; x \ T\ \ a \ T \ b \ T" unfolding face_of_def by blast lemma face_of_translation_eq [simp]: "((+) a ` T face_of (+) a ` S) \ T face_of S" proof - have *: "\a T S. T face_of S \ ((+) a ` T face_of (+) a ` S)" apply (simp add: face_of_def Ball_def, clarify) by (meson imageI open_segment_translation_eq) show ?thesis apply (rule iffI) apply (force simp: image_comp o_def dest: * [where a = "-a"]) apply (blast intro: *) done qed lemma face_of_linear_image: assumes "linear f" "inj f" shows "(f ` c face_of f ` S) \ c face_of S" by (simp add: face_of_def inj_image_subset_iff inj_image_mem_iff open_segment_linear_image assms) lemma face_of_refl: "convex S \ S face_of S" by (auto simp: face_of_def) lemma face_of_refl_eq: "S face_of S \ convex S" by (auto simp: face_of_def) lemma empty_face_of [iff]: "{} face_of S" by (simp add: face_of_def) lemma face_of_empty [simp]: "S face_of {} \ S = {}" by (meson empty_face_of face_of_def subset_empty) lemma face_of_trans [trans]: "\S face_of T; T face_of u\ \ S face_of u" unfolding face_of_def by (safe; blast) lemma face_of_face: "T face_of S \ (f face_of T \ f face_of S \ f \ T)" unfolding face_of_def by (safe; blast) lemma face_of_subset: "\F face_of S; F \ T; T \ S\ \ F face_of T" unfolding face_of_def by (safe; blast) lemma face_of_slice: "\F face_of S; convex T\ \ (F \ T) face_of (S \ T)" unfolding face_of_def by (blast intro: convex_Int) lemma face_of_Int: "\t1 face_of S; t2 face_of S\ \ (t1 \ t2) face_of S" unfolding face_of_def by (blast intro: convex_Int) lemma face_of_Inter: "\A \ {}; \T. T \ A \ T face_of S\ \ (\ A) face_of S" unfolding face_of_def by (blast intro: convex_Inter) lemma face_of_Int_Int: "\F face_of T; F' face_of t'\ \ (F \ F') face_of (T \ t')" unfolding face_of_def by (blast intro: convex_Int) lemma face_of_imp_subset: "T face_of S \ T \ S" unfolding face_of_def by blast proposition face_of_imp_eq_affine_Int: fixes S :: "'a::euclidean_space set" assumes S: "convex S" and T: "T face_of S" shows "T = (affine hull T) \ S" proof - have "convex T" using T by (simp add: face_of_def) have *: False if x: "x \ affine hull T" and "x \ S" "x \ T" and y: "y \ rel_interior T" for x y proof - obtain e where "e>0" and e: "cball y e \ affine hull T \ T" using y by (auto simp: rel_interior_cball) have "y \ x" "y \ S" "y \ T" using face_of_imp_subset rel_interior_subset T that by blast+ then have zne: "\u. \u \ {0<..<1}; (1 - u) *\<^sub>R y + u *\<^sub>R x \ T\ \ False" using \x \ S\ \x \ T\ \T face_of S\ unfolding face_of_def apply clarify apply (drule_tac x=x in bspec, assumption) apply (drule_tac x=y in bspec, assumption) apply (subst (asm) open_segment_commute) apply (force simp: open_segment_image_interval image_def) done have in01: "min (1/2) (e / norm (x - y)) \ {0<..<1}" using \y \ x\ \e > 0\ by simp show ?thesis apply (rule zne [OF in01]) apply (rule e [THEN subsetD]) apply (rule IntI) using \y \ x\ \e > 0\ apply (simp add: cball_def dist_norm algebra_simps) apply (simp add: Real_Vector_Spaces.scaleR_diff_right [symmetric] norm_minus_commute min_mult_distrib_right) apply (rule mem_affine [OF affine_affine_hull _ x]) using \y \ T\ apply (auto simp: hull_inc) done qed show ?thesis apply (rule subset_antisym) using assms apply (simp add: hull_subset face_of_imp_subset) apply (cases "T={}", simp) apply (force simp: rel_interior_eq_empty [symmetric] \convex T\ intro: *) done qed lemma face_of_imp_closed: fixes S :: "'a::euclidean_space set" assumes "convex S" "closed S" "T face_of S" shows "closed T" by (metis affine_affine_hull affine_closed closed_Int face_of_imp_eq_affine_Int assms) lemma face_of_Int_supporting_hyperplane_le_strong: assumes "convex(S \ {x. a \ x = b})" and aleb: "\x. x \ S \ a \ x \ b" shows "(S \ {x. a \ x = b}) face_of S" proof - have *: "a \ u = a \ x" if "x \ open_segment u v" "u \ S" "v \ S" and b: "b = a \ x" for u v x proof (rule antisym) show "a \ u \ a \ x" using aleb \u \ S\ \b = a \ x\ by blast next obtain \ where "b = a \ ((1 - \) *\<^sub>R u + \ *\<^sub>R v)" "0 < \" "\ < 1" using \b = a \ x\ \x \ open_segment u v\ in_segment by (auto simp: open_segment_image_interval split: if_split_asm) then have "b + \ * (a \ u) \ a \ u + \ * b" using aleb [OF \v \ S\] by (simp add: algebra_simps) then have "(1 - \) * b \ (1 - \) * (a \ u)" by (simp add: algebra_simps) then have "b \ a \ u" using \\ < 1\ by auto with b show "a \ x \ a \ u" by simp qed show ?thesis apply (simp add: face_of_def assms) using "*" open_segment_commute by blast qed lemma face_of_Int_supporting_hyperplane_ge_strong: "\convex(S \ {x. a \ x = b}); \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) face_of S" using face_of_Int_supporting_hyperplane_le_strong [of S "-a" "-b"] by simp lemma face_of_Int_supporting_hyperplane_le: "\convex S; \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) face_of S" by (simp add: convex_Int convex_hyperplane face_of_Int_supporting_hyperplane_le_strong) lemma face_of_Int_supporting_hyperplane_ge: "\convex S; \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) face_of S" by (simp add: convex_Int convex_hyperplane face_of_Int_supporting_hyperplane_ge_strong) lemma face_of_imp_convex: "T face_of S \ convex T" using face_of_def by blast lemma face_of_imp_compact: fixes S :: "'a::euclidean_space set" shows "\convex S; compact S; T face_of S\ \ compact T" by (meson bounded_subset compact_eq_bounded_closed face_of_imp_closed face_of_imp_subset) lemma face_of_Int_subface: "\A \ B face_of A; A \ B face_of B; C face_of A; D face_of B\ \ (C \ D) face_of C \ (C \ D) face_of D" by (meson face_of_Int_Int face_of_face inf_le1 inf_le2) lemma subset_of_face_of: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "u \ S" "T \ (rel_interior u) \ {}" shows "u \ T" proof fix c assume "c \ u" obtain b where "b \ T" "b \ rel_interior u" using assms by auto then obtain e where "e>0" "b \ u" and e: "cball b e \ affine hull u \ u" by (auto simp: rel_interior_cball) show "c \ T" proof (cases "b=c") case True with \b \ T\ show ?thesis by blast next case False define d where "d = b + (e / norm(b - c)) *\<^sub>R (b - c)" have "d \ cball b e \ affine hull u" using \e > 0\ \b \ u\ \c \ u\ by (simp add: d_def dist_norm hull_inc mem_affine_3_minus False) with e have "d \ u" by blast have nbc: "norm (b - c) + e > 0" using \e > 0\ by (metis add.commute le_less_trans less_add_same_cancel2 norm_ge_zero) then have [simp]: "d \ c" using False scaleR_cancel_left [of "1 + (e / norm (b - c))" b c] by (simp add: algebra_simps d_def) (simp add: field_split_simps) have [simp]: "((e - e * e / (e + norm (b - c))) / norm (b - c)) = (e / (e + norm (b - c)))" using False nbc by (simp add: divide_simps) (simp add: algebra_simps) have "b \ open_segment d c" apply (simp add: open_segment_image_interval) apply (simp add: d_def algebra_simps image_def) apply (rule_tac x="e / (e + norm (b - c))" in bexI) using False nbc \0 < e\ apply (auto simp: algebra_simps) done then have "d \ T \ c \ T" apply (rule face_ofD [OF \T face_of S\]) using \d \ u\ \c \ u\ \u \ S\ \b \ T\ apply auto done then show ?thesis .. qed qed lemma face_of_eq: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "u face_of S" "(rel_interior T) \ (rel_interior u) \ {}" shows "T = u" apply (rule subset_antisym) apply (metis assms disjoint_iff_not_equal face_of_imp_subset rel_interior_subset subsetCE subset_of_face_of) by (metis assms disjoint_iff_not_equal face_of_imp_subset rel_interior_subset subset_iff subset_of_face_of) lemma face_of_disjoint_rel_interior: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "T \ S" shows "T \ rel_interior S = {}" by (meson assms subset_of_face_of face_of_imp_subset order_refl subset_antisym) lemma face_of_disjoint_interior: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "T \ S" shows "T \ interior S = {}" proof - have "T \ interior S \ rel_interior S" by (meson inf_sup_ord(2) interior_subset_rel_interior order.trans) thus ?thesis by (metis (no_types) Int_greatest assms face_of_disjoint_rel_interior inf_sup_ord(1) subset_empty) qed lemma face_of_subset_rel_boundary: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "T \ S" shows "T \ (S - rel_interior S)" by (meson DiffI assms disjoint_iff_not_equal face_of_disjoint_rel_interior face_of_imp_subset rev_subsetD subsetI) lemma face_of_subset_rel_frontier: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "T \ S" shows "T \ rel_frontier S" using assms closure_subset face_of_disjoint_rel_interior face_of_imp_subset rel_frontier_def by fastforce lemma face_of_aff_dim_lt: fixes S :: "'a::euclidean_space set" assumes "convex S" "T face_of S" "T \ S" shows "aff_dim T < aff_dim S" proof - have "aff_dim T \ aff_dim S" by (simp add: face_of_imp_subset aff_dim_subset assms) moreover have "aff_dim T \ aff_dim S" proof (cases "T = {}") case True then show ?thesis by (metis aff_dim_empty \T \ S\) next case False then show ?thesis by (metis Set.set_insert assms convex_rel_frontier_aff_dim dual_order.irrefl face_of_imp_convex face_of_subset_rel_frontier insert_not_empty subsetI) qed ultimately show ?thesis by simp qed lemma subset_of_face_of_affine_hull: fixes S :: "'a::euclidean_space set" assumes T: "T face_of S" and "convex S" "U \ S" and dis: "\ disjnt (affine hull T) (rel_interior U)" shows "U \ T" apply (rule subset_of_face_of [OF T \U \ S\]) using face_of_imp_eq_affine_Int [OF \convex S\ T] using rel_interior_subset [of U] dis using \U \ S\ disjnt_def by fastforce lemma affine_hull_face_of_disjoint_rel_interior: fixes S :: "'a::euclidean_space set" assumes "convex S" "F face_of S" "F \ S" shows "affine hull F \ rel_interior S = {}" by (metis assms disjnt_def face_of_imp_subset order_refl subset_antisym subset_of_face_of_affine_hull) lemma affine_diff_divide: assumes "affine S" "k \ 0" "k \ 1" and xy: "x \ S" "y /\<^sub>R (1 - k) \ S" shows "(x - y) /\<^sub>R k \ S" proof - have "inverse(k) *\<^sub>R (x - y) = (1 - inverse k) *\<^sub>R inverse(1 - k) *\<^sub>R y + inverse(k) *\<^sub>R x" using assms by (simp add: algebra_simps) (simp add: scaleR_left_distrib [symmetric] field_split_simps) then show ?thesis using \affine S\ xy by (auto simp: affine_alt) qed proposition face_of_convex_hulls: assumes S: "finite S" "T \ S" and disj: "affine hull T \ convex hull (S - T) = {}" shows "(convex hull T) face_of (convex hull S)" proof - have fin: "finite T" "finite (S - T)" using assms by (auto simp: finite_subset) have *: "x \ convex hull T" if x: "x \ convex hull S" and y: "y \ convex hull S" and w: "w \ convex hull T" "w \ open_segment x y" for x y w proof - have waff: "w \ affine hull T" using convex_hull_subset_affine_hull w by blast obtain a b where a: "\i. i \ S \ 0 \ a i" and asum: "sum a S = 1" and aeqx: "(\i\S. a i *\<^sub>R i) = x" and b: "\i. i \ S \ 0 \ b i" and bsum: "sum b S = 1" and beqy: "(\i\S. b i *\<^sub>R i) = y" using x y by (auto simp: assms convex_hull_finite) obtain u where "(1 - u) *\<^sub>R x + u *\<^sub>R y \ convex hull T" "x \ y" and weq: "w = (1 - u) *\<^sub>R x + u *\<^sub>R y" and u01: "0 < u" "u < 1" using w by (auto simp: open_segment_image_interval split: if_split_asm) define c where "c i = (1 - u) * a i + u * b i" for i have cge0: "\i. i \ S \ 0 \ c i" using a b u01 by (simp add: c_def) have sumc1: "sum c S = 1" by (simp add: c_def sum.distrib sum_distrib_left [symmetric] asum bsum) have sumci_xy: "(\i\S. c i *\<^sub>R i) = (1 - u) *\<^sub>R x + u *\<^sub>R y" apply (simp add: c_def sum.distrib scaleR_left_distrib) by (simp only: scaleR_scaleR [symmetric] Real_Vector_Spaces.scaleR_right.sum [symmetric] aeqx beqy) show ?thesis proof (cases "sum c (S - T) = 0") case True have ci0: "\i. i \ (S - T) \ c i = 0" using True cge0 fin(2) sum_nonneg_eq_0_iff by auto have a0: "a i = 0" if "i \ (S - T)" for i using ci0 [OF that] u01 a [of i] b [of i] that by (simp add: c_def Groups.ordered_comm_monoid_add_class.add_nonneg_eq_0_iff) have [simp]: "sum a T = 1" using assms by (metis sum.mono_neutral_cong_right a0 asum) show ?thesis apply (simp add: convex_hull_finite \finite T\) apply (rule_tac x=a in exI) using a0 assms apply (auto simp: cge0 a aeqx [symmetric] sum.mono_neutral_right) done next case False define k where "k = sum c (S - T)" have "k > 0" using False unfolding k_def by (metis DiffD1 antisym_conv cge0 sum_nonneg not_less) have weq_sumsum: "w = sum (\x. c x *\<^sub>R x) T + sum (\x. c x *\<^sub>R x) (S - T)" by (metis (no_types) add.commute S(1) S(2) sum.subset_diff sumci_xy weq) show ?thesis proof (cases "k = 1") case True then have "sum c T = 0" by (simp add: S k_def sum_diff sumc1) then have [simp]: "sum c (S - T) = 1" by (simp add: S sum_diff sumc1) have ci0: "\i. i \ T \ c i = 0" by (meson \finite T\ \sum c T = 0\ \T \ S\ cge0 sum_nonneg_eq_0_iff subsetCE) then have [simp]: "(\i\S-T. c i *\<^sub>R i) = w" by (simp add: weq_sumsum) have "w \ convex hull (S - T)" apply (simp add: convex_hull_finite fin) apply (rule_tac x=c in exI) apply (auto simp: cge0 weq True k_def) done then show ?thesis using disj waff by blast next case False then have sumcf: "sum c T = 1 - k" by (simp add: S k_def sum_diff sumc1) have "(\i\T. c i *\<^sub>R i) /\<^sub>R (1 - k) \ convex hull T" apply (simp add: convex_hull_finite fin) apply (rule_tac x="\i. inverse (1-k) * c i" in exI) apply auto apply (metis sumcf cge0 inverse_nonnegative_iff_nonnegative mult_nonneg_nonneg S(2) sum_nonneg subsetCE) apply (metis False mult.commute right_inverse right_minus_eq sum_distrib_left sumcf) by (metis (mono_tags, lifting) scaleR_right.sum scaleR_scaleR sum.cong) with \0 < k\ have "inverse(k) *\<^sub>R (w - sum (\i. c i *\<^sub>R i) T) \ affine hull T" by (simp add: affine_diff_divide [OF affine_affine_hull] False waff convex_hull_subset_affine_hull [THEN subsetD]) moreover have "inverse(k) *\<^sub>R (w - sum (\x. c x *\<^sub>R x) T) \ convex hull (S - T)" apply (simp add: weq_sumsum convex_hull_finite fin) apply (rule_tac x="\i. inverse k * c i" in exI) using \k > 0\ cge0 apply (auto simp: scaleR_right.sum sum_distrib_left [symmetric] k_def [symmetric]) done ultimately show ?thesis using disj by blast qed qed qed have [simp]: "convex hull T \ convex hull S" by (simp add: \T \ S\ hull_mono) show ?thesis using open_segment_commute by (auto simp: face_of_def intro: *) qed proposition face_of_convex_hull_insert: "\finite S; a \ affine hull S; T face_of convex hull S\ \ T face_of convex hull insert a S" apply (rule face_of_trans, blast) apply (rule face_of_convex_hulls; force simp: insert_Diff_if) done proposition face_of_affine_trivial: assumes "affine S" "T face_of S" shows "T = {} \ T = S" proof (rule ccontr, clarsimp) assume "T \ {}" "T \ S" then obtain a where "a \ T" by auto then have "a \ S" using \T face_of S\ face_of_imp_subset by blast have "S \ T" proof fix b assume "b \ S" show "b \ T" proof (cases "a = b") case True with \a \ T\ show ?thesis by auto next case False then have "a \ open_segment (2 *\<^sub>R a - b) b" apply (auto simp: open_segment_def closed_segment_def) apply (rule_tac x="1/2" in exI) apply (simp add: algebra_simps) by (simp add: scaleR_2) moreover have "2 *\<^sub>R a - b \ S" by (rule mem_affine [OF \affine S\ \a \ S\ \b \ S\, of 2 "-1", simplified]) moreover note \b \ S\ \a \ T\ ultimately show ?thesis by (rule face_ofD [OF \T face_of S\, THEN conjunct2]) qed qed then show False using \T \ S\ \T face_of S\ face_of_imp_subset by blast qed lemma face_of_affine_eq: "affine S \ (T face_of S \ T = {} \ T = S)" using affine_imp_convex face_of_affine_trivial face_of_refl by auto proposition Inter_faces_finite_altbound: fixes T :: "'a::euclidean_space set set" assumes cfaI: "\c. c \ T \ c face_of S" shows "\F'. finite F' \ F' \ T \ card F' \ DIM('a) + 2 \ \F' = \T" proof (cases "\F'. finite F' \ F' \ T \ card F' \ DIM('a) + 2 \ (\c. c \ T \ c \ (\F') \ (\F'))") case True then obtain c where c: "\F'. \finite F'; F' \ T; card F' \ DIM('a) + 2\ \ c F' \ T \ c F' \ (\F') \ (\F')" by metis define d where "d = rec_nat {c{}} (\n r. insert (c r) r)" have [simp]: "d 0 = {c {}}" by (simp add: d_def) have dSuc [simp]: "\n. d (Suc n) = insert (c (d n)) (d n)" by (simp add: d_def) have dn_notempty: "d n \ {}" for n by (induction n) auto have dn_le_Suc: "d n \ T \ finite(d n) \ card(d n) \ Suc n" if "n \ DIM('a) + 2" for n using that proof (induction n) case 0 then show ?case by (simp add: c) next case (Suc n) then show ?case by (auto simp: c card_insert_if) qed have aff_dim_le: "aff_dim(\(d n)) \ DIM('a) - int n" if "n \ DIM('a) + 2" for n using that proof (induction n) case 0 then show ?case by (simp add: aff_dim_le_DIM) next case (Suc n) have fs: "\(d (Suc n)) face_of S" by (meson Suc.prems cfaI dn_le_Suc dn_notempty face_of_Inter subsetCE) have condn: "convex (\(d n))" using Suc.prems nat_le_linear not_less_eq_eq by (blast intro: face_of_imp_convex cfaI convex_Inter dest: dn_le_Suc) have fdn: "\(d (Suc n)) face_of \(d n)" by (metis (no_types, lifting) Inter_anti_mono Suc.prems dSuc cfaI dn_le_Suc dn_notempty face_of_Inter face_of_imp_subset face_of_subset subset_iff subset_insertI) have ne: "\(d (Suc n)) \ \(d n)" by (metis (no_types, lifting) Suc.prems Suc_leD c complete_lattice_class.Inf_insert dSuc dn_le_Suc less_irrefl order.trans) have *: "\m::int. \d. \d'::int. d < d' \ d' \ m - n \ d \ m - of_nat(n+1)" by arith have "aff_dim (\(d (Suc n))) < aff_dim (\(d n))" by (rule face_of_aff_dim_lt [OF condn fdn ne]) moreover have "aff_dim (\(d n)) \ int (DIM('a)) - int n" using Suc by auto ultimately have "aff_dim (\(d (Suc n))) \ int (DIM('a)) - (n+1)" by arith then show ?case by linarith qed have "aff_dim (\(d (DIM('a) + 2))) \ -2" using aff_dim_le [OF order_refl] by simp with aff_dim_geq [of "\(d (DIM('a) + 2))"] show ?thesis using order.trans by fastforce next case False then show ?thesis apply simp apply (erule ex_forward) by blast qed lemma faces_of_translation: "{F. F face_of image (\x. a + x) S} = image (image (\x. a + x)) {F. F face_of S}" apply (rule subset_antisym, clarify) apply (auto simp: image_iff) apply (metis face_of_imp_subset face_of_translation_eq subset_imageE) done proposition face_of_Times: assumes "F face_of S" and "F' face_of S'" shows "(F \ F') face_of (S \ S')" proof - have "F \ F' \ S \ S'" using assms [unfolded face_of_def] by blast moreover have "convex (F \ F')" using assms [unfolded face_of_def] by (blast intro: convex_Times) moreover have "a \ F \ a' \ F' \ b \ F \ b' \ F'" if "a \ S" "b \ S" "a' \ S'" "b' \ S'" "x \ F \ F'" "x \ open_segment (a,a') (b,b')" for a b a' b' x proof (cases "b=a \ b'=a'") case True with that show ?thesis using assms by (force simp: in_segment dest: face_ofD) next case False with assms [unfolded face_of_def] that show ?thesis by (blast dest!: open_segment_PairD) qed ultimately show ?thesis unfolding face_of_def by blast qed corollary face_of_Times_decomp: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" shows "c face_of (S \ S') \ (\F F'. F face_of S \ F' face_of S' \ c = F \ F')" (is "?lhs = ?rhs") proof assume c: ?lhs show ?rhs proof (cases "c = {}") case True then show ?thesis by auto next case False have 1: "fst ` c \ S" "snd ` c \ S'" using c face_of_imp_subset by fastforce+ have "convex c" using c by (metis face_of_imp_convex) have conv: "convex (fst ` c)" "convex (snd ` c)" by (simp_all add: \convex c\ convex_linear_image linear_fst linear_snd) have fstab: "a \ fst ` c \ b \ fst ` c" if "a \ S" "b \ S" "x \ open_segment a b" "(x,x') \ c" for a b x x' proof - have *: "(x,x') \ open_segment (a,x') (b,x')" using that by (auto simp: in_segment) show ?thesis using face_ofD [OF c *] that face_of_imp_subset [OF c] by force qed have fst: "fst ` c face_of S" by (force simp: face_of_def 1 conv fstab) have sndab: "a' \ snd ` c \ b' \ snd ` c" if "a' \ S'" "b' \ S'" "x' \ open_segment a' b'" "(x,x') \ c" for a' b' x x' proof - have *: "(x,x') \ open_segment (x,a') (x,b')" using that by (auto simp: in_segment) show ?thesis using face_ofD [OF c *] that face_of_imp_subset [OF c] by force qed have snd: "snd ` c face_of S'" by (force simp: face_of_def 1 conv sndab) have cc: "rel_interior c \ rel_interior (fst ` c) \ rel_interior (snd ` c)" by (force simp: face_of_Times rel_interior_Times conv fst snd \convex c\ linear_fst linear_snd rel_interior_convex_linear_image [symmetric]) have "c = fst ` c \ snd ` c" apply (rule face_of_eq [OF c]) apply (simp_all add: face_of_Times rel_interior_Times conv fst snd) using False rel_interior_eq_empty \convex c\ cc apply blast done with fst snd show ?thesis by metis qed next assume ?rhs with face_of_Times show ?lhs by auto qed lemma face_of_Times_eq: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" shows "(F \ F') face_of (S \ S') \ F = {} \ F' = {} \ F face_of S \ F' face_of S'" by (auto simp: face_of_Times_decomp times_eq_iff) lemma hyperplane_face_of_halfspace_le: "{x. a \ x = b} face_of {x. a \ x \ b}" proof - have "{x. a \ x \ b} \ {x. a \ x = b} = {x. a \ x = b}" by auto with face_of_Int_supporting_hyperplane_le [OF convex_halfspace_le [of a b], of a b] show ?thesis by auto qed lemma hyperplane_face_of_halfspace_ge: "{x. a \ x = b} face_of {x. a \ x \ b}" proof - have "{x. a \ x \ b} \ {x. a \ x = b} = {x. a \ x = b}" by auto with face_of_Int_supporting_hyperplane_ge [OF convex_halfspace_ge [of b a], of b a] show ?thesis by auto qed lemma face_of_halfspace_le: fixes a :: "'n::euclidean_space" shows "F face_of {x. a \ x \ b} \ F = {} \ F = {x. a \ x = b} \ F = {x. a \ x \ b}" (is "?lhs = ?rhs") proof (cases "a = 0") case True then show ?thesis using face_of_affine_eq affine_UNIV by auto next case False then have ine: "interior {x. a \ x \ b} \ {}" using halfspace_eq_empty_lt interior_halfspace_le by blast show ?thesis proof assume L: ?lhs have "F \ {x. a \ x \ b} \ F face_of {x. a \ x = b}" using False apply (simp add: frontier_halfspace_le [symmetric] rel_frontier_nonempty_interior [OF ine, symmetric]) apply (rule face_of_subset [OF L]) apply (simp add: face_of_subset_rel_frontier [OF L]) apply (force simp: rel_frontier_def closed_halfspace_le) done with L show ?rhs using affine_hyperplane face_of_affine_eq by blast next assume ?rhs then show ?lhs by (metis convex_halfspace_le empty_face_of face_of_refl hyperplane_face_of_halfspace_le) qed qed lemma face_of_halfspace_ge: fixes a :: "'n::euclidean_space" shows "F face_of {x. a \ x \ b} \ F = {} \ F = {x. a \ x = b} \ F = {x. a \ x \ b}" using face_of_halfspace_le [of F "-a" "-b"] by simp subsection\Exposed faces\ text\That is, faces that are intersection with supporting hyperplane\ definition\<^marker>\tag important\ exposed_face_of :: "['a::euclidean_space set, 'a set] \ bool" (infixr "(exposed'_face'_of)" 50) where "T exposed_face_of S \ T face_of S \ (\a b. S \ {x. a \ x \ b} \ T = S \ {x. a \ x = b})" lemma empty_exposed_face_of [iff]: "{} exposed_face_of S" apply (simp add: exposed_face_of_def) apply (rule_tac x=0 in exI) apply (rule_tac x=1 in exI, force) done lemma exposed_face_of_refl_eq [simp]: "S exposed_face_of S \ convex S" apply (simp add: exposed_face_of_def face_of_refl_eq, auto) apply (rule_tac x=0 in exI)+ apply force done lemma exposed_face_of_refl: "convex S \ S exposed_face_of S" by simp lemma exposed_face_of: "T exposed_face_of S \ T face_of S \ (T = {} \ T = S \ (\a b. a \ 0 \ S \ {x. a \ x \ b} \ T = S \ {x. a \ x = b}))" proof (cases "T = {}") case True then show ?thesis by simp next case False show ?thesis proof (cases "T = S") case True then show ?thesis by (simp add: face_of_refl_eq) next case False with \T \ {}\ show ?thesis apply (auto simp: exposed_face_of_def) apply (metis inner_zero_left) done qed qed lemma exposed_face_of_Int_supporting_hyperplane_le: "\convex S; \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) exposed_face_of S" by (force simp: exposed_face_of_def face_of_Int_supporting_hyperplane_le) lemma exposed_face_of_Int_supporting_hyperplane_ge: "\convex S; \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) exposed_face_of S" using exposed_face_of_Int_supporting_hyperplane_le [of S "-a" "-b"] by simp proposition exposed_face_of_Int: assumes "T exposed_face_of S" and "u exposed_face_of S" shows "(T \ u) exposed_face_of S" proof - obtain a b where T: "S \ {x. a \ x = b} face_of S" and S: "S \ {x. a \ x \ b}" and teq: "T = S \ {x. a \ x = b}" using assms by (auto simp: exposed_face_of_def) obtain a' b' where u: "S \ {x. a' \ x = b'} face_of S" and s': "S \ {x. a' \ x \ b'}" and ueq: "u = S \ {x. a' \ x = b'}" using assms by (auto simp: exposed_face_of_def) have tu: "T \ u face_of S" using T teq u ueq by (simp add: face_of_Int) have ss: "S \ {x. (a + a') \ x \ b + b'}" using S s' by (force simp: inner_left_distrib) show ?thesis apply (simp add: exposed_face_of_def tu) apply (rule_tac x="a+a'" in exI) apply (rule_tac x="b+b'" in exI) using S s' apply (fastforce simp: ss inner_left_distrib teq ueq) done qed proposition exposed_face_of_Inter: fixes P :: "'a::euclidean_space set set" assumes "P \ {}" and "\T. T \ P \ T exposed_face_of S" shows "\P exposed_face_of S" proof - obtain Q where "finite Q" and QsubP: "Q \ P" "card Q \ DIM('a) + 2" and IntQ: "\Q = \P" using Inter_faces_finite_altbound [of P S] assms [unfolded exposed_face_of] by force show ?thesis proof (cases "Q = {}") case True then show ?thesis by (metis IntQ Inter_UNIV_conv(2) assms(1) assms(2) ex_in_conv) next case False have "Q \ {T. T exposed_face_of S}" using QsubP assms by blast moreover have "Q \ {T. T exposed_face_of S} \ \Q exposed_face_of S" using \finite Q\ False apply (induction Q rule: finite_induct) using exposed_face_of_Int apply fastforce+ done ultimately show ?thesis by (simp add: IntQ) qed qed proposition exposed_face_of_sums: assumes "convex S" and "convex T" and "F exposed_face_of {x + y | x y. x \ S \ y \ T}" (is "F exposed_face_of ?ST") obtains k l where "k exposed_face_of S" "l exposed_face_of T" "F = {x + y | x y. x \ k \ y \ l}" proof (cases "F = {}") case True then show ?thesis using that by blast next case False show ?thesis proof (cases "F = ?ST") case True then show ?thesis using assms exposed_face_of_refl_eq that by blast next case False obtain p where "p \ F" using \F \ {}\ by blast moreover obtain u z where T: "?ST \ {x. u \ x = z} face_of ?ST" and S: "?ST \ {x. u \ x \ z}" and feq: "F = ?ST \ {x. u \ x = z}" using assms by (auto simp: exposed_face_of_def) ultimately obtain a0 b0 where p: "p = a0 + b0" and "a0 \ S" "b0 \ T" and z: "u \ p = z" by auto have lez: "u \ (x + y) \ z" if "x \ S" "y \ T" for x y using S that by auto have sef: "S \ {x. u \ x = u \ a0} exposed_face_of S" apply (rule exposed_face_of_Int_supporting_hyperplane_le [OF \convex S\]) apply (metis p z add_le_cancel_right inner_right_distrib lez [OF _ \b0 \ T\]) done have tef: "T \ {x. u \ x = u \ b0} exposed_face_of T" apply (rule exposed_face_of_Int_supporting_hyperplane_le [OF \convex T\]) apply (metis p z add.commute add_le_cancel_right inner_right_distrib lez [OF \a0 \ S\]) done have "{x + y |x y. x \ S \ u \ x = u \ a0 \ y \ T \ u \ y = u \ b0} \ F" by (auto simp: feq) (metis inner_right_distrib p z) moreover have "F \ {x + y |x y. x \ S \ u \ x = u \ a0 \ y \ T \ u \ y = u \ b0}" apply (auto simp: feq) apply (rename_tac x y) apply (rule_tac x=x in exI) apply (rule_tac x=y in exI, simp) using z p \a0 \ S\ \b0 \ T\ apply clarify apply (simp add: inner_right_distrib) apply (metis add_le_cancel_right antisym lez [unfolded inner_right_distrib] add.commute) done ultimately have "F = {x + y |x y. x \ S \ {x. u \ x = u \ a0} \ y \ T \ {x. u \ x = u \ b0}}" by blast then show ?thesis by (rule that [OF sef tef]) qed qed proposition exposed_face_of_parallel: "T exposed_face_of S \ T face_of S \ (\a b. S \ {x. a \ x \ b} \ T = S \ {x. a \ x = b} \ (T \ {} \ T \ S \ a \ 0) \ (T \ S \ (\w \ affine hull S. (w + a) \ affine hull S)))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs proof (clarsimp simp: exposed_face_of_def) fix a b assume faceS: "S \ {x. a \ x = b} face_of S" and Ssub: "S \ {x. a \ x \ b}" show "\c d. S \ {x. c \ x \ d} \ S \ {x. a \ x = b} = S \ {x. c \ x = d} \ (S \ {x. a \ x = b} \ {} \ S \ {x. a \ x = b} \ S \ c \ 0) \ (S \ {x. a \ x = b} \ S \ (\w \ affine hull S. w + c \ affine hull S))" proof (cases "affine hull S \ {x. -a \ x \ -b} = {} \ affine hull S \ {x. - a \ x \ - b}") case True then show ?thesis proof assume "affine hull S \ {x. - a \ x \ - b} = {}" then show ?thesis apply (rule_tac x="0" in exI) apply (rule_tac x="1" in exI) using hull_subset by fastforce next assume "affine hull S \ {x. - a \ x \ - b}" then show ?thesis apply (rule_tac x="0" in exI) apply (rule_tac x="0" in exI) using Ssub hull_subset by fastforce qed next case False then obtain a' b' where "a' \ 0" and le: "affine hull S \ {x. a' \ x \ b'} = affine hull S \ {x. - a \ x \ - b}" and eq: "affine hull S \ {x. a' \ x = b'} = affine hull S \ {x. - a \ x = - b}" and mem: "\w. w \ affine hull S \ w + a' \ affine hull S" using affine_parallel_slice affine_affine_hull by metis show ?thesis proof (intro conjI impI allI ballI exI) have *: "S \ - (affine hull S \ {x. P x}) \ affine hull S \ {x. Q x} \ S \ {x. \ P x \ Q x}" for P Q using hull_subset by fastforce have "S \ {x. \ (a' \ x \ b') \ a' \ x = b'}" apply (rule *) apply (simp only: le eq) using Ssub by auto then show "S \ {x. - a' \ x \ - b'}" by auto show "S \ {x. a \ x = b} = S \ {x. - a' \ x = - b'}" using eq hull_subset [of S affine] by force show "\S \ {x. a \ x = b} \ {}; S \ {x. a \ x = b} \ S\ \ - a' \ 0" using \a' \ 0\ by auto show "w + - a' \ affine hull S" if "S \ {x. a \ x = b} \ S" "w \ affine hull S" for w proof - have "w + 1 *\<^sub>R (w - (w + a')) \ affine hull S" using affine_affine_hull mem mem_affine_3_minus that(2) by blast then show ?thesis by simp qed qed qed qed next assume ?rhs then show ?lhs unfolding exposed_face_of_def by blast qed subsection\Extreme points of a set: its singleton faces\ definition\<^marker>\tag important\ extreme_point_of :: "['a::real_vector, 'a set] \ bool" (infixr "(extreme'_point'_of)" 50) where "x extreme_point_of S \ x \ S \ (\a \ S. \b \ S. x \ open_segment a b)" lemma extreme_point_of_stillconvex: "convex S \ (x extreme_point_of S \ x \ S \ convex(S - {x}))" by (fastforce simp add: convex_contains_segment extreme_point_of_def open_segment_def) lemma face_of_singleton: "{x} face_of S \ x extreme_point_of S" by (fastforce simp add: extreme_point_of_def face_of_def) lemma extreme_point_not_in_REL_INTERIOR: fixes S :: "'a::real_normed_vector set" shows "\x extreme_point_of S; S \ {x}\ \ x \ rel_interior S" apply (simp add: face_of_singleton [symmetric]) apply (blast dest: face_of_disjoint_rel_interior) done lemma extreme_point_not_in_interior: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "x extreme_point_of S \ x \ interior S" apply (case_tac "S = {x}") apply (simp add: empty_interior_finite) by (meson contra_subsetD extreme_point_not_in_REL_INTERIOR interior_subset_rel_interior) lemma extreme_point_of_face: "F face_of S \ v extreme_point_of F \ v extreme_point_of S \ v \ F" by (meson empty_subsetI face_of_face face_of_singleton insert_subset) lemma extreme_point_of_convex_hull: "x extreme_point_of (convex hull S) \ x \ S" apply (simp add: extreme_point_of_stillconvex) using hull_minimal [of S "(convex hull S) - {x}" convex] using hull_subset [of S convex] apply blast done proposition extreme_points_of_convex_hull: "{x. x extreme_point_of (convex hull S)} \ S" using extreme_point_of_convex_hull by auto lemma extreme_point_of_empty [simp]: "\ (x extreme_point_of {})" by (simp add: extreme_point_of_def) lemma extreme_point_of_singleton [iff]: "x extreme_point_of {a} \ x = a" using extreme_point_of_stillconvex by auto lemma extreme_point_of_translation_eq: "(a + x) extreme_point_of (image (\x. a + x) S) \ x extreme_point_of S" by (auto simp: extreme_point_of_def) lemma extreme_points_of_translation: "{x. x extreme_point_of (image (\x. a + x) S)} = (\x. a + x) ` {x. x extreme_point_of S}" using extreme_point_of_translation_eq by auto (metis (no_types, lifting) image_iff mem_Collect_eq minus_add_cancel) lemma extreme_points_of_translation_subtract: "{x. x extreme_point_of (image (\x. x - a) S)} = (\x. x - a) ` {x. x extreme_point_of S}" using extreme_points_of_translation [of "- a" S] by simp lemma extreme_point_of_Int: "\x extreme_point_of S; x extreme_point_of T\ \ x extreme_point_of (S \ T)" by (simp add: extreme_point_of_def) lemma extreme_point_of_Int_supporting_hyperplane_le: "\S \ {x. a \ x = b} = {c}; \x. x \ S \ a \ x \ b\ \ c extreme_point_of S" apply (simp add: face_of_singleton [symmetric]) by (metis face_of_Int_supporting_hyperplane_le_strong convex_singleton) lemma extreme_point_of_Int_supporting_hyperplane_ge: "\S \ {x. a \ x = b} = {c}; \x. x \ S \ a \ x \ b\ \ c extreme_point_of S" apply (simp add: face_of_singleton [symmetric]) by (metis face_of_Int_supporting_hyperplane_ge_strong convex_singleton) lemma exposed_point_of_Int_supporting_hyperplane_le: "\S \ {x. a \ x = b} = {c}; \x. x \ S \ a \ x \ b\ \ {c} exposed_face_of S" apply (simp add: exposed_face_of_def face_of_singleton) apply (force simp: extreme_point_of_Int_supporting_hyperplane_le) done lemma exposed_point_of_Int_supporting_hyperplane_ge: "\S \ {x. a \ x = b} = {c}; \x. x \ S \ a \ x \ b\ \ {c} exposed_face_of S" using exposed_point_of_Int_supporting_hyperplane_le [of S "-a" "-b" c] by simp lemma extreme_point_of_convex_hull_insert: "\finite S; a \ convex hull S\ \ a extreme_point_of (convex hull (insert a S))" apply (case_tac "a \ S") apply (simp add: hull_inc) using face_of_convex_hulls [of "insert a S" "{a}"] apply (auto simp: face_of_singleton hull_same) done subsection\Facets\ definition\<^marker>\tag important\ facet_of :: "['a::euclidean_space set, 'a set] \ bool" (infixr "(facet'_of)" 50) where "F facet_of S \ F face_of S \ F \ {} \ aff_dim F = aff_dim S - 1" lemma facet_of_empty [simp]: "\ S facet_of {}" by (simp add: facet_of_def) lemma facet_of_irrefl [simp]: "\ S facet_of S " by (simp add: facet_of_def) lemma facet_of_imp_face_of: "F facet_of S \ F face_of S" by (simp add: facet_of_def) lemma facet_of_imp_subset: "F facet_of S \ F \ S" by (simp add: face_of_imp_subset facet_of_def) lemma hyperplane_facet_of_halfspace_le: "a \ 0 \ {x. a \ x = b} facet_of {x. a \ x \ b}" unfolding facet_of_def hyperplane_eq_empty by (auto simp: hyperplane_face_of_halfspace_ge hyperplane_face_of_halfspace_le Suc_leI of_nat_diff aff_dim_halfspace_le) lemma hyperplane_facet_of_halfspace_ge: "a \ 0 \ {x. a \ x = b} facet_of {x. a \ x \ b}" unfolding facet_of_def hyperplane_eq_empty by (auto simp: hyperplane_face_of_halfspace_le hyperplane_face_of_halfspace_ge Suc_leI of_nat_diff aff_dim_halfspace_ge) lemma facet_of_halfspace_le: "F facet_of {x. a \ x \ b} \ a \ 0 \ F = {x. a \ x = b}" (is "?lhs = ?rhs") proof assume c: ?lhs with c facet_of_irrefl show ?rhs by (force simp: aff_dim_halfspace_le facet_of_def face_of_halfspace_le cong: conj_cong split: if_split_asm) next assume ?rhs then show ?lhs by (simp add: hyperplane_facet_of_halfspace_le) qed lemma facet_of_halfspace_ge: "F facet_of {x. a \ x \ b} \ a \ 0 \ F = {x. a \ x = b}" using facet_of_halfspace_le [of F "-a" "-b"] by simp subsection \Edges: faces of affine dimension 1\ (*FIXME too small subsection, rearrange? *) definition\<^marker>\tag important\ edge_of :: "['a::euclidean_space set, 'a set] \ bool" (infixr "(edge'_of)" 50) where "e edge_of S \ e face_of S \ aff_dim e = 1" lemma edge_of_imp_subset: "S edge_of T \ S \ T" by (simp add: edge_of_def face_of_imp_subset) subsection\Existence of extreme points\ proposition different_norm_3_collinear_points: fixes a :: "'a::euclidean_space" assumes "x \ open_segment a b" "norm(a) = norm(b)" "norm(x) = norm(b)" shows False proof - obtain u where "norm ((1 - u) *\<^sub>R a + u *\<^sub>R b) = norm b" and "a \ b" and u01: "0 < u" "u < 1" using assms by (auto simp: open_segment_image_interval if_splits) then have "(1 - u) *\<^sub>R a \ (1 - u) *\<^sub>R a + ((1 - u) * 2) *\<^sub>R a \ u *\<^sub>R b = (1 - u * u) *\<^sub>R (a \ a)" using assms by (simp add: norm_eq algebra_simps inner_commute) then have "(1 - u) *\<^sub>R ((1 - u) *\<^sub>R a \ a + (2 * u) *\<^sub>R a \ b) = (1 - u) *\<^sub>R ((1 + u) *\<^sub>R (a \ a))" by (simp add: algebra_simps) then have "(1 - u) *\<^sub>R (a \ a) + (2 * u) *\<^sub>R (a \ b) = (1 + u) *\<^sub>R (a \ a)" using u01 by auto then have "a \ b = a \ a" using u01 by (simp add: algebra_simps) then have "a = b" using \norm(a) = norm(b)\ norm_eq vector_eq by fastforce then show ?thesis using \a \ b\ by force qed proposition extreme_point_exists_convex: fixes S :: "'a::euclidean_space set" assumes "compact S" "convex S" "S \ {}" obtains x where "x extreme_point_of S" proof - obtain x where "x \ S" and xsup: "\y. y \ S \ norm y \ norm x" using distance_attains_sup [of S 0] assms by auto have False if "a \ S" "b \ S" and x: "x \ open_segment a b" for a b proof - have noax: "norm a \ norm x" and nobx: "norm b \ norm x" using xsup that by auto have "a \ b" using empty_iff open_segment_idem x by auto have *: "(1 - u) * na + u * nb < norm x" if "na < norm x" "nb \ norm x" "0 < u" "u < 1" for na nb u proof - have "(1 - u) * na + u * nb < (1 - u) * norm x + u * nb" by (simp add: that) also have "... \ (1 - u) * norm x + u * norm x" by (simp add: that) finally have "(1 - u) * na + u * nb < (1 - u) * norm x + u * norm x" . then show ?thesis using scaleR_collapse [symmetric, of "norm x" u] by auto qed have "norm x < norm x" if "norm a < norm x" using x apply (clarsimp simp only: open_segment_image_interval \a \ b\ if_False) apply (rule norm_triangle_lt) apply (simp add: norm_mult) using * [of "norm a" "norm b"] nobx that apply blast done moreover have "norm x < norm x" if "norm b < norm x" using x apply (clarsimp simp only: open_segment_image_interval \a \ b\ if_False) apply (rule norm_triangle_lt) apply (simp add: norm_mult) using * [of "norm b" "norm a" "1-u" for u] noax that apply (simp add: add.commute) done ultimately have "\ (norm a < norm x) \ \ (norm b < norm x)" by auto then show ?thesis using different_norm_3_collinear_points noax nobx that(3) by fastforce qed then show ?thesis apply (rule_tac x=x in that) apply (force simp: extreme_point_of_def \x \ S\) done qed subsection\Krein-Milman, the weaker form\ proposition Krein_Milman: fixes S :: "'a::euclidean_space set" assumes "compact S" "convex S" shows "S = closure(convex hull {x. x extreme_point_of S})" proof (cases "S = {}") case True then show ?thesis by simp next case False have "closed S" by (simp add: \compact S\ compact_imp_closed) have "closure (convex hull {x. x extreme_point_of S}) \ S" apply (rule closure_minimal [OF hull_minimal \closed S\]) using assms apply (auto simp: extreme_point_of_def) done moreover have "u \ closure (convex hull {x. x extreme_point_of S})" if "u \ S" for u proof (rule ccontr) assume unot: "u \ closure(convex hull {x. x extreme_point_of S})" then obtain a b where "a \ u < b" and ab: "\x. x \ closure(convex hull {x. x extreme_point_of S}) \ b < a \ x" using separating_hyperplane_closed_point [of "closure(convex hull {x. x extreme_point_of S})"] by blast have "continuous_on S ((\) a)" by (rule continuous_intros)+ then obtain m where "m \ S" and m: "\y. y \ S \ a \ m \ a \ y" using continuous_attains_inf [of S "\x. a \ x"] \compact S\ \u \ S\ by auto define T where "T = S \ {x. a \ x = a \ m}" have "m \ T" by (simp add: T_def \m \ S\) moreover have "compact T" by (simp add: T_def compact_Int_closed [OF \compact S\ closed_hyperplane]) moreover have "convex T" by (simp add: T_def convex_Int [OF \convex S\ convex_hyperplane]) ultimately obtain v where v: "v extreme_point_of T" using extreme_point_exists_convex [of T] by auto then have "{v} face_of T" by (simp add: face_of_singleton) also have "T face_of S" by (simp add: T_def m face_of_Int_supporting_hyperplane_ge [OF \convex S\]) finally have "v extreme_point_of S" by (simp add: face_of_singleton) then have "b < a \ v" using closure_subset by (simp add: closure_hull hull_inc ab) then show False using \a \ u < b\ \{v} face_of T\ face_of_imp_subset m T_def that by fastforce qed ultimately show ?thesis by blast qed text\Now the sharper form.\ lemma Krein_Milman_Minkowski_aux: fixes S :: "'a::euclidean_space set" assumes n: "dim S = n" and S: "compact S" "convex S" "0 \ S" shows "0 \ convex hull {x. x extreme_point_of S}" using n S proof (induction n arbitrary: S rule: less_induct) case (less n S) show ?case proof (cases "0 \ rel_interior S") case True with Krein_Milman show ?thesis by (metis subsetD convex_convex_hull convex_rel_interior_closure less.prems(2) less.prems(3) rel_interior_subset) next case False have "rel_interior S \ {}" by (simp add: rel_interior_convex_nonempty_aux less) then obtain c where c: "c \ rel_interior S" by blast obtain a where "a \ 0" and le_ay: "\y. y \ S \ a \ 0 \ a \ y" and less_ay: "\y. y \ rel_interior S \ a \ 0 < a \ y" by (blast intro: supporting_hyperplane_rel_boundary intro!: less False) have face: "S \ {x. a \ x = 0} face_of S" apply (rule face_of_Int_supporting_hyperplane_ge [OF \convex S\]) using le_ay by auto then have co: "compact (S \ {x. a \ x = 0})" "convex (S \ {x. a \ x = 0})" using less.prems by (blast intro: face_of_imp_compact face_of_imp_convex)+ have "a \ y = 0" if "y \ span (S \ {x. a \ x = 0})" for y proof - have "y \ span {x. a \ x = 0}" by (metis inf.cobounded2 span_mono subsetCE that) then show ?thesis by (blast intro: span_induct [OF _ subspace_hyperplane]) qed then have "dim (S \ {x. a \ x = 0}) < n" by (metis (no_types) less_ay c subsetD dim_eq_span inf.strict_order_iff inf_le1 \dim S = n\ not_le rel_interior_subset span_0 span_base) then have "0 \ convex hull {x. x extreme_point_of (S \ {x. a \ x = 0})}" by (rule less.IH) (auto simp: co less.prems) then show ?thesis by (metis (mono_tags, lifting) Collect_mono_iff \S \ {x. a \ x = 0} face_of S\ extreme_point_of_face hull_mono subset_iff) qed qed theorem Krein_Milman_Minkowski: fixes S :: "'a::euclidean_space set" assumes "compact S" "convex S" shows "S = convex hull {x. x extreme_point_of S}" proof show "S \ convex hull {x. x extreme_point_of S}" proof fix a assume [simp]: "a \ S" have 1: "compact ((+) (- a) ` S)" by (simp add: \compact S\ compact_translation_subtract cong: image_cong_simp) have 2: "convex ((+) (- a) ` S)" by (simp add: \convex S\ compact_translation_subtract) show a_invex: "a \ convex hull {x. x extreme_point_of S}" using Krein_Milman_Minkowski_aux [OF refl 1 2] convex_hull_translation [of "-a"] by (auto simp: extreme_points_of_translation_subtract translation_assoc cong: image_cong_simp) qed next show "convex hull {x. x extreme_point_of S} \ S" proof - have "{a. a extreme_point_of S} \ S" using extreme_point_of_def by blast then show ?thesis by (simp add: \convex S\ hull_minimal) qed qed subsection\Applying it to convex hulls of explicitly indicated finite sets\ corollary Krein_Milman_polytope: fixes S :: "'a::euclidean_space set" shows "finite S \ convex hull S = convex hull {x. x extreme_point_of (convex hull S)}" by (simp add: Krein_Milman_Minkowski finite_imp_compact_convex_hull) lemma extreme_points_of_convex_hull_eq: fixes S :: "'a::euclidean_space set" shows "\compact S; \T. T \ S \ convex hull T \ convex hull S\ \ {x. x extreme_point_of (convex hull S)} = S" by (metis (full_types) Krein_Milman_Minkowski compact_convex_hull convex_convex_hull extreme_points_of_convex_hull psubsetI) lemma extreme_point_of_convex_hull_eq: fixes S :: "'a::euclidean_space set" shows "\compact S; \T. T \ S \ convex hull T \ convex hull S\ \ (x extreme_point_of (convex hull S) \ x \ S)" using extreme_points_of_convex_hull_eq by auto lemma extreme_point_of_convex_hull_convex_independent: fixes S :: "'a::euclidean_space set" assumes "compact S" and S: "\a. a \ S \ a \ convex hull (S - {a})" shows "(x extreme_point_of (convex hull S) \ x \ S)" proof - have "convex hull T \ convex hull S" if "T \ S" for T proof - obtain a where "T \ S" "a \ S" "a \ T" using \T \ S\ by blast then show ?thesis by (metis (full_types) Diff_eq_empty_iff Diff_insert0 S hull_mono hull_subset insert_Diff_single subsetCE) qed then show ?thesis by (rule extreme_point_of_convex_hull_eq [OF \compact S\]) qed lemma extreme_point_of_convex_hull_affine_independent: fixes S :: "'a::euclidean_space set" shows "\ affine_dependent S \ (x extreme_point_of (convex hull S) \ x \ S)" by (metis aff_independent_finite affine_dependent_def affine_hull_convex_hull extreme_point_of_convex_hull_convex_independent finite_imp_compact hull_inc) text\Elementary proofs exist, not requiring Euclidean spaces and all this development\ lemma extreme_point_of_convex_hull_2: fixes x :: "'a::euclidean_space" shows "x extreme_point_of (convex hull {a,b}) \ x = a \ x = b" proof - have "x extreme_point_of (convex hull {a,b}) \ x \ {a,b}" by (intro extreme_point_of_convex_hull_affine_independent affine_independent_2) then show ?thesis by simp qed lemma extreme_point_of_segment: fixes x :: "'a::euclidean_space" shows "x extreme_point_of closed_segment a b \ x = a \ x = b" by (simp add: extreme_point_of_convex_hull_2 segment_convex_hull) lemma face_of_convex_hull_subset: fixes S :: "'a::euclidean_space set" assumes "compact S" and T: "T face_of (convex hull S)" obtains s' where "s' \ S" "T = convex hull s'" apply (rule_tac s' = "{x. x extreme_point_of T}" in that) using T extreme_point_of_convex_hull extreme_point_of_face apply blast by (metis (no_types) Krein_Milman_Minkowski assms compact_convex_hull convex_convex_hull face_of_imp_compact face_of_imp_convex) lemma face_of_convex_hull_aux: assumes eq: "x *\<^sub>R p = u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c" and x: "u + v + w = x" "x \ 0" and S: "affine S" "a \ S" "b \ S" "c \ S" shows "p \ S" proof - have "p = (u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c) /\<^sub>R x" by (metis \x \ 0\ eq mult.commute right_inverse scaleR_one scaleR_scaleR) moreover have "affine hull {a,b,c} \ S" by (simp add: S hull_minimal) moreover have "(u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c) /\<^sub>R x \ affine hull {a,b,c}" apply (simp add: affine_hull_3) apply (rule_tac x="u/x" in exI) apply (rule_tac x="v/x" in exI) apply (rule_tac x="w/x" in exI) using x apply (auto simp: field_split_simps) done ultimately show ?thesis by force qed proposition face_of_convex_hull_insert_eq: fixes a :: "'a :: euclidean_space" assumes "finite S" and a: "a \ affine hull S" shows "(F face_of (convex hull (insert a S)) \ F face_of (convex hull S) \ (\F'. F' face_of (convex hull S) \ F = convex hull (insert a F')))" (is "F face_of ?CAS \ _") proof safe assume F: "F face_of ?CAS" and *: "\F'. F' face_of convex hull S \ F = convex hull insert a F'" obtain T where T: "T \ insert a S" and FeqT: "F = convex hull T" by (metis F \finite S\ compact_insert finite_imp_compact face_of_convex_hull_subset) show "F face_of convex hull S" proof (cases "a \ T") case True have "F = convex hull insert a (convex hull T \ convex hull S)" proof have "T \ insert a (convex hull T \ convex hull S)" using T hull_subset by fastforce then show "F \ convex hull insert a (convex hull T \ convex hull S)" by (simp add: FeqT hull_mono) show "convex hull insert a (convex hull T \ convex hull S) \ F" apply (rule hull_minimal) using True by (auto simp: \F = convex hull T\ hull_inc) qed moreover have "convex hull T \ convex hull S face_of convex hull S" by (metis F FeqT convex_convex_hull face_of_slice hull_mono inf.absorb_iff2 subset_insertI) ultimately show ?thesis using * by force next case False then show ?thesis by (metis FeqT F T face_of_subset hull_mono subset_insert subset_insertI) qed next assume "F face_of convex hull S" show "F face_of ?CAS" by (simp add: \F face_of convex hull S\ a face_of_convex_hull_insert \finite S\) next fix F assume F: "F face_of convex hull S" show "convex hull insert a F face_of ?CAS" proof (cases "S = {}") case True then show ?thesis using F face_of_affine_eq by auto next case False have anotc: "a \ convex hull S" by (metis (no_types) a affine_hull_convex_hull hull_inc) show ?thesis proof (cases "F = {}") case True show ?thesis using anotc by (simp add: \F = {}\ \finite S\ extreme_point_of_convex_hull_insert face_of_singleton) next case False have "convex hull insert a F \ ?CAS" by (simp add: F a \finite S\ convex_hull_subset face_of_convex_hull_insert face_of_imp_subset hull_inc) moreover have "(\y v. (1 - ub) *\<^sub>R a + ub *\<^sub>R b = (1 - v) *\<^sub>R a + v *\<^sub>R y \ 0 \ v \ v \ 1 \ y \ F) \ (\x u. (1 - uc) *\<^sub>R a + uc *\<^sub>R c = (1 - u) *\<^sub>R a + u *\<^sub>R x \ 0 \ u \ u \ 1 \ x \ F)" if *: "(1 - ux) *\<^sub>R a + ux *\<^sub>R x \ open_segment ((1 - ub) *\<^sub>R a + ub *\<^sub>R b) ((1 - uc) *\<^sub>R a + uc *\<^sub>R c)" and "0 \ ub" "ub \ 1" "0 \ uc" "uc \ 1" "0 \ ux" "ux \ 1" and b: "b \ convex hull S" and c: "c \ convex hull S" and "x \ F" for b c ub uc ux x proof - obtain v where ne: "(1 - ub) *\<^sub>R a + ub *\<^sub>R b \ (1 - uc) *\<^sub>R a + uc *\<^sub>R c" and eq: "(1 - ux) *\<^sub>R a + ux *\<^sub>R x = (1 - v) *\<^sub>R ((1 - ub) *\<^sub>R a + ub *\<^sub>R b) + v *\<^sub>R ((1 - uc) *\<^sub>R a + uc *\<^sub>R c)" and "0 < v" "v < 1" using * by (auto simp: in_segment) then have 0: "((1 - ux) - ((1 - v) * (1 - ub) + v * (1 - uc))) *\<^sub>R a + (ux *\<^sub>R x - (((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c)) = 0" by (auto simp: algebra_simps) then have "((1 - ux) - ((1 - v) * (1 - ub) + v * (1 - uc))) *\<^sub>R a = ((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c + (-ux) *\<^sub>R x" by (auto simp: algebra_simps) then have "a \ affine hull S" if "1 - ux - ((1 - v) * (1 - ub) + v * (1 - uc)) \ 0" apply (rule face_of_convex_hull_aux) using b c that apply (auto simp: algebra_simps) using F convex_hull_subset_affine_hull face_of_imp_subset \x \ F\ apply blast+ done then have "1 - ux - ((1 - v) * (1 - ub) + v * (1 - uc)) = 0" using a by blast with 0 have equx: "(1 - v) * ub + v * uc = ux" and uxx: "ux *\<^sub>R x = (((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c)" by auto (auto simp: algebra_simps) show ?thesis proof (cases "uc = 0") case True then show ?thesis using equx 0 \0 \ ub\ \ub \ 1\ \v < 1\ \x \ F\ apply (auto simp: algebra_simps) apply (rule_tac x=x in exI, simp) apply (rule_tac x=ub in exI, auto) apply (metis add.left_neutral diff_eq_eq less_irrefl mult.commute mult_cancel_right1 real_vector.scale_cancel_left real_vector.scale_left_diff_distrib) using \x \ F\ \uc \ 1\ apply blast done next case False show ?thesis proof (cases "ub = 0") case True then show ?thesis using equx 0 \0 \ uc\ \uc \ 1\ \0 < v\ \x \ F\ \uc \ 0\ by (force simp: algebra_simps) next case False then have "0 < ub" "0 < uc" using \uc \ 0\ \0 \ ub\ \0 \ uc\ by auto then have "ux \ 0" by (metis \0 < v\ \v < 1\ diff_ge_0_iff_ge dual_order.strict_implies_order equx leD le_add_same_cancel2 zero_le_mult_iff zero_less_mult_iff) have "b \ F \ c \ F" proof (cases "b = c") case True then show ?thesis by (metis \ux \ 0\ equx real_vector.scale_cancel_left scaleR_add_left uxx \x \ F\) next case False have "x = (((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c) /\<^sub>R ux" by (metis \ux \ 0\ uxx mult.commute right_inverse scaleR_one scaleR_scaleR) also have "... = (1 - v * uc / ux) *\<^sub>R b + (v * uc / ux) *\<^sub>R c" using \ux \ 0\ equx apply (auto simp: field_split_simps) by (metis add.commute add_diff_eq add_divide_distrib diff_add_cancel scaleR_add_left) finally have "x = (1 - v * uc / ux) *\<^sub>R b + (v * uc / ux) *\<^sub>R c" . then have "x \ open_segment b c" apply (simp add: in_segment \b \ c\) apply (rule_tac x="(v * uc) / ux" in exI) using \0 \ ux\ \ux \ 0\ \0 < uc\ \0 < v\ \0 < ub\ \v < 1\ equx apply (force simp: field_split_simps) done then show ?thesis by (rule face_ofD [OF F _ b c \x \ F\]) qed with \0 \ ub\ \ub \ 1\ \0 \ uc\ \uc \ 1\ show ?thesis by blast qed qed qed moreover have "convex hull F = F" by (meson F convex_hull_eq face_of_imp_convex) ultimately show ?thesis unfolding face_of_def by (fastforce simp: convex_hull_insert_alt \S \ {}\ \F \ {}\) qed qed qed lemma face_of_convex_hull_insert2: fixes a :: "'a :: euclidean_space" assumes S: "finite S" and a: "a \ affine hull S" and F: "F face_of convex hull S" shows "convex hull (insert a F) face_of convex hull (insert a S)" by (metis F face_of_convex_hull_insert_eq [OF S a]) proposition face_of_convex_hull_affine_independent: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "(T face_of (convex hull S) \ (\c. c \ S \ T = convex hull c))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (meson \T face_of convex hull S\ aff_independent_finite assms face_of_convex_hull_subset finite_imp_compact) next assume ?rhs then obtain c where "c \ S" and T: "T = convex hull c" by blast have "affine hull c \ affine hull (S - c) = {}" apply (rule disjoint_affine_hull [OF assms \c \ S\], auto) done then have "affine hull c \ convex hull (S - c) = {}" using convex_hull_subset_affine_hull by fastforce then show ?lhs by (metis face_of_convex_hulls \c \ S\ aff_independent_finite assms T) qed lemma facet_of_convex_hull_affine_independent: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "T facet_of (convex hull S) \ T \ {} \ (\u. u \ S \ T = convex hull (S - {u}))" (is "?lhs = ?rhs") proof assume ?lhs then have "T face_of (convex hull S)" "T \ {}" and afft: "aff_dim T = aff_dim (convex hull S) - 1" by (auto simp: facet_of_def) then obtain c where "c \ S" and c: "T = convex hull c" by (auto simp: face_of_convex_hull_affine_independent [OF assms]) then have affs: "aff_dim S = aff_dim c + 1" by (metis aff_dim_convex_hull afft eq_diff_eq) have "\ affine_dependent c" using \c \ S\ affine_dependent_subset assms by blast with affs have "card (S - c) = 1" apply (simp add: aff_dim_affine_independent [symmetric] aff_dim_convex_hull) by (metis aff_dim_affine_independent aff_independent_finite One_nat_def \c \ S\ add.commute add_diff_cancel_right' assms card_Diff_subset card_mono of_nat_1 of_nat_diff of_nat_eq_iff) then obtain u where u: "u \ S - c" by (metis DiffI \c \ S\ aff_independent_finite assms cancel_comm_monoid_add_class.diff_cancel card_Diff_subset subsetI subset_antisym zero_neq_one) then have u: "S = insert u c" by (metis Diff_subset \c \ S\ \card (S - c) = 1\ card_1_singletonE double_diff insert_Diff insert_subset singletonD) have "T = convex hull (c - {u})" by (metis Diff_empty Diff_insert0 \T facet_of convex hull S\ c facet_of_irrefl insert_absorb u) with \T \ {}\ show ?rhs using c u by auto next assume ?rhs then obtain u where "T \ {}" "u \ S" and u: "T = convex hull (S - {u})" by (force simp: facet_of_def) then have "\ S \ {u}" using \T \ {}\ u by auto have [simp]: "aff_dim (convex hull (S - {u})) = aff_dim (convex hull S) - 1" using assms \u \ S\ apply (simp add: aff_dim_convex_hull affine_dependent_def) apply (drule bspec, assumption) by (metis add_diff_cancel_right' aff_dim_insert insert_Diff [of u S]) show ?lhs apply (subst u) apply (simp add: \\ S \ {u}\ facet_of_def face_of_convex_hull_affine_independent [OF assms], blast) done qed lemma facet_of_convex_hull_affine_independent_alt: fixes S :: "'a::euclidean_space set" shows "\affine_dependent S \ (T facet_of (convex hull S) \ 2 \ card S \ (\u. u \ S \ T = convex hull (S - {u})))" apply (simp add: facet_of_convex_hull_affine_independent) apply (auto simp: Set.subset_singleton_iff) apply (metis Diff_cancel Int_empty_right Int_insert_right_if1 aff_independent_finite card_eq_0_iff card_insert_if card_mono card_subset_eq convex_hull_eq_empty eq_iff equals0D finite_insert finite_subset inf.absorb_iff2 insert_absorb insert_not_empty not_less_eq_eq numeral_2_eq_2) done lemma segment_face_of: assumes "(closed_segment a b) face_of S" shows "a extreme_point_of S" "b extreme_point_of S" proof - have as: "{a} face_of S" by (metis (no_types) assms convex_hull_singleton empty_iff extreme_point_of_convex_hull_insert face_of_face face_of_singleton finite.emptyI finite.insertI insert_absorb insert_iff segment_convex_hull) moreover have "{b} face_of S" proof - have "b \ convex hull {a} \ b extreme_point_of convex hull {b, a}" by (meson extreme_point_of_convex_hull_insert finite.emptyI finite.insertI) moreover have "closed_segment a b = convex hull {b, a}" using closed_segment_commute segment_convex_hull by blast ultimately show ?thesis by (metis as assms face_of_face convex_hull_singleton empty_iff face_of_singleton insertE) qed ultimately show "a extreme_point_of S" "b extreme_point_of S" using face_of_singleton by blast+ qed proposition Krein_Milman_frontier: fixes S :: "'a::euclidean_space set" assumes "convex S" "compact S" shows "S = convex hull (frontier S)" (is "?lhs = ?rhs") proof have "?lhs \ convex hull {x. x extreme_point_of S}" using Krein_Milman_Minkowski assms by blast also have "... \ ?rhs" apply (rule hull_mono) apply (auto simp: frontier_def extreme_point_not_in_interior) using closure_subset apply (force simp: extreme_point_of_def) done finally show "?lhs \ ?rhs" . next have "?rhs \ convex hull S" by (metis Diff_subset \compact S\ closure_closed compact_eq_bounded_closed frontier_def hull_mono) also have "... \ ?lhs" by (simp add: \convex S\ hull_same) finally show "?rhs \ ?lhs" . qed subsection\Polytopes\ definition\<^marker>\tag important\ polytope where "polytope S \ \v. finite v \ S = convex hull v" lemma polytope_translation_eq: "polytope (image (\x. a + x) S) \ polytope S" apply (simp add: polytope_def, safe) apply (metis convex_hull_translation finite_imageI translation_galois) by (metis convex_hull_translation finite_imageI) lemma polytope_linear_image: "\linear f; polytope p\ \ polytope(image f p)" unfolding polytope_def using convex_hull_linear_image by blast lemma polytope_empty: "polytope {}" using convex_hull_empty polytope_def by blast lemma polytope_convex_hull: "finite S \ polytope(convex hull S)" using polytope_def by auto lemma polytope_Times: "\polytope S; polytope T\ \ polytope(S \ T)" unfolding polytope_def by (metis finite_cartesian_product convex_hull_Times) lemma face_of_polytope_polytope: fixes S :: "'a::euclidean_space set" shows "\polytope S; F face_of S\ \ polytope F" unfolding polytope_def by (meson face_of_convex_hull_subset finite_imp_compact finite_subset) lemma finite_polytope_faces: fixes S :: "'a::euclidean_space set" assumes "polytope S" shows "finite {F. F face_of S}" proof - obtain v where "finite v" "S = convex hull v" using assms polytope_def by auto have "finite ((hull) convex ` {T. T \ v})" by (simp add: \finite v\) moreover have "{F. F face_of S} \ ((hull) convex ` {T. T \ v})" by (metis (no_types, lifting) \finite v\ \S = convex hull v\ face_of_convex_hull_subset finite_imp_compact image_eqI mem_Collect_eq subsetI) ultimately show ?thesis by (blast intro: finite_subset) qed lemma finite_polytope_facets: assumes "polytope S" shows "finite {T. T facet_of S}" by (simp add: assms facet_of_def finite_polytope_faces) lemma polytope_scaling: assumes "polytope S" shows "polytope (image (\x. c *\<^sub>R x) S)" by (simp add: assms polytope_linear_image) lemma polytope_imp_compact: fixes S :: "'a::real_normed_vector set" shows "polytope S \ compact S" by (metis finite_imp_compact_convex_hull polytope_def) lemma polytope_imp_convex: "polytope S \ convex S" by (metis convex_convex_hull polytope_def) lemma polytope_imp_closed: fixes S :: "'a::real_normed_vector set" shows "polytope S \ closed S" by (simp add: compact_imp_closed polytope_imp_compact) lemma polytope_imp_bounded: fixes S :: "'a::real_normed_vector set" shows "polytope S \ bounded S" by (simp add: compact_imp_bounded polytope_imp_compact) lemma polytope_interval: "polytope(cbox a b)" unfolding polytope_def by (meson closed_interval_as_convex_hull) lemma polytope_sing: "polytope {a}" using polytope_def by force lemma face_of_polytope_insert: "\polytope S; a \ affine hull S; F face_of S\ \ F face_of convex hull (insert a S)" by (metis (no_types, lifting) affine_hull_convex_hull face_of_convex_hull_insert hull_insert polytope_def) proposition face_of_polytope_insert2: fixes a :: "'a :: euclidean_space" assumes "polytope S" "a \ affine hull S" "F face_of S" shows "convex hull (insert a F) face_of convex hull (insert a S)" proof - obtain V where "finite V" "S = convex hull V" using assms by (auto simp: polytope_def) then have "convex hull (insert a F) face_of convex hull (insert a V)" using affine_hull_convex_hull assms face_of_convex_hull_insert2 by blast then show ?thesis by (metis \S = convex hull V\ hull_insert) qed subsection\Polyhedra\ definition\<^marker>\tag important\ polyhedron where "polyhedron S \ \F. finite F \ S = \ F \ (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b})" lemma polyhedron_Int [intro,simp]: "\polyhedron S; polyhedron T\ \ polyhedron (S \ T)" apply (simp add: polyhedron_def, clarify) apply (rename_tac F G) apply (rule_tac x="F \ G" in exI, auto) done lemma polyhedron_UNIV [iff]: "polyhedron UNIV" unfolding polyhedron_def by (rule_tac x="{}" in exI) auto lemma polyhedron_Inter [intro,simp]: "\finite F; \S. S \ F \ polyhedron S\ \ polyhedron(\F)" by (induction F rule: finite_induct) auto lemma polyhedron_empty [iff]: "polyhedron ({} :: 'a :: euclidean_space set)" proof - have "\a. a \ 0 \ (\b. {x. (SOME i. i \ Basis) \ x \ - 1} = {x. a \ x \ b})" by (rule_tac x="(SOME i. i \ Basis)" in exI) (force simp: SOME_Basis nonzero_Basis) moreover have "\a b. a \ 0 \ {x. - (SOME i. i \ Basis) \ x \ - 1} = {x. a \ x \ b}" apply (rule_tac x="-(SOME i. i \ Basis)" in exI) apply (rule_tac x="-1" in exI) apply (simp add: SOME_Basis nonzero_Basis) done ultimately show ?thesis unfolding polyhedron_def apply (rule_tac x="{{x. (SOME i. i \ Basis) \ x \ -1}, {x. -(SOME i. i \ Basis) \ x \ -1}}" in exI) apply force done qed lemma polyhedron_halfspace_le: fixes a :: "'a :: euclidean_space" shows "polyhedron {x. a \ x \ b}" proof (cases "a = 0") case True then show ?thesis by auto next case False then show ?thesis unfolding polyhedron_def by (rule_tac x="{{x. a \ x \ b}}" in exI) auto qed lemma polyhedron_halfspace_ge: fixes a :: "'a :: euclidean_space" shows "polyhedron {x. a \ x \ b}" using polyhedron_halfspace_le [of "-a" "-b"] by simp lemma polyhedron_hyperplane: fixes a :: "'a :: euclidean_space" shows "polyhedron {x. a \ x = b}" proof - have "{x. a \ x = b} = {x. a \ x \ b} \ {x. a \ x \ b}" by force then show ?thesis by (simp add: polyhedron_halfspace_ge polyhedron_halfspace_le) qed lemma affine_imp_polyhedron: fixes S :: "'a :: euclidean_space set" shows "affine S \ polyhedron S" by (metis affine_hull_eq polyhedron_Inter polyhedron_hyperplane affine_hull_finite_intersection_hyperplanes [of S]) lemma polyhedron_imp_closed: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ closed S" apply (simp add: polyhedron_def) using closed_halfspace_le by fastforce lemma polyhedron_imp_convex: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ convex S" apply (simp add: polyhedron_def) using convex_Inter convex_halfspace_le by fastforce lemma polyhedron_affine_hull: fixes S :: "'a :: euclidean_space set" shows "polyhedron(affine hull S)" by (simp add: affine_imp_polyhedron) subsection\Canonical polyhedron representation making facial structure explicit\ proposition polyhedron_Int_affine: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ (\F. finite F \ S = (affine hull S) \ \F \ (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b}))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (simp add: polyhedron_def) apply (erule ex_forward) using hull_subset apply force done next assume ?rhs then show ?lhs apply clarify apply (erule ssubst) apply (force intro: polyhedron_affine_hull polyhedron_halfspace_le) done qed proposition rel_interior_polyhedron_explicit: assumes "finite F" and seq: "S = affine hull S \ \F" and faceq: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" and psub: "\F'. F' \ F \ S \ affine hull S \ \F'" shows "rel_interior S = {x \ S. \h \ F. a h \ x < b h}" proof - have rels: "\x. x \ rel_interior S \ x \ S" by (meson IntE mem_rel_interior) moreover have "a i \ x < b i" if x: "x \ rel_interior S" and "i \ F" for x i proof - have fif: "F - {i} \ F" using \i \ F\ Diff_insert_absorb Diff_subset set_insert psubsetI by blast then have "S \ affine hull S \ \(F - {i})" by (rule psub) then obtain z where ssub: "S \ \(F - {i})" and zint: "z \ \(F - {i})" and "z \ S" and zaff: "z \ affine hull S" by auto have "z \ x" using \z \ S\ rels x by blast have "z \ affine hull S \ \F" using \z \ S\ seq by auto then have aiz: "a i \ z > b i" using faceq zint zaff by fastforce obtain e where "e > 0" "x \ S" and e: "ball x e \ affine hull S \ S" using x by (auto simp: mem_rel_interior_ball) then have ins: "\y. \norm (x - y) < e; y \ affine hull S\ \ y \ S" by (metis IntI subsetD dist_norm mem_ball) define \ where "\ = min (1/2) (e / 2 / norm(z - x))" have "norm (\ *\<^sub>R x - \ *\<^sub>R z) = norm (\ *\<^sub>R (x - z))" by (simp add: \_def algebra_simps norm_mult) also have "... = \ * norm (x - z)" using \e > 0\ by (simp add: \_def) also have "... < e" using \z \ x\ \e > 0\ by (simp add: \_def min_def field_split_simps norm_minus_commute) finally have lte: "norm (\ *\<^sub>R x - \ *\<^sub>R z) < e" . have \_aff: "\ *\<^sub>R z + (1 - \) *\<^sub>R x \ affine hull S" by (metis \x \ S\ add.commute affine_affine_hull diff_add_cancel hull_inc mem_affine zaff) have "\ *\<^sub>R z + (1 - \) *\<^sub>R x \ S" apply (rule ins [OF _ \_aff]) apply (simp add: algebra_simps lte) done then obtain l where l: "0 < l" "l < 1" and ls: "(l *\<^sub>R z + (1 - l) *\<^sub>R x) \ S" apply (rule_tac l = \ in that) using \e > 0\ \z \ x\ apply (auto simp: \_def) done then have i: "l *\<^sub>R z + (1 - l) *\<^sub>R x \ i" using seq \i \ F\ by auto have "b i * l + (a i \ x) * (1 - l) < a i \ (l *\<^sub>R z + (1 - l) *\<^sub>R x)" using l by (simp add: algebra_simps aiz) also have "\ \ b i" using i l using faceq mem_Collect_eq \i \ F\ by blast finally have "(a i \ x) * (1 - l) < b i * (1 - l)" by (simp add: algebra_simps) with l show ?thesis by simp qed moreover have "x \ rel_interior S" if "x \ S" and less: "\h. h \ F \ a h \ x < b h" for x proof - have 1: "\h. h \ F \ x \ interior h" by (metis interior_halfspace_le mem_Collect_eq less faceq) have 2: "\y. \\h\F. y \ interior h; y \ affine hull S\ \ y \ S" by (metis IntI Inter_iff contra_subsetD interior_subset seq) show ?thesis apply (simp add: rel_interior \x \ S\) apply (rule_tac x="\h\F. interior h" in exI) apply (auto simp: \finite F\ open_INT 1 2) done qed ultimately show ?thesis by blast qed lemma polyhedron_Int_affine_parallel: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ (\F. finite F \ S = (affine hull S) \ (\F) \ (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b} \ (\x \ affine hull S. (x + a) \ affine hull S)))" (is "?lhs = ?rhs") proof assume ?lhs then obtain F where "finite F" and seq: "S = (affine hull S) \ \F" and faces: "\h. h \ F \ \a b. a \ 0 \ h = {x. a \ x \ b}" by (fastforce simp add: polyhedron_Int_affine) then obtain a b where ab: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" by metis show ?rhs proof - have "\a' b'. a' \ 0 \ affine hull S \ {x. a' \ x \ b'} = affine hull S \ h \ (\w \ affine hull S. (w + a') \ affine hull S)" if "h \ F" "\(affine hull S \ h)" for h proof - have "a h \ 0" and "h = {x. a h \ x \ b h}" "h \ \F = \F" using \h \ F\ ab by auto then have "(affine hull S) \ {x. a h \ x \ b h} \ {}" by (metis (no_types) affine_hull_eq_empty inf.absorb_iff2 inf_assoc inf_bot_right inf_commute seq that(2)) moreover have "\ (affine hull S \ {x. a h \ x \ b h})" using \h = {x. a h \ x \ b h}\ that(2) by blast ultimately show ?thesis using affine_parallel_slice [of "affine hull S"] by (metis \h = {x. a h \ x \ b h}\ affine_affine_hull) qed then obtain a b where ab: "\h. \h \ F; \ (affine hull S \ h)\ \ a h \ 0 \ affine hull S \ {x. a h \ x \ b h} = affine hull S \ h \ (\w \ affine hull S. (w + a h) \ affine hull S)" by metis have seq2: "S = affine hull S \ (\h\{h \ F. \ affine hull S \ h}. {x. a h \ x \ b h})" by (subst seq) (auto simp: ab INT_extend_simps) show ?thesis apply (rule_tac x="(\h. {x. a h \ x \ b h}) ` {h. h \ F \ \(affine hull S \ h)}" in exI) apply (intro conjI seq2) using \finite F\ apply force using ab apply blast done qed next assume ?rhs then show ?lhs apply (simp add: polyhedron_Int_affine) by metis qed proposition polyhedron_Int_affine_parallel_minimal: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ (\F. finite F \ S = (affine hull S) \ (\F) \ (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b} \ (\x \ affine hull S. (x + a) \ affine hull S)) \ (\F'. F' \ F \ S \ (affine hull S) \ (\F')))" (is "?lhs = ?rhs") proof assume ?lhs then obtain f0 where f0: "finite f0" "S = (affine hull S) \ (\f0)" (is "?P f0") "\h \ f0. \a b. a \ 0 \ h = {x. a \ x \ b} \ (\x \ affine hull S. (x + a) \ affine hull S)" (is "?Q f0") by (force simp: polyhedron_Int_affine_parallel) define n where "n = (LEAST n. \F. card F = n \ finite F \ ?P F \ ?Q F)" have nf: "\F. card F = n \ finite F \ ?P F \ ?Q F" apply (simp add: n_def) apply (rule LeastI [where k = "card f0"]) using f0 apply auto done then obtain F where F: "card F = n" "finite F" and seq: "?P F" and aff: "?Q F" by blast then have "\ (finite g \ ?P g \ ?Q g)" if "card g < n" for g using that by (auto simp: n_def dest!: not_less_Least) then have *: "\ (?P g \ ?Q g)" if "g \ F" for g using that \finite F\ psubset_card_mono \card F = n\ by (metis finite_Int inf.strict_order_iff) have 1: "\F'. F' \ F \ S \ affine hull S \ \F'" by (subst seq) blast have 2: "\F'. F' \ F \ S \ affine hull S \ \F'" apply (frule *) by (metis aff subsetCE subset_iff_psubset_eq) show ?rhs by (metis \finite F\ seq aff psubsetI 1 2) next assume ?rhs then show ?lhs by (auto simp: polyhedron_Int_affine_parallel) qed lemma polyhedron_Int_affine_minimal: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ (\F. finite F \ S = (affine hull S) \ \F \ (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b}) \ (\F'. F' \ F \ S \ (affine hull S) \ \F'))" apply (rule iffI) apply (force simp: polyhedron_Int_affine_parallel_minimal elim!: ex_forward) apply (auto simp: polyhedron_Int_affine elim!: ex_forward) done proposition facet_of_polyhedron_explicit: assumes "finite F" and seq: "S = affine hull S \ \F" and faceq: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" and psub: "\F'. F' \ F \ S \ affine hull S \ \F'" shows "c facet_of S \ (\h. h \ F \ c = S \ {x. a h \ x = b h})" proof (cases "S = {}") case True with psub show ?thesis by force next case False have "polyhedron S" apply (simp add: polyhedron_Int_affine) apply (rule_tac x=F in exI) using assms apply force done then have "convex S" by (rule polyhedron_imp_convex) with False rel_interior_eq_empty have "rel_interior S \ {}" by blast then obtain x where "x \ rel_interior S" by auto then obtain T where "open T" "x \ T" "x \ S" "T \ affine hull S \ S" by (force simp: mem_rel_interior) then have xaff: "x \ affine hull S" and xint: "x \ \F" using seq hull_inc by auto have "rel_interior S = {x \ S. \h\F. a h \ x < b h}" by (rule rel_interior_polyhedron_explicit [OF \finite F\ seq faceq psub]) with \x \ rel_interior S\ have [simp]: "\h. h\F \ a h \ x < b h" by blast have *: "(S \ {x. a h \ x = b h}) facet_of S" if "h \ F" for h proof - have "S \ affine hull S \ \(F - {h})" using psub that by (metis Diff_disjoint Diff_subset insert_disjoint(2) psubsetI) then obtain z where zaff: "z \ affine hull S" and zint: "z \ \(F - {h})" and "z \ S" by force then have "z \ x" "z \ h" using seq \x \ S\ by auto have "x \ h" using that xint by auto then have able: "a h \ x \ b h" using faceq that by blast also have "... < a h \ z" using \z \ h\ faceq [OF that] xint by auto finally have xltz: "a h \ x < a h \ z" . define l where "l = (b h - a h \ x) / (a h \ z - a h \ x)" define w where "w = (1 - l) *\<^sub>R x + l *\<^sub>R z" have "0 < l" "l < 1" using able xltz \b h < a h \ z\ \h \ F\ by (auto simp: l_def field_split_simps) have awlt: "a i \ w < b i" if "i \ F" "i \ h" for i proof - have "(1 - l) * (a i \ x) < (1 - l) * b i" by (simp add: \l < 1\ \i \ F\) moreover have "l * (a i \ z) \ l * b i" apply (rule mult_left_mono) apply (metis Diff_insert_absorb Inter_iff Set.set_insert \h \ F\ faceq insertE mem_Collect_eq that zint) using \0 < l\ apply simp done ultimately show ?thesis by (simp add: w_def algebra_simps) qed have weq: "a h \ w = b h" using xltz unfolding w_def l_def by (simp add: algebra_simps) (simp add: field_simps) have "w \ affine hull S" by (simp add: w_def mem_affine xaff zaff) moreover have "w \ \F" using \a h \ w = b h\ awlt faceq less_eq_real_def by blast ultimately have "w \ S" using seq by blast with weq have "S \ {x. a h \ x = b h} \ {}" by blast moreover have "S \ {x. a h \ x = b h} face_of S" apply (rule face_of_Int_supporting_hyperplane_le) apply (rule \convex S\) apply (subst (asm) seq) using faceq that apply fastforce done moreover have "affine hull (S \ {x. a h \ x = b h}) = (affine hull S) \ {x. a h \ x = b h}" proof show "affine hull (S \ {x. a h \ x = b h}) \ affine hull S \ {x. a h \ x = b h}" apply (intro Int_greatest hull_mono Int_lower1) apply (metis affine_hull_eq affine_hyperplane hull_mono inf_le2) done next show "affine hull S \ {x. a h \ x = b h} \ affine hull (S \ {x. a h \ x = b h})" proof fix y assume yaff: "y \ affine hull S \ {y. a h \ y = b h}" obtain T where "0 < T" and T: "\j. \j \ F; j \ h\ \ T * (a j \ y - a j \ w) \ b j - a j \ w" proof (cases "F - {h} = {}") case True then show ?thesis by (rule_tac T=1 in that) auto next case False then obtain h' where h': "h' \ F - {h}" by auto let ?body = "(\j. if 0 < a j \ y - a j \ w then (b j - a j \ w) / (a j \ y - a j \ w) else 1) ` (F - {h})" define inff where "inff = Inf ?body" from \finite F\ have "finite ?body" by blast moreover from h' have "?body \ {}" by blast moreover have "j > 0" if "j \ ?body" for j proof - from that obtain x where "x \ F" and "x \ h" and *: "j = (if 0 < a x \ y - a x \ w then (b x - a x \ w) / (a x \ y - a x \ w) else 1)" by blast with awlt [of x] have "a x \ w < b x" by simp with * show ?thesis by simp qed ultimately have "0 < inff" by (simp_all add: finite_less_Inf_iff inff_def) moreover have "inff * (a j \ y - a j \ w) \ b j - a j \ w" if "j \ F" "j \ h" for j proof (cases "a j \ w < a j \ y") case True then have "inff \ (b j - a j \ w) / (a j \ y - a j \ w)" apply (simp add: inff_def) apply (rule cInf_le_finite) using \finite F\ apply blast apply (simp add: that split: if_split_asm) done then show ?thesis using \0 < inff\ awlt [OF that] mult_strict_left_mono by (fastforce simp add: field_split_simps split: if_split_asm) next case False with \0 < inff\ have "inff * (a j \ y - a j \ w) \ 0" by (simp add: mult_le_0_iff) also have "... < b j - a j \ w" by (simp add: awlt that) finally show ?thesis by simp qed ultimately show ?thesis by (blast intro: that) qed define c where "c = (1 - T) *\<^sub>R w + T *\<^sub>R y" have "(1 - T) *\<^sub>R w + T *\<^sub>R y \ j" if "j \ F" for j proof (cases "j = h") case True have "(1 - T) *\<^sub>R w + T *\<^sub>R y \ {x. a h \ x \ b h}" using weq yaff by (auto simp: algebra_simps) with True faceq [OF that] show ?thesis by metis next case False with T that have "(1 - T) *\<^sub>R w + T *\<^sub>R y \ {x. a j \ x \ b j}" by (simp add: algebra_simps) with faceq [OF that] show ?thesis by simp qed moreover have "(1 - T) *\<^sub>R w + T *\<^sub>R y \ affine hull S" apply (rule affine_affine_hull [simplified affine_alt, rule_format]) apply (simp add: \w \ affine hull S\) using yaff apply blast done ultimately have "c \ S" using seq by (force simp: c_def) moreover have "a h \ c = b h" using yaff by (force simp: c_def algebra_simps weq) ultimately have caff: "c \ affine hull (S \ {y. a h \ y = b h})" by (simp add: hull_inc) have waff: "w \ affine hull (S \ {y. a h \ y = b h})" using \w \ S\ weq by (blast intro: hull_inc) have yeq: "y = (1 - inverse T) *\<^sub>R w + c /\<^sub>R T" using \0 < T\ by (simp add: c_def algebra_simps) show "y \ affine hull (S \ {y. a h \ y = b h})" by (metis yeq affine_affine_hull [simplified affine_alt, rule_format, OF waff caff]) qed qed ultimately show ?thesis apply (simp add: facet_of_def) apply (subst aff_dim_affine_hull [symmetric]) using \b h < a h \ z\ zaff apply (force simp: aff_dim_affine_Int_hyperplane) done qed show ?thesis proof show "\h. h \ F \ c = S \ {x. a h \ x = b h} \ c facet_of S" using * by blast next assume "c facet_of S" then have "c face_of S" "convex c" "c \ {}" and affc: "aff_dim c = aff_dim S - 1" by (auto simp: facet_of_def face_of_imp_convex) then obtain x where x: "x \ rel_interior c" by (force simp: rel_interior_eq_empty) then have "x \ c" by (meson subsetD rel_interior_subset) then have "x \ S" using \c facet_of S\ facet_of_imp_subset by blast have rels: "rel_interior S = {x \ S. \h\F. a h \ x < b h}" by (rule rel_interior_polyhedron_explicit [OF assms]) have "c \ S" using \c facet_of S\ facet_of_irrefl by blast then have "x \ rel_interior S" by (metis IntI empty_iff \x \ c\ \c \ S\ \c face_of S\ face_of_disjoint_rel_interior) with rels \x \ S\ obtain i where "i \ F" and i: "a i \ x \ b i" by force have "x \ {u. a i \ u \ b i}" by (metis IntD2 InterE \i \ F\ \x \ S\ faceq seq) then have "a i \ x \ b i" by simp then have "a i \ x = b i" using i by auto have "c \ S \ {x. a i \ x = b i}" apply (rule subset_of_face_of [of _ S]) apply (simp add: "*" \i \ F\ facet_of_imp_face_of) apply (simp add: \c face_of S\ face_of_imp_subset) using \a i \ x = b i\ \x \ S\ x by blast then have cface: "c face_of (S \ {x. a i \ x = b i})" by (meson \c face_of S\ face_of_subset inf_le1) have con: "convex (S \ {x. a i \ x = b i})" by (simp add: \convex S\ convex_Int convex_hyperplane) show "\h. h \ F \ c = S \ {x. a h \ x = b h}" apply (rule_tac x=i in exI) apply (simp add: \i \ F\) by (metis (no_types) * \i \ F\ affc facet_of_def less_irrefl face_of_aff_dim_lt [OF con cface]) qed qed lemma face_of_polyhedron_subset_explicit: fixes S :: "'a :: euclidean_space set" assumes "finite F" and seq: "S = affine hull S \ \F" and faceq: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" and psub: "\F'. F' \ F \ S \ affine hull S \ \F'" and c: "c face_of S" and "c \ {}" "c \ S" obtains h where "h \ F" "c \ S \ {x. a h \ x = b h}" proof - have "c \ S" using \c face_of S\ by (simp add: face_of_imp_subset) have "polyhedron S" apply (simp add: polyhedron_Int_affine) by (metis \finite F\ faceq seq) then have "convex S" by (simp add: polyhedron_imp_convex) then have *: "(S \ {x. a h \ x = b h}) face_of S" if "h \ F" for h apply (rule face_of_Int_supporting_hyperplane_le) using faceq seq that by fastforce have "rel_interior c \ {}" using c \c \ {}\ face_of_imp_convex rel_interior_eq_empty by blast then obtain x where "x \ rel_interior c" by auto have rels: "rel_interior S = {x \ S. \h\F. a h \ x < b h}" by (rule rel_interior_polyhedron_explicit [OF \finite F\ seq faceq psub]) then have xnot: "x \ rel_interior S" by (metis IntI \x \ rel_interior c\ c \c \ S\ contra_subsetD empty_iff face_of_disjoint_rel_interior rel_interior_subset) then have "x \ S" using \c \ S\ \x \ rel_interior c\ rel_interior_subset by auto then have xint: "x \ \F" using seq by blast have "F \ {}" using assms by (metis affine_Int affine_Inter affine_affine_hull ex_in_conv face_of_affine_trivial) then obtain i where "i \ F" "\ (a i \ x < b i)" using \x \ S\ rels xnot by auto with xint have "a i \ x = b i" by (metis eq_iff mem_Collect_eq not_le Inter_iff faceq) have face: "S \ {x. a i \ x = b i} face_of S" by (simp add: "*" \i \ F\) show ?thesis apply (rule_tac h = i in that) apply (rule \i \ F\) apply (rule subset_of_face_of [OF face \c \ S\]) using \a i \ x = b i\ \x \ rel_interior c\ \x \ S\ apply blast done qed text\Initial part of proof duplicates that above\ proposition face_of_polyhedron_explicit: fixes S :: "'a :: euclidean_space set" assumes "finite F" and seq: "S = affine hull S \ \F" and faceq: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" and psub: "\F'. F' \ F \ S \ affine hull S \ \F'" and c: "c face_of S" and "c \ {}" "c \ S" shows "c = \{S \ {x. a h \ x = b h} | h. h \ F \ c \ S \ {x. a h \ x = b h}}" proof - let ?ab = "\h. {x. a h \ x = b h}" have "c \ S" using \c face_of S\ by (simp add: face_of_imp_subset) have "polyhedron S" apply (simp add: polyhedron_Int_affine) by (metis \finite F\ faceq seq) then have "convex S" by (simp add: polyhedron_imp_convex) then have *: "(S \ ?ab h) face_of S" if "h \ F" for h apply (rule face_of_Int_supporting_hyperplane_le) using faceq seq that by fastforce have "rel_interior c \ {}" using c \c \ {}\ face_of_imp_convex rel_interior_eq_empty by blast then obtain z where z: "z \ rel_interior c" by auto have rels: "rel_interior S = {z \ S. \h\F. a h \ z < b h}" by (rule rel_interior_polyhedron_explicit [OF \finite F\ seq faceq psub]) then have xnot: "z \ rel_interior S" by (metis IntI \z \ rel_interior c\ c \c \ S\ contra_subsetD empty_iff face_of_disjoint_rel_interior rel_interior_subset) then have "z \ S" using \c \ S\ \z \ rel_interior c\ rel_interior_subset by auto with seq have xint: "z \ \F" by blast have "open (\h\{h \ F. a h \ z < b h}. {w. a h \ w < b h})" by (auto simp: \finite F\ open_halfspace_lt open_INT) then obtain e where "0 < e" "ball z e \ (\h\{h \ F. a h \ z < b h}. {w. a h \ w < b h})" by (auto intro: openE [of _ z]) then have e: "\h. \h \ F; a h \ z < b h\ \ ball z e \ {w. a h \ w < b h}" by blast have "c \ (S \ ?ab h) \ z \ S \ ?ab h" if "h \ F" for h proof show "z \ S \ ?ab h \ c \ S \ ?ab h" apply (rule subset_of_face_of [of _ S]) using that \c \ S\ \z \ rel_interior c\ using facet_of_polyhedron_explicit [OF \finite F\ seq faceq psub] unfolding facet_of_def apply auto done next show "c \ S \ ?ab h \ z \ S \ ?ab h" using \z \ rel_interior c\ rel_interior_subset by force qed then have **: "{S \ ?ab h | h. h \ F \ c \ S \ c \ ?ab h} = {S \ ?ab h |h. h \ F \ z \ S \ ?ab h}" by blast have bsub: "ball z e \ affine hull \{S \ ?ab h |h. h \ F \ a h \ z = b h} \ affine hull S \ \F \ \{?ab h |h. h \ F \ a h \ z = b h}" if "i \ F" and i: "a i \ z = b i" for i proof - have sub: "ball z e \ \{?ab h |h. h \ F \ a h \ z = b h} \ j" if "j \ F" for j proof - have "a j \ z \ b j" using faceq that xint by auto then consider "a j \ z < b j" | "a j \ z = b j" by linarith then have "\G. G \ {?ab h |h. h \ F \ a h \ z = b h} \ ball z e \ G \ j" proof cases assume "a j \ z < b j" then have "ball z e \ {x. a i \ x = b i} \ j" using e [OF \j \ F\] faceq that by (fastforce simp: ball_def) then show ?thesis by (rule_tac x="{x. a i \ x = b i}" in exI) (force simp: \i \ F\ i) next assume eq: "a j \ z = b j" with faceq that show ?thesis by (rule_tac x="{x. a j \ x = b j}" in exI) (fastforce simp add: \j \ F\) qed then show ?thesis by blast qed have 1: "affine hull \{S \ ?ab h |h. h \ F \ a h \ z = b h} \ affine hull S" apply (rule hull_mono) using that \z \ S\ by auto have 2: "affine hull \{S \ ?ab h |h. h \ F \ a h \ z = b h} \ \{?ab h |h. h \ F \ a h \ z = b h}" by (rule hull_minimal) (auto intro: affine_hyperplane) have 3: "ball z e \ \{?ab h |h. h \ F \ a h \ z = b h} \ \F" by (iprover intro: sub Inter_greatest) have *: "\A \ (B :: 'a set); A \ C; E \ C \ D\ \ E \ A \ (B \ D) \ C" for A B C D E by blast show ?thesis by (intro * 1 2 3) qed have "\h. h \ F \ c \ ?ab h" apply (rule face_of_polyhedron_subset_explicit [OF \finite F\ seq faceq psub]) using assms by auto then have fac: "\{S \ ?ab h |h. h \ F \ c \ S \ ?ab h} face_of S" using * by (force simp: \c \ S\ intro: face_of_Inter) have red: "(\a. P a \ T \ S \ \{F x |x. P x}) \ T \ \{S \ F x |x. P x}" for P T F by blast have "ball z e \ affine hull \{S \ ?ab h |h. h \ F \ a h \ z = b h} \ \{S \ ?ab h |h. h \ F \ a h \ z = b h}" apply (rule red) apply (metis seq bsub) done with \0 < e\ have zinrel: "z \ rel_interior (\{S \ ?ab h |h. h \ F \ z \ S \ a h \ z = b h})" by (auto simp: mem_rel_interior_ball \z \ S\) show ?thesis apply (rule face_of_eq [OF c fac]) using z zinrel apply (force simp: **) done qed subsection\More general corollaries from the explicit representation\ corollary facet_of_polyhedron: assumes "polyhedron S" and "c facet_of S" obtains a b where "a \ 0" "S \ {x. a \ x \ b}" "c = S \ {x. a \ x = b}" proof - obtain F where "finite F" and seq: "S = affine hull S \ \F" and faces: "\h. h \ F \ \a b. a \ 0 \ h = {x. a \ x \ b}" and min: "\F'. F' \ F \ S \ (affine hull S) \ \F'" using assms by (simp add: polyhedron_Int_affine_minimal) meson then obtain a b where ab: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" by metis obtain i where "i \ F" and c: "c = S \ {x. a i \ x = b i}" using facet_of_polyhedron_explicit [OF \finite F\ seq ab min] assms by force moreover have ssub: "S \ {x. a i \ x \ b i}" apply (subst seq) using \i \ F\ ab by auto ultimately show ?thesis by (rule_tac a = "a i" and b = "b i" in that) (simp_all add: ab) qed corollary face_of_polyhedron: assumes "polyhedron S" and "c face_of S" and "c \ {}" and "c \ S" shows "c = \{F. F facet_of S \ c \ F}" proof - obtain F where "finite F" and seq: "S = affine hull S \ \F" and faces: "\h. h \ F \ \a b. a \ 0 \ h = {x. a \ x \ b}" and min: "\F'. F' \ F \ S \ (affine hull S) \ \F'" using assms by (simp add: polyhedron_Int_affine_minimal) meson then obtain a b where ab: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" by metis show ?thesis apply (subst face_of_polyhedron_explicit [OF \finite F\ seq ab min]) apply (auto simp: assms facet_of_polyhedron_explicit [OF \finite F\ seq ab min] cong: Collect_cong) done qed lemma face_of_polyhedron_subset_facet: assumes "polyhedron S" and "c face_of S" and "c \ {}" and "c \ S" obtains F where "F facet_of S" "c \ F" using face_of_polyhedron assms by (metis (no_types, lifting) Inf_greatest antisym_conv face_of_imp_subset mem_Collect_eq) lemma exposed_face_of_polyhedron: assumes "polyhedron S" shows "F exposed_face_of S \ F face_of S" proof show "F exposed_face_of S \ F face_of S" by (simp add: exposed_face_of_def) next assume "F face_of S" show "F exposed_face_of S" proof (cases "F = {} \ F = S") case True then show ?thesis using \F face_of S\ exposed_face_of by blast next case False then have "{g. g facet_of S \ F \ g} \ {}" by (metis Collect_empty_eq_bot \F face_of S\ assms empty_def face_of_polyhedron_subset_facet) moreover have "\T. \T facet_of S; F \ T\ \ T exposed_face_of S" by (metis assms exposed_face_of facet_of_imp_face_of facet_of_polyhedron) ultimately have "\{fa. fa facet_of S \ F \ fa} exposed_face_of S" by (metis (no_types, lifting) mem_Collect_eq exposed_face_of_Inter) then show ?thesis using False apply (subst face_of_polyhedron [OF assms \F face_of S\], auto) done qed qed lemma face_of_polyhedron_polyhedron: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" "c face_of S" shows "polyhedron c" by (metis assms face_of_imp_eq_affine_Int polyhedron_Int polyhedron_affine_hull polyhedron_imp_convex) lemma finite_polyhedron_faces: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" shows "finite {F. F face_of S}" proof - obtain F where "finite F" and seq: "S = affine hull S \ \F" and faces: "\h. h \ F \ \a b. a \ 0 \ h = {x. a \ x \ b}" and min: "\F'. F' \ F \ S \ (affine hull S) \ \F'" using assms by (simp add: polyhedron_Int_affine_minimal) meson then obtain a b where ab: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" by metis have "finite {\{S \ {x. a h \ x = b h} |h. h \ F'}| F'. F' \ Pow F}" by (simp add: \finite F\) moreover have "{F. F face_of S} - {{}, S} \ {\{S \ {x. a h \ x = b h} |h. h \ F'}| F'. F' \ Pow F}" apply clarify apply (rename_tac c) apply (drule face_of_polyhedron_explicit [OF \finite F\ seq ab min, simplified], simp_all) apply (erule ssubst) apply (rule_tac x="{h \ F. c \ S \ {x. a h \ x = b h}}" in exI, auto) done ultimately show ?thesis by (meson finite.emptyI finite.insertI finite_Diff2 finite_subset) qed lemma finite_polyhedron_exposed_faces: "polyhedron S \ finite {F. F exposed_face_of S}" using exposed_face_of_polyhedron finite_polyhedron_faces by fastforce lemma finite_polyhedron_extreme_points: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ finite {v. v extreme_point_of S}" apply (simp add: face_of_singleton [symmetric]) apply (rule finite_subset [OF _ finite_vimageI [OF finite_polyhedron_faces]], auto) done lemma finite_polyhedron_facets: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ finite {F. F facet_of S}" unfolding facet_of_def by (blast intro: finite_subset [OF _ finite_polyhedron_faces]) proposition rel_interior_of_polyhedron: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" shows "rel_interior S = S - \{F. F facet_of S}" proof - obtain F where "finite F" and seq: "S = affine hull S \ \F" and faces: "\h. h \ F \ \a b. a \ 0 \ h = {x. a \ x \ b}" and min: "\F'. F' \ F \ S \ (affine hull S) \ \F'" using assms by (simp add: polyhedron_Int_affine_minimal) meson then obtain a b where ab: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" by metis have facet: "(c facet_of S) \ (\h. h \ F \ c = S \ {x. a h \ x = b h})" for c by (rule facet_of_polyhedron_explicit [OF \finite F\ seq ab min]) have rel: "rel_interior S = {x \ S. \h\F. a h \ x < b h}" by (rule rel_interior_polyhedron_explicit [OF \finite F\ seq ab min]) have "a h \ x < b h" if "x \ S" "h \ F" and xnot: "x \ \{F. F facet_of S}" for x h proof - have "x \ \F" using seq that by force with \h \ F\ ab have "a h \ x \ b h" by auto then consider "a h \ x < b h" | "a h \ x = b h" by linarith then show ?thesis proof cases case 1 then show ?thesis . next case 2 have "Collect ((\) x) \ Collect ((\) (\{A. A facet_of S}))" using xnot by fastforce then have "F \ Collect ((\) h)" using 2 \x \ S\ facet by blast with \h \ F\ have "\F \ S \ {x. a h \ x = b h}" by blast with 2 that \x \ \F\ show ?thesis apply simp apply (drule_tac x="\F" in spec) apply (simp add: facet) apply (drule_tac x=h in spec) using seq by auto qed qed moreover have "\h\F. a h \ x \ b h" if "x \ \{F. F facet_of S}" for x using that by (force simp: facet) ultimately show ?thesis by (force simp: rel) qed lemma rel_boundary_of_polyhedron: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" shows "S - rel_interior S = \ {F. F facet_of S}" using facet_of_imp_subset by (fastforce simp add: rel_interior_of_polyhedron assms) lemma rel_frontier_of_polyhedron: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" shows "rel_frontier S = \ {F. F facet_of S}" by (simp add: assms rel_frontier_def polyhedron_imp_closed rel_boundary_of_polyhedron) lemma rel_frontier_of_polyhedron_alt: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" shows "rel_frontier S = \ {F. F face_of S \ (F \ S)}" apply (rule subset_antisym) apply (force simp: rel_frontier_of_polyhedron facet_of_def assms) using face_of_subset_rel_frontier by fastforce text\A characterization of polyhedra as having finitely many faces\ proposition polyhedron_eq_finite_exposed_faces: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ closed S \ convex S \ finite {F. F exposed_face_of S}" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (auto simp: polyhedron_imp_closed polyhedron_imp_convex finite_polyhedron_exposed_faces) next assume ?rhs then have "closed S" "convex S" and fin: "finite {F. F exposed_face_of S}" by auto show ?lhs proof (cases "S = {}") case True then show ?thesis by auto next case False define F where "F = {h. h exposed_face_of S \ h \ {} \ h \ S}" have "finite F" by (simp add: fin F_def) have hface: "h face_of S" and "\a b. a \ 0 \ S \ {x. a \ x \ b} \ h = S \ {x. a \ x = b}" if "h \ F" for h - using exposed_face_of F_def that by simp_all auto + using exposed_face_of F_def that by blast+ then obtain a b where ab: "\h. h \ F \ a h \ 0 \ S \ {x. a h \ x \ b h} \ h = S \ {x. a h \ x = b h}" by metis have *: "False" if paff: "p \ affine hull S" and "p \ S" and pint: "p \ \{{x. a h \ x \ b h} |h. h \ F}" for p proof - have "rel_interior S \ {}" by (simp add: \S \ {}\ \convex S\ rel_interior_eq_empty) then obtain c where c: "c \ rel_interior S" by auto with rel_interior_subset have "c \ S" by blast have ccp: "closed_segment c p \ affine hull S" by (meson affine_affine_hull affine_imp_convex c closed_segment_subset hull_subset paff rel_interior_subset subsetCE) obtain x where xcl: "x \ closed_segment c p" and "x \ S" and xnot: "x \ rel_interior S" using connected_openin [of "closed_segment c p"] apply simp apply (drule_tac x="closed_segment c p \ rel_interior S" in spec) apply (erule impE) apply (force simp: openin_rel_interior openin_Int intro: openin_subtopology_Int_subset [OF _ ccp]) apply (drule_tac x="closed_segment c p \ (- S)" in spec) using rel_interior_subset \closed S\ c \p \ S\ apply blast done then obtain \ where "0 \ \" "\ \ 1" and xeq: "x = (1 - \) *\<^sub>R c + \ *\<^sub>R p" by (auto simp: in_segment) show False proof (cases "\=0 \ \=1") case True with xeq c xnot \x \ S\ \p \ S\ show False by auto next case False then have xos: "x \ open_segment c p" using \x \ S\ c open_segment_def that(2) xcl xnot by auto have xclo: "x \ closure S" using \x \ S\ closure_subset by blast obtain d where "d \ 0" and dle: "\y. y \ closure S \ d \ x \ d \ y" and dless: "\y. y \ rel_interior S \ d \ x < d \ y" by (metis supporting_hyperplane_relative_frontier [OF \convex S\ xclo xnot]) have sex: "S \ {y. d \ y = d \ x} exposed_face_of S" by (simp add: \closed S\ dle exposed_face_of_Int_supporting_hyperplane_ge [OF \convex S\]) have sne: "S \ {y. d \ y = d \ x} \ {}" using \x \ S\ by blast have sns: "S \ {y. d \ y = d \ x} \ S" by (metis (mono_tags) Int_Collect c subsetD dless not_le order_refl rel_interior_subset) obtain h where "h \ F" "x \ h" apply (rule_tac h="S \ {y. d \ y = d \ x}" in that) apply (simp_all add: F_def sex sne sns \x \ S\) done have abface: "{y. a h \ y = b h} face_of {y. a h \ y \ b h}" using hyperplane_face_of_halfspace_le by blast then have "c \ h" using face_ofD [OF abface xos] \c \ S\ \h \ F\ ab pint \x \ h\ by blast with c have "h \ rel_interior S \ {}" by blast then show False using \h \ F\ F_def face_of_disjoint_rel_interior hface by auto qed qed have "S \ affine hull S \ \{{x. a h \ x \ b h} |h. h \ F}" using ab by (auto simp: hull_subset) moreover have "affine hull S \ \{{x. a h \ x \ b h} |h. h \ F} \ S" using * by blast ultimately have "S = affine hull S \ \ {{x. a h \ x \ b h} |h. h \ F}" .. then show ?thesis apply (rule ssubst) apply (force intro: polyhedron_affine_hull polyhedron_halfspace_le simp: \finite F\) done qed qed corollary polyhedron_eq_finite_faces: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ closed S \ convex S \ finite {F. F face_of S}" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (simp add: finite_polyhedron_faces polyhedron_imp_closed polyhedron_imp_convex) next assume ?rhs then show ?lhs by (force simp: polyhedron_eq_finite_exposed_faces exposed_face_of intro: finite_subset) qed lemma polyhedron_linear_image_eq: fixes h :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "linear h" "bij h" shows "polyhedron (h ` S) \ polyhedron S" proof - have *: "{f. P f} = (image h) ` {f. P (h ` f)}" for P apply safe apply (rule_tac x="inv h ` x" in image_eqI) apply (auto simp: \bij h\ bij_is_surj image_f_inv_f) done have "inj h" using bij_is_inj assms by blast then have injim: "inj_on ((`) h) A" for A by (simp add: inj_on_def inj_image_eq_iff) show ?thesis using \linear h\ \inj h\ apply (simp add: polyhedron_eq_finite_faces closed_injective_linear_image_eq) apply (simp add: * face_of_linear_image [of h _ S, symmetric] finite_image_iff injim) done qed lemma polyhedron_negations: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ polyhedron(image uminus S)" by (subst polyhedron_linear_image_eq) (auto simp: bij_uminus intro!: linear_uminus) subsection\Relation between polytopes and polyhedra\ proposition polytope_eq_bounded_polyhedron: fixes S :: "'a :: euclidean_space set" shows "polytope S \ polyhedron S \ bounded S" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (simp add: finite_polytope_faces polyhedron_eq_finite_faces polytope_imp_closed polytope_imp_convex polytope_imp_bounded) next assume ?rhs then show ?lhs unfolding polytope_def apply (rule_tac x="{v. v extreme_point_of S}" in exI) apply (simp add: finite_polyhedron_extreme_points Krein_Milman_Minkowski compact_eq_bounded_closed polyhedron_imp_closed polyhedron_imp_convex) done qed lemma polytope_Int: fixes S :: "'a :: euclidean_space set" shows "\polytope S; polytope T\ \ polytope(S \ T)" by (simp add: polytope_eq_bounded_polyhedron bounded_Int) lemma polytope_Int_polyhedron: fixes S :: "'a :: euclidean_space set" shows "\polytope S; polyhedron T\ \ polytope(S \ T)" by (simp add: bounded_Int polytope_eq_bounded_polyhedron) lemma polyhedron_Int_polytope: fixes S :: "'a :: euclidean_space set" shows "\polyhedron S; polytope T\ \ polytope(S \ T)" by (simp add: bounded_Int polytope_eq_bounded_polyhedron) lemma polytope_imp_polyhedron: fixes S :: "'a :: euclidean_space set" shows "polytope S \ polyhedron S" by (simp add: polytope_eq_bounded_polyhedron) lemma polytope_facet_exists: fixes p :: "'a :: euclidean_space set" assumes "polytope p" "0 < aff_dim p" obtains F where "F facet_of p" proof (cases "p = {}") case True with assms show ?thesis by auto next case False then obtain v where "v extreme_point_of p" using extreme_point_exists_convex by (blast intro: \polytope p\ polytope_imp_compact polytope_imp_convex) then show ?thesis by (metis face_of_polyhedron_subset_facet polytope_imp_polyhedron aff_dim_sing all_not_in_conv assms face_of_singleton less_irrefl singletonI that) qed lemma polyhedron_interval [iff]: "polyhedron(cbox a b)" by (metis polytope_imp_polyhedron polytope_interval) lemma polyhedron_convex_hull: fixes S :: "'a :: euclidean_space set" shows "finite S \ polyhedron(convex hull S)" by (simp add: polytope_convex_hull polytope_imp_polyhedron) subsection\Relative and absolute frontier of a polytope\ lemma rel_boundary_of_convex_hull: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "(convex hull S) - rel_interior(convex hull S) = (\a\S. convex hull (S - {a}))" proof - have "finite S" by (metis assms aff_independent_finite) then consider "card S = 0" | "card S = 1" | "2 \ card S" by arith then show ?thesis proof cases case 1 then have "S = {}" by (simp add: \finite S\) then show ?thesis by simp next case 2 show ?thesis by (auto intro: card_1_singletonE [OF \card S = 1\]) next case 3 with assms show ?thesis by (auto simp: polyhedron_convex_hull rel_boundary_of_polyhedron facet_of_convex_hull_affine_independent_alt \finite S\) qed qed proposition frontier_of_convex_hull: fixes S :: "'a::euclidean_space set" assumes "card S = Suc (DIM('a))" shows "frontier(convex hull S) = \ {convex hull (S - {a}) | a. a \ S}" proof (cases "affine_dependent S") case True have [iff]: "finite S" using assms using card_infinite by force then have ccs: "closed (convex hull S)" by (simp add: compact_imp_closed finite_imp_compact_convex_hull) { fix x T assume "finite T" "T \ S" "int (card T) \ aff_dim S + 1" "x \ convex hull T" then have "S \ T" using True \finite S\ aff_dim_le_card affine_independent_iff_card by fastforce then obtain a where "a \ S" "a \ T" using \T \ S\ by blast then have "x \ (\a\S. convex hull (S - {a}))" using True affine_independent_iff_card [of S] apply simp apply (metis (no_types, hide_lams) Diff_eq_empty_iff Diff_insert0 \a \ T\ \T \ S\ \x \ convex hull T\ hull_mono insert_Diff_single subsetCE) done } note * = this have 1: "convex hull S \ (\ a\S. convex hull (S - {a}))" apply (subst caratheodory_aff_dim) apply (blast intro: *) done have 2: "\((\a. convex hull (S - {a})) ` S) \ convex hull S" by (rule Union_least) (metis (no_types, lifting) Diff_subset hull_mono imageE) show ?thesis using True apply (simp add: segment_convex_hull frontier_def) using interior_convex_hull_eq_empty [OF assms] apply (simp add: closure_closed [OF ccs]) apply (rule subset_antisym) using 1 apply blast using 2 apply blast done next case False then have "frontier (convex hull S) = (convex hull S) - rel_interior(convex hull S)" apply (simp add: rel_boundary_of_convex_hull [symmetric] frontier_def) by (metis aff_independent_finite assms closure_convex_hull finite_imp_compact_convex_hull hull_hull interior_convex_hull_eq_empty rel_interior_nonempty_interior) also have "... = \{convex hull (S - {a}) |a. a \ S}" proof - have "convex hull S - rel_interior (convex hull S) = rel_frontier (convex hull S)" by (simp add: False aff_independent_finite polyhedron_convex_hull rel_boundary_of_polyhedron rel_frontier_of_polyhedron) then show ?thesis by (simp add: False rel_frontier_convex_hull_cases) qed finally show ?thesis . qed subsection\Special case of a triangle\ proposition frontier_of_triangle: fixes a :: "'a::euclidean_space" assumes "DIM('a) = 2" shows "frontier(convex hull {a,b,c}) = closed_segment a b \ closed_segment b c \ closed_segment c a" (is "?lhs = ?rhs") proof (cases "b = a \ c = a \ c = b") case True then show ?thesis by (auto simp: assms segment_convex_hull frontier_def empty_interior_convex_hull insert_commute card_insert_le_m1 hull_inc insert_absorb) next case False then have [simp]: "card {a, b, c} = Suc (DIM('a))" by (simp add: card_insert Set.insert_Diff_if assms) show ?thesis proof show "?lhs \ ?rhs" using False by (force simp: segment_convex_hull frontier_of_convex_hull insert_Diff_if insert_commute split: if_split_asm) show "?rhs \ ?lhs" using False apply (simp add: frontier_of_convex_hull segment_convex_hull) apply (intro conjI subsetI) apply (rule_tac X="convex hull {a,b}" in UnionI; force simp: Set.insert_Diff_if) apply (rule_tac X="convex hull {b,c}" in UnionI; force) apply (rule_tac X="convex hull {a,c}" in UnionI; force simp: insert_commute Set.insert_Diff_if) done qed qed corollary inside_of_triangle: fixes a :: "'a::euclidean_space" assumes "DIM('a) = 2" shows "inside (closed_segment a b \ closed_segment b c \ closed_segment c a) = interior(convex hull {a,b,c})" by (metis assms frontier_of_triangle bounded_empty bounded_insert convex_convex_hull inside_frontier_eq_interior bounded_convex_hull) corollary interior_of_triangle: fixes a :: "'a::euclidean_space" assumes "DIM('a) = 2" shows "interior(convex hull {a,b,c}) = convex hull {a,b,c} - (closed_segment a b \ closed_segment b c \ closed_segment c a)" using interior_subset by (force simp: frontier_of_triangle [OF assms, symmetric] frontier_def Diff_Diff_Int) subsection\Subdividing a cell complex\ lemma subdivide_interval: fixes x::real assumes "a < \x - y\" "0 < a" obtains n where "n \ \" "x < n * a \ n * a < y \ y < n * a \ n * a < x" proof - consider "a + x < y" | "a + y < x" using assms by linarith then show ?thesis proof cases case 1 let ?n = "of_int (floor (x/a)) + 1" have x: "x < ?n * a" by (meson \0 < a\ divide_less_eq floor_eq_iff) have "?n * a \ a + x" apply (simp add: algebra_simps) by (metis \0 < a\ floor_correct less_irrefl nonzero_mult_div_cancel_left real_mult_le_cancel_iff2 times_divide_eq_right) also have "... < y" by (rule 1) finally have "?n * a < y" . with x show ?thesis using Ints_1 Ints_add Ints_of_int that by blast next case 2 let ?n = "of_int (floor (y/a)) + 1" have y: "y < ?n * a" by (meson \0 < a\ divide_less_eq floor_eq_iff) have "?n * a \ a + y" apply (simp add: algebra_simps) by (metis \0 < a\ floor_correct less_irrefl nonzero_mult_div_cancel_left real_mult_le_cancel_iff2 times_divide_eq_right) also have "... < x" by (rule 2) finally have "?n * a < x" . then show ?thesis using Ints_1 Ints_add Ints_of_int that y by blast qed qed lemma cell_subdivision_lemma: assumes "finite \" and "\X. X \ \ \ polytope X" and "\X. X \ \ \ aff_dim X \ d" and "\X Y. \X \ \; Y \ \\ \ (X \ Y) face_of X" and "finite I" shows "\\. \\ = \\ \ finite \ \ (\C \ \. \D. D \ \ \ C \ D) \ (\C \ \. \x \ C. \D. D \ \ \ x \ D \ D \ C) \ (\X \ \. polytope X) \ (\X \ \. aff_dim X \ d) \ (\X \ \. \Y \ \. X \ Y face_of X) \ (\X \ \. \x \ X. \y \ X. \a b. (a,b) \ I \ a \ x \ b \ a \ y \ b \ a \ x \ b \ a \ y \ b)" using \finite I\ proof induction case empty then show ?case by (rule_tac x="\" in exI) (auto simp: assms) next case (insert ab I) then obtain \ where eq: "\\ = \\" and "finite \" and sub1: "\C. C \ \ \ \D. D \ \ \ C \ D" and sub2: "\C x. C \ \ \ x \ C \ \D. D \ \ \ x \ D \ D \ C" and poly: "\X. X \ \ \ polytope X" and aff: "\X. X \ \ \ aff_dim X \ d" and face: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X" and I: "\X x y a b. \X \ \; x \ X; y \ X; (a,b) \ I\ \ a \ x \ b \ a \ y \ b \ a \ x \ b \ a \ y \ b" by (auto simp: that) obtain a b where "ab = (a,b)" by fastforce let ?\ = "(\X. X \ {x. a \ x \ b}) ` \ \ (\X. X \ {x. a \ x \ b}) ` \" have eqInt: "(S \ Collect P) \ (T \ Collect Q) = (S \ T) \ (Collect P \ Collect Q)" for S T::"'a set" and P Q by blast show ?case proof (intro conjI exI) show "\?\ = \\" by (force simp: eq [symmetric]) show "finite ?\" using \finite \\ by force show "\X \ ?\. polytope X" by (force simp: poly polytope_Int_polyhedron polyhedron_halfspace_le polyhedron_halfspace_ge) show "\X \ ?\. aff_dim X \ d" by (auto; metis order_trans aff aff_dim_subset inf_le1) show "\X \ ?\. \x \ X. \y \ X. \a b. (a,b) \ insert ab I \ a \ x \ b \ a \ y \ b \ a \ x \ b \ a \ y \ b" using \ab = (a, b)\ I by fastforce show "\X \ ?\. \Y \ ?\. X \ Y face_of X" by (auto simp: eqInt halfspace_Int_eq face_of_Int_Int face face_of_halfspace_le face_of_halfspace_ge) show "\C \ ?\. \D. D \ \ \ C \ D" using sub1 by force show "\C\\. \x\C. \D. D \ ?\ \ x \ D \ D \ C" proof (intro ballI) fix C z assume "C \ \" "z \ C" with sub2 obtain D where D: "D \ \" "z \ D" "D \ C" by blast have "D \ \ \ z \ D \ {x. a \ x \ b} \ D \ {x. a \ x \ b} \ C \ D \ \ \ z \ D \ {x. a \ x \ b} \ D \ {x. a \ x \ b} \ C" using linorder_class.linear [of "a \ z" b] D by blast then show "\D. D \ ?\ \ z \ D \ D \ C" by blast qed qed qed proposition cell_complex_subdivision_exists: fixes \ :: "'a::euclidean_space set set" assumes "0 < e" "finite \" and poly: "\X. X \ \ \ polytope X" and aff: "\X. X \ \ \ aff_dim X \ d" and face: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X" obtains "\'" where "finite \'" "\\' = \\" "\X. X \ \' \ diameter X < e" "\X. X \ \' \ polytope X" "\X. X \ \' \ aff_dim X \ d" "\X Y. \X \ \'; Y \ \'\ \ X \ Y face_of X" "\C. C \ \' \ \D. D \ \ \ C \ D" "\C x. C \ \ \ x \ C \ \D. D \ \' \ x \ D \ D \ C" proof - have "bounded(\\)" by (simp add: \finite \\ poly bounded_Union polytope_imp_bounded) then obtain B where "B > 0" and B: "\x. x \ \\ \ norm x < B" by (meson bounded_pos_less) define C where "C \ {z \ \. \z * e / 2 / real DIM('a)\ \ B}" define I where "I \ \i \ Basis. \j \ C. { (i::'a, j * e / 2 / DIM('a)) }" have "finite C" using finite_int_segment [of "-B / (e / 2 / DIM('a))" "B / (e / 2 / DIM('a))"] apply (simp add: C_def) apply (erule rev_finite_subset) using \0 < e\ apply (auto simp: field_split_simps) done then have "finite I" by (simp add: I_def) obtain \' where eq: "\\' = \\" and "finite \'" and poly: "\X. X \ \' \ polytope X" and aff: "\X. X \ \' \ aff_dim X \ d" and face: "\X Y. \X \ \'; Y \ \'\ \ X \ Y face_of X" and I: "\X x y a b. \X \ \'; x \ X; y \ X; (a,b) \ I\ \ a \ x \ b \ a \ y \ b \ a \ x \ b \ a \ y \ b" and sub1: "\C. C \ \' \ \D. D \ \ \ C \ D" and sub2: "\C x. C \ \ \ x \ C \ \D. D \ \' \ x \ D \ D \ C" apply (rule exE [OF cell_subdivision_lemma]) using assms \finite I\ apply auto done show ?thesis proof (rule_tac \'="\'" in that) show "diameter X < e" if "X \ \'" for X proof - have "diameter X \ e/2" proof (rule diameter_le) show "norm (x - y) \ e / 2" if "x \ X" "y \ X" for x y proof - have "norm x < B" "norm y < B" - using B \X \ \'\ eq that by fastforce+ + using B \X \ \'\ eq that by blast+ have "norm (x - y) \ (\b\Basis. \(x-y) \ b\)" by (rule norm_le_l1) also have "... \ of_nat (DIM('a)) * (e / 2 / DIM('a))" proof (rule sum_bounded_above) fix i::'a assume "i \ Basis" then have I': "\z b. \z \ C; b = z * e / (2 * real DIM('a))\ \ i \ x \ b \ i \ y \ b \ i \ x \ b \ i \ y \ b" - using I \X \ \'\ that - by (fastforce simp: I_def) + using I[of X x y] \X \ \'\ that unfolding I_def by auto show "\(x - y) \ i\ \ e / 2 / real DIM('a)" proof (rule ccontr) assume "\ \(x - y) \ i\ \ e / 2 / real DIM('a)" then have xyi: "\i \ x - i \ y\ > e / 2 / real DIM('a)" by (simp add: inner_commute inner_diff_right) obtain n where "n \ \" and n: "i \ x < n * (e / 2 / real DIM('a)) \ n * (e / 2 / real DIM('a)) < i \ y \ i \ y < n * (e / 2 / real DIM('a)) \ n * (e / 2 / real DIM('a)) < i \ x" using subdivide_interval [OF xyi] DIM_positive \0 < e\ by (auto simp: zero_less_divide_iff) have "\i \ x\ < B" by (metis \i \ Basis\ \norm x < B\ inner_commute norm_bound_Basis_lt) have "\i \ y\ < B" by (metis \i \ Basis\ \norm y < B\ inner_commute norm_bound_Basis_lt) have *: "\n * e\ \ B * (2 * real DIM('a))" if "\ix\ < B" "\iy\ < B" and ix: "ix * (2 * real DIM('a)) < n * e" and iy: "n * e < iy * (2 * real DIM('a))" for ix iy proof (rule abs_leI) have "iy * (2 * real DIM('a)) \ B * (2 * real DIM('a))" by (rule mult_right_mono) (use \\iy\ < B\ in linarith)+ then show "n * e \ B * (2 * real DIM('a))" using iy by linarith next have "- ix * (2 * real DIM('a)) \ B * (2 * real DIM('a))" by (rule mult_right_mono) (use \\ix\ < B\ in linarith)+ then show "- (n * e) \ B * (2 * real DIM('a))" using ix by linarith qed have "n \ C" using \n \ \\ n by (auto simp: C_def divide_simps intro: * \\i \ x\ < B\ \\i \ y\ < B\) show False using I' [OF \n \ C\ refl] n by auto qed qed also have "... = e / 2" by simp finally show ?thesis . qed qed (use \0 < e\ in force) also have "... < e" by (simp add: \0 < e\) finally show ?thesis . qed qed (auto simp: eq poly aff face sub1 sub2 \finite \'\) qed subsection\Simplexes\ text\The notion of n-simplex for integer \<^term>\n \ -1\\ definition\<^marker>\tag important\ simplex :: "int \ 'a::euclidean_space set \ bool" (infix "simplex" 50) where "n simplex S \ \C. \ affine_dependent C \ int(card C) = n + 1 \ S = convex hull C" lemma simplex: "n simplex S \ (\C. finite C \ \ affine_dependent C \ int(card C) = n + 1 \ S = convex hull C)" by (auto simp add: simplex_def intro: aff_independent_finite) lemma simplex_convex_hull: "\ affine_dependent C \ int(card C) = n + 1 \ n simplex (convex hull C)" by (auto simp add: simplex_def) lemma convex_simplex: "n simplex S \ convex S" by (metis convex_convex_hull simplex_def) lemma compact_simplex: "n simplex S \ compact S" unfolding simplex using finite_imp_compact_convex_hull by blast lemma closed_simplex: "n simplex S \ closed S" by (simp add: compact_imp_closed compact_simplex) lemma simplex_imp_polytope: "n simplex S \ polytope S" unfolding simplex_def polytope_def using aff_independent_finite by blast lemma simplex_imp_polyhedron: "n simplex S \ polyhedron S" by (simp add: polytope_imp_polyhedron simplex_imp_polytope) lemma simplex_dim_ge: "n simplex S \ -1 \ n" by (metis (no_types, hide_lams) aff_dim_geq affine_independent_iff_card diff_add_cancel diff_diff_eq2 simplex_def) lemma simplex_empty [simp]: "n simplex {} \ n = -1" proof assume "n simplex {}" then show "n = -1" unfolding simplex by (metis card_empty convex_hull_eq_empty diff_0 diff_eq_eq of_nat_0) next assume "n = -1" then show "n simplex {}" by (fastforce simp: simplex) qed lemma simplex_minus_1 [simp]: "-1 simplex S \ S = {}" by (metis simplex cancel_comm_monoid_add_class.diff_cancel card_0_eq diff_minus_eq_add of_nat_eq_0_iff simplex_empty) lemma aff_dim_simplex: "n simplex S \ aff_dim S = n" by (metis simplex add.commute add_diff_cancel_left' aff_dim_convex_hull affine_independent_iff_card) lemma zero_simplex_sing: "0 simplex {a}" apply (simp add: simplex_def) by (metis affine_independent_1 card_empty card_insert_disjoint convex_hull_singleton empty_iff finite.emptyI) lemma simplex_sing [simp]: "n simplex {a} \ n = 0" using aff_dim_simplex aff_dim_sing zero_simplex_sing by blast lemma simplex_zero: "0 simplex S \ (\a. S = {a})" apply (auto simp: ) using aff_dim_eq_0 aff_dim_simplex by blast lemma one_simplex_segment: "a \ b \ 1 simplex closed_segment a b" apply (simp add: simplex_def) apply (rule_tac x="{a,b}" in exI) apply (auto simp: segment_convex_hull) done lemma simplex_segment_cases: "(if a = b then 0 else 1) simplex closed_segment a b" by (auto simp: one_simplex_segment) lemma simplex_segment: "\n. n simplex closed_segment a b" using simplex_segment_cases by metis lemma polytope_lowdim_imp_simplex: assumes "polytope P" "aff_dim P \ 1" obtains n where "n simplex P" proof (cases "P = {}") case True then show ?thesis by (simp add: that) next case False then show ?thesis by (metis assms compact_convex_collinear_segment collinear_aff_dim polytope_imp_compact polytope_imp_convex simplex_segment_cases that) qed lemma simplex_insert_dimplus1: fixes n::int assumes "n simplex S" and a: "a \ affine hull S" shows "(n+1) simplex (convex hull (insert a S))" proof - obtain C where C: "finite C" "\ affine_dependent C" "int(card C) = n+1" and S: "S = convex hull C" using assms unfolding simplex by force show ?thesis unfolding simplex proof (intro exI conjI) have "aff_dim S = n" using aff_dim_simplex assms(1) by blast moreover have "a \ affine hull C" using S a affine_hull_convex_hull by blast moreover have "a \ C" using S a hull_inc by fastforce ultimately show "\ affine_dependent (insert a C)" by (simp add: C S aff_dim_convex_hull aff_dim_insert affine_independent_iff_card) next have "a \ C" using S a hull_inc by fastforce then show "int (card (insert a C)) = n + 1 + 1" by (simp add: C) next show "convex hull insert a S = convex hull (insert a C)" by (simp add: S convex_hull_insert_segments) qed (use C in auto) qed subsection \Simplicial complexes and triangulations\ definition\<^marker>\tag important\ simplicial_complex where "simplicial_complex \ \ finite \ \ (\S \ \. \n. n simplex S) \ (\F S. S \ \ \ F face_of S \ F \ \) \ (\S S'. S \ \ \ S' \ \ \ (S \ S') face_of S)" definition\<^marker>\tag important\ triangulation where "triangulation \ \ finite \ \ (\T \ \. \n. n simplex T) \ (\T T'. T \ \ \ T' \ \ \ (T \ T') face_of T)" subsection\Refining a cell complex to a simplicial complex\ proposition convex_hull_insert_Int_eq: fixes z :: "'a :: euclidean_space" assumes z: "z \ rel_interior S" and T: "T \ rel_frontier S" and U: "U \ rel_frontier S" and "convex S" "convex T" "convex U" shows "convex hull (insert z T) \ convex hull (insert z U) = convex hull (insert z (T \ U))" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" proof (cases "T={} \ U={}") case True then show ?thesis by auto next case False then have "T \ {}" "U \ {}" by auto have TU: "convex (T \ U)" by (simp add: \convex T\ \convex U\ convex_Int) have "(\x\T. closed_segment z x) \ (\x\U. closed_segment z x) \ (if T \ U = {} then {z} else \((closed_segment z) ` (T \ U)))" (is "_ \ ?IF") proof clarify fix x t u assume xt: "x \ closed_segment z t" and xu: "x \ closed_segment z u" and "t \ T" "u \ U" then have ne: "t \ z" "u \ z" using T U z unfolding rel_frontier_def by blast+ show "x \ ?IF" proof (cases "x = z") case True then show ?thesis by auto next case False have t: "t \ closure S" using T \t \ T\ rel_frontier_def by auto have u: "u \ closure S" using U \u \ U\ rel_frontier_def by auto show ?thesis proof (cases "t = u") case True then show ?thesis using \t \ T\ \u \ U\ xt by auto next case False have tnot: "t \ closed_segment u z" proof - have "t \ closure S - rel_interior S" using T \t \ T\ rel_frontier_def by blast then have "t \ open_segment z u" by (meson DiffD2 rel_interior_closure_convex_segment [OF \convex S\ z u] subsetD) then show ?thesis by (simp add: \t \ u\ \t \ z\ open_segment_commute open_segment_def) qed moreover have "u \ closed_segment z t" using rel_interior_closure_convex_segment [OF \convex S\ z t] \u \ U\ \u \ z\ U [unfolded rel_frontier_def] tnot by (auto simp: closed_segment_eq_open) ultimately have "\(between (t,u) z | between (u,z) t | between (z,t) u)" if "x \ z" using that xt xu apply (simp add: between_mem_segment [symmetric]) by (metis between_commute between_trans_2 between_antisym) then have "\ collinear {t, z, u}" if "x \ z" by (auto simp: that collinear_between_cases between_commute) moreover have "collinear {t, z, x}" by (metis closed_segment_commute collinear_2 collinear_closed_segment collinear_triples ends_in_segment(1) insert_absorb insert_absorb2 xt) moreover have "collinear {z, x, u}" by (metis closed_segment_commute collinear_2 collinear_closed_segment collinear_triples ends_in_segment(1) insert_absorb insert_absorb2 xu) ultimately have False using collinear_3_trans [of t z x u] \x \ z\ by blast then show ?thesis by metis qed qed qed then show ?thesis using False \convex T\ \convex U\ TU by (simp add: convex_hull_insert_segments hull_same split: if_split_asm) qed show "?rhs \ ?lhs" by (metis inf_greatest hull_mono inf.cobounded1 inf.cobounded2 insert_mono) qed lemma simplicial_subdivision_aux: assumes "finite \" and "\C. C \ \ \ polytope C" and "\C. C \ \ \ aff_dim C \ of_nat n" and "\C F. \C \ \; F face_of C\ \ F \ \" and "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1" shows "\\. simplicial_complex \ \ (\K \ \. aff_dim K \ of_nat n) \ \\ = \\ \ (\C \ \. \F. finite F \ F \ \ \ C = \F) \ (\K \ \. \C. C \ \ \ K \ C)" using assms proof (induction n arbitrary: \ rule: less_induct) case (less n) then have poly\: "\C. C \ \ \ polytope C" and aff\: "\C. C \ \ \ aff_dim C \ of_nat n" and face\: "\C F. \C \ \; F face_of C\ \ F \ \" and intface\: "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1" by metis+ show ?case proof (cases "n \ 1") case True have "\s. \n \ 1; s \ \\ \ \m. m simplex s" using poly\ aff\ by (force intro: polytope_lowdim_imp_simplex) then show ?thesis unfolding simplicial_complex_def apply (rule_tac x="\" in exI) using True by (auto simp: less.prems) next case False define \ where "\ \ {C \ \. aff_dim C < n}" have "finite \" "\C. C \ \ \ polytope C" "\C. C \ \ \ aff_dim C \ int (n - 1)" "\C F. \C \ \; F face_of C\ \ F \ \" "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1" using less.prems apply (auto simp: \_def) by (metis aff_dim_subset face_of_imp_subset less_le not_le) with less.IH [of "n-1" \] False obtain \ where "simplicial_complex \" and aff_dim\: "\K. K \ \ \ aff_dim K \ int (n - 1)" and "\\ = \\" and fin\: "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" and C\: "\K. K \ \ \ \C. C \ \ \ K \ C" by auto then have "finite \" and simpl\: "\S. S \ \ \ \n. n simplex S" and face\: "\F S. \S \ \; F face_of S\ \ F \ \" and faceI\: "\S S'. \S \ \; S' \ \\ \ (S \ S') face_of S" by (auto simp: simplicial_complex_def) define \ where "\ \ {C \ \. aff_dim C = n}" have "finite \" by (simp add: \_def less.prems(1)) have poly\: "\C. C \ \ \ polytope C" and convex\: "\C. C \ \ \ convex C" and closed\: "\C. C \ \ \ closed C" by (auto simp: \_def poly\ polytope_imp_convex polytope_imp_closed) have in_rel_interior: "(SOME z. z \ rel_interior C) \ rel_interior C" if "C \ \" for C using that poly\ polytope_imp_convex rel_interior_aff_dim some_in_eq by (fastforce simp: \_def) have *: "\T. \ affine_dependent T \ card T \ n \ aff_dim K < n \ K = convex hull T" if "K \ \" for K proof - obtain r where r: "r simplex K" using \K \ \\ simpl\ by blast have "r = aff_dim K" using \r simplex K\ aff_dim_simplex by blast with r show ?thesis unfolding simplex_def using False \\K. K \ \ \ aff_dim K \ int (n - 1)\ that by fastforce qed have ahK_C_disjoint: "affine hull K \ rel_interior C = {}" if "C \ \" "K \ \" "K \ rel_frontier C" for C K proof - have "convex C" "closed C" by (auto simp: convex\ closed\ \C \ \\) obtain F where F: "F face_of C" and "F \ C" "K \ F" proof - obtain L where "L \ \" "K \ L" using \K \ \\ C\ by blast have "K \ rel_frontier C" by (simp add: \K \ rel_frontier C\) also have "... \ C" by (simp add: \closed C\ rel_frontier_def subset_iff) finally have "K \ C" . have "L \ C face_of C" using \_def \_def \C \ \\ \L \ \\ intface\ by (simp add: inf_commute) moreover have "L \ C \ C" using \C \ \\ \L \ \\ apply (clarsimp simp: \_def \_def) by (metis aff_dim_subset inf_le1 not_le) moreover have "K \ L \ C" using \C \ \\ \L \ \\ \K \ C\ \K \ L\ by (auto simp: \_def \_def) ultimately show ?thesis using that by metis qed have "affine hull F \ rel_interior C = {}" by (rule affine_hull_face_of_disjoint_rel_interior [OF \convex C\ F \F \ C\]) with hull_mono [OF \K \ F\] show "affine hull K \ rel_interior C = {}" by fastforce qed let ?\ = "(\C \ \. \K \ \ \ Pow (rel_frontier C). {convex hull (insert (SOME z. z \ rel_interior C) K)})" have "\\. simplicial_complex \ \ (\K \ \. aff_dim K \ of_nat n) \ (\C \ \. \F. F \ \ \ C = \F) \ (\K \ \. \C. C \ \ \ K \ C)" proof (rule exI, intro conjI ballI) show "simplicial_complex (\ \ ?\)" unfolding simplicial_complex_def proof (intro conjI impI ballI allI) show "finite (\ \ ?\)" using \finite \\ \finite \\ by simp show "\n. n simplex S" if "S \ \ \ ?\" for S using that ahK_C_disjoint in_rel_interior simpl\ simplex_insert_dimplus1 by fastforce show "F \ \ \ ?\" if S: "S \ \ \ ?\ \ F face_of S" for F S proof - have "F \ \" if "S \ \" using S face\ that by blast moreover have "F \ \ \ ?\" if "F face_of S" "C \ \" "K \ \" and "K \ rel_frontier C" and S: "S = convex hull insert (SOME z. z \ rel_interior C) K" for C K proof - let ?z = "SOME z. z \ rel_interior C" have "?z \ rel_interior C" by (simp add: in_rel_interior \C \ \\) moreover obtain I where "\ affine_dependent I" "card I \ n" "aff_dim K < int n" "K = convex hull I" using * [OF \K \ \\] by auto ultimately have "?z \ affine hull I" using ahK_C_disjoint affine_hull_convex_hull that by blast have "compact I" "finite I" by (auto simp: \\ affine_dependent I\ aff_independent_finite finite_imp_compact) moreover have "F face_of convex hull insert ?z I" by (metis S \F face_of S\ \K = convex hull I\ convex_hull_eq_empty convex_hull_insert_segments hull_hull) ultimately obtain J where "J \ insert ?z I" "F = convex hull J" using face_of_convex_hull_subset [of "insert ?z I" F] by auto show ?thesis proof (cases "?z \ J") case True have "F \ (\K\\ \ Pow (rel_frontier C). {convex hull insert ?z K})" proof have "convex hull (J - {?z}) face_of K" by (metis True \J \ insert ?z I\ \K = convex hull I\ \\ affine_dependent I\ face_of_convex_hull_affine_independent subset_insert_iff) then have "convex hull (J - {?z}) \ \" by (rule face\ [OF \K \ \\]) moreover have "\x. x \ convex hull (J - {?z}) \ x \ rel_frontier C" by (metis True \J \ insert ?z I\ \K = convex hull I\ subsetD hull_mono subset_insert_iff that(4)) ultimately show "convex hull (J - {?z}) \ \ \ Pow (rel_frontier C)" by auto let ?F = "convex hull insert ?z (convex hull (J - {?z}))" have "F \ ?F" apply (clarsimp simp: \F = convex hull J\) by (metis True subsetD hull_mono hull_subset subset_insert_iff) moreover have "?F \ F" apply (clarsimp simp: \F = convex hull J\) by (metis (no_types, lifting) True convex_hull_eq_empty convex_hull_insert_segments hull_hull insert_Diff) ultimately show "F \ {?F}" by auto qed with \C\\\ show ?thesis by auto next case False then have "F \ \" using face_of_convex_hull_affine_independent [OF \\ affine_dependent I\] by (metis Int_absorb2 Int_insert_right_if0 \F = convex hull J\ \J \ insert ?z I\ \K = convex hull I\ face\ inf_le2 \K \ \\) then show "F \ \ \ ?\" by blast qed qed ultimately show ?thesis using that by auto qed have \
: "X \ Y face_of X \ X \ Y face_of Y" if XY: "X \ \" "Y \ ?\" for X Y proof - obtain C K where "C \ \" "K \ \" "K \ rel_frontier C" and Y: "Y = convex hull insert (SOME z. z \ rel_interior C) K" using XY by blast have "convex C" by (simp add: \C \ \\ convex\) have "K \ C" by (metis DiffE \C \ \\ \K \ rel_frontier C\ closed\ closure_closed rel_frontier_def subset_iff) let ?z = "(SOME z. z \ rel_interior C)" have z: "?z \ rel_interior C" using \C \ \\ in_rel_interior by blast obtain D where "D \ \" "X \ D" using C\ \X \ \\ by blast have "D \ rel_interior C = (C \ D) \ rel_interior C" using rel_interior_subset by blast also have "(C \ D) \ rel_interior C = {}" proof (rule face_of_disjoint_rel_interior) show "C \ D face_of C" using \_def \_def \C \ \\ \D \ \\ intface\ by blast show "C \ D \ C" by (metis (mono_tags, lifting) Int_lower2 \_def \_def \C \ \\ \D \ \\ aff_dim_subset mem_Collect_eq not_le) qed finally have DC: "D \ rel_interior C = {}" . have eq: "X \ convex hull (insert ?z K) = X \ convex hull K" apply (rule Int_convex_hull_insert_rel_exterior [OF \convex C\ \K \ C\ z]) using DC by (meson \X \ D\ disjnt_def disjnt_subset1) obtain I where I: "\ affine_dependent I" and Keq: "K = convex hull I" and [simp]: "convex hull K = K" using "*" \K \ \\ by force then have "?z \ affine hull I" using ahK_C_disjoint \C \ \\ \K \ \\ \K \ rel_frontier C\ affine_hull_convex_hull z by blast have "X \ K face_of K" by (simp add: XY(1) \K \ \\ faceI\ inf_commute) also have "... face_of convex hull insert ?z K" by (metis I Keq \?z \ affine hull I\ aff_independent_finite convex_convex_hull face_of_convex_hull_insert face_of_refl hull_insert) finally have "X \ K face_of convex hull insert ?z K" . then show ?thesis by (simp add: XY(1) Y \K \ \\ eq faceI\) qed show "S \ S' face_of S" if "S \ \ \ ?\ \ S' \ \ \ ?\" for S S' using that proof (elim conjE UnE) fix X Y assume "X \ \" and "Y \ \" then show "X \ Y face_of X" by (simp add: faceI\) next fix X Y assume XY: "X \ \" "Y \ ?\" then show "X \ Y face_of X" "Y \ X face_of Y" using \
[OF XY] by (auto simp: Int_commute) next fix X Y assume XY: "X \ ?\" "Y \ ?\" show "X \ Y face_of X" proof - obtain C K D L where "C \ \" "K \ \" "K \ rel_frontier C" and X: "X = convex hull insert (SOME z. z \ rel_interior C) K" and "D \ \" "L \ \" "L \ rel_frontier D" and Y: "Y = convex hull insert (SOME z. z \ rel_interior D) L" using XY by blast let ?z = "(SOME z. z \ rel_interior C)" have z: "?z \ rel_interior C" using \C \ \\ in_rel_interior by blast have "convex C" by (simp add: \C \ \\ convex\) have "convex K" using "*" \K \ \\ by blast have "convex L" by (meson \L \ \\ convex_simplex simpl\) show ?thesis proof (cases "D=C") case True then have "L \ rel_frontier C" using \L \ rel_frontier D\ by auto show ?thesis apply (simp add: X Y True) apply (simp add: convex_hull_insert_Int_eq [OF z] \K \ rel_frontier C\ \L \ rel_frontier C\ \convex C\ \convex K\ \convex L\) using face_of_polytope_insert2 by (metis "*" IntI \C \ \\ \K \ \\ \L \ \\\K \ rel_frontier C\ \L \ rel_frontier C\ aff_independent_finite ahK_C_disjoint empty_iff faceI\ polytope_convex_hull z) next case False have "convex D" by (simp add: \D \ \\ convex\) have "K \ C" by (metis DiffE \C \ \\ \K \ rel_frontier C\ closed\ closure_closed rel_frontier_def subset_eq) have "L \ D" by (metis DiffE \D \ \\ \L \ rel_frontier D\ closed\ closure_closed rel_frontier_def subset_eq) let ?w = "(SOME w. w \ rel_interior D)" have w: "?w \ rel_interior D" using \D \ \\ in_rel_interior by blast have "C \ rel_interior D = (D \ C) \ rel_interior D" using rel_interior_subset by blast also have "(D \ C) \ rel_interior D = {}" proof (rule face_of_disjoint_rel_interior) show "D \ C face_of D" using \_def \C \ \\ \D \ \\ intface\ by blast have "D \ \ \ aff_dim D = int n" using \_def \D \ \\ by blast moreover have "C \ \ \ aff_dim C = int n" using \_def \C \ \\ by blast ultimately show "D \ C \ D" by (metis Int_commute False face_of_aff_dim_lt inf.idem inf_le1 intface\ not_le poly\ polytope_imp_convex) qed finally have CD: "C \ (rel_interior D) = {}" . have zKC: "(convex hull insert ?z K) \ C" by (metis DiffE \C \ \\ \K \ rel_frontier C\ closed\ closure_closed convex\ hull_minimal insert_subset rel_frontier_def rel_interior_subset subset_iff z) have eq: "convex hull (insert ?z K) \ convex hull (insert ?w L) = convex hull (insert ?z K) \ convex hull L" apply (rule Int_convex_hull_insert_rel_exterior [OF \convex D\ \L \ D\ w]) using zKC CD apply (force simp: disjnt_def) done have ch_id: "convex hull K = K" "convex hull L = L" using "*" \K \ \\ \L \ \\ hull_same by auto have "convex C" by (simp add: \C \ \\ convex\) have "convex hull (insert ?z K) \ L = L \ convex hull (insert ?z K)" by blast also have "... = convex hull K \ L" proof (subst Int_convex_hull_insert_rel_exterior [OF \convex C\ \K \ C\ z]) have "(C \ D) \ rel_interior C = {}" proof (rule face_of_disjoint_rel_interior) show "C \ D face_of C" using \_def \C \ \\ \D \ \\ intface\ by blast have "D \ \" "aff_dim D = int n" using \_def \D \ \\ by fastforce+ moreover have "C \ \" "aff_dim C = int n" using \_def \C \ \\ by fastforce+ ultimately have "aff_dim D + - 1 * aff_dim C \ 0" by fastforce then have "\ C face_of D" using False \convex D\ face_of_aff_dim_lt by fastforce show "C \ D \ C" by (metis inf_commute \C \ \\ \D \ \\ \\ C face_of D\ intface\) qed then have "D \ rel_interior C = {}" by (metis inf.absorb_iff2 inf_assoc inf_sup_aci(1) rel_interior_subset) then show "disjnt L (rel_interior C)" by (meson \L \ D\ disjnt_def disjnt_subset1) next show "L \ convex hull K = convex hull K \ L" by force qed finally have chKL: "convex hull (insert ?z K) \ L = convex hull K \ L" . have "convex hull insert ?z K \ convex hull L face_of K" by (simp add: \K \ \\ \L \ \\ ch_id chKL faceI\) also have "... face_of convex hull insert ?z K" proof - obtain I where I: "\ affine_dependent I" "K = convex hull I" using * [OF \K \ \\] by auto then have "\a. a \ rel_interior C \ a \ affine hull I" using ahK_C_disjoint \C \ \\ \K \ \\ \K \ rel_frontier C\ affine_hull_convex_hull by blast then show ?thesis by (metis I affine_independent_insert face_of_convex_hull_affine_independent hull_insert subset_insertI z) qed finally have 1: "convex hull insert ?z K \ convex hull L face_of convex hull insert ?z K" . have "convex hull insert ?z K \ convex hull L face_of L" by (metis \K \ \\ \L \ \\ chKL ch_id faceI\ inf_commute) also have "... face_of convex hull insert ?w L" proof - obtain I where I: "\ affine_dependent I" "L = convex hull I" using * [OF \L \ \\] by auto then have "\a. a \ rel_interior D \ a \ affine hull I" using \D \ \\ \L \ \\ \L \ rel_frontier D\ affine_hull_convex_hull ahK_C_disjoint by blast then show ?thesis by (metis I aff_independent_finite convex_convex_hull face_of_convex_hull_insert face_of_refl hull_insert w) qed finally have 2: "convex hull insert ?z K \ convex hull L face_of convex hull insert ?w L" . show ?thesis by (simp add: X Y eq 1 2) qed qed qed qed show "\F \ \ \ ?\. C = \F" if "C \ \" for C proof (cases "C \ \") case True then show ?thesis by (meson UnCI fin\ subsetD subsetI) next case False then have "C \ \" by (simp add: \_def \_def aff\ less_le that) let ?z = "SOME z. z \ rel_interior C" have z: "?z \ rel_interior C" using \C \ \\ in_rel_interior by blast let ?F = "\K \ \ \ Pow (rel_frontier C). {convex hull (insert ?z K)}" have "?F \ ?\" using \C \ \\ by blast moreover have "C \ \?F" proof fix x assume "x \ C" have "convex C" using \C \ \\ convex\ by blast have "bounded C" using \C \ \\ by (simp add: poly\ polytope_imp_bounded that) have "polytope C" using \C \ \\ poly\ by auto have "\ (?z = x \ C = {?z})" using \C \ \\ aff_dim_sing [of ?z] \\ n \ 1\ by (force simp: \_def) then obtain y where y: "y \ rel_frontier C" and xzy: "x \ closed_segment ?z y" and sub: "open_segment ?z y \ rel_interior C" by (blast intro: segment_to_rel_frontier [OF \convex C\ \bounded C\ z \x \ C\]) then obtain F where "y \ F" "F face_of C" "F \ C" by (auto simp: rel_frontier_of_polyhedron_alt [OF polytope_imp_polyhedron [OF \polytope C\]]) then obtain \ where "finite \" "\ \ \" "F = \\" by (metis (mono_tags, lifting) \_def \C \ \\ \convex C\ aff\ face\ face_of_aff_dim_lt fin\ le_less_trans mem_Collect_eq not_less) then obtain K where "y \ K" "K \ \" using \y \ F\ by blast moreover have x: "x \ convex hull {?z,y}" using segment_convex_hull xzy by auto moreover have "convex hull {?z,y} \ convex hull insert ?z K" by (metis (full_types) \y \ K\ hull_mono empty_subsetI insertCI insert_subset) moreover have "K \ \" using \K \ \\ \\ \ \\ by blast moreover have "K \ rel_frontier C" using \F = \\\ \F \ C\ \F face_of C\ \K \ \\ face_of_subset_rel_frontier by fastforce ultimately show "x \ \?F" by force qed moreover have "convex hull insert (SOME z. z \ rel_interior C) K \ C" if "K \ \" "K \ rel_frontier C" for K proof (rule hull_minimal) show "insert (SOME z. z \ rel_interior C) K \ C" using that \C \ \\ in_rel_interior rel_interior_subset by (force simp: closure_eq rel_frontier_def closed\) show "convex C" by (simp add: \C \ \\ convex\) qed then have "\?F \ C" by auto ultimately show ?thesis by blast qed have "(\C. C \ \ \ L \ C) \ aff_dim L \ int n" if "L \ \ \ ?\" for L using that proof assume "L \ \" then show ?thesis using C\ \_def "*" by fastforce next assume "L \ ?\" then obtain C K where "C \ \" and L: "L = convex hull insert (SOME z. z \ rel_interior C) K" and K: "K \ \" "K \ rel_frontier C" by auto then have "convex hull C = C" by (meson convex\ convex_hull_eq) then have "convex C" by (metis (no_types) convex_convex_hull) have "rel_frontier C \ C" by (metis DiffE closed\ \C \ \\ closure_closed rel_frontier_def subsetI) have "K \ C" using K \rel_frontier C \ C\ by blast have "C \ \" using \_def \C \ \\ by auto moreover have "L \ C" using K L \C \ \\ by (metis \K \ C\ \convex hull C = C\ contra_subsetD hull_mono in_rel_interior insert_subset rel_interior_subset) ultimately show ?thesis using \rel_frontier C \ C\ \L \ C\ aff\ aff_dim_subset \C \ \\ dual_order.trans by blast qed then show "\C. C \ \ \ L \ C" "aff_dim L \ int n" if "L \ \ \ ?\" for L using that by auto qed then show ?thesis apply (rule ex_forward, safe) apply (meson Union_iff subsetCE, fastforce) by (meson infinite_super simplicial_complex_def) qed qed lemma simplicial_subdivision_of_cell_complex_lowdim: assumes "finite \" and poly: "\C. C \ \ \ polytope C" and face: "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1" and aff: "\C. C \ \ \ aff_dim C \ d" obtains \ where "simplicial_complex \" "\K. K \ \ \ aff_dim K \ d" "\\ = \\" "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" "\K. K \ \ \ \C. C \ \ \ K \ C" proof (cases "d \ 0") case True then obtain n where n: "d = of_nat n" using zero_le_imp_eq_int by blast have "\\. simplicial_complex \ \ (\K\\. aff_dim K \ int n) \ \\ = \(\C\\. {F. F face_of C}) \ (\C\\C\\. {F. F face_of C}. \F. finite F \ F \ \ \ C = \F) \ (\K\\. \C. C \ (\C\\. {F. F face_of C}) \ K \ C)" proof (rule simplicial_subdivision_aux) show "finite (\C\\. {F. F face_of C})" using \finite \\ poly polyhedron_eq_finite_faces polytope_imp_polyhedron by fastforce show "polytope F" if "F \ (\C\\. {F. F face_of C})" for F using poly that face_of_polytope_polytope by blast show "aff_dim F \ int n" if "F \ (\C\\. {F. F face_of C})" for F using that by clarify (metis n aff_dim_subset aff face_of_imp_subset order_trans) show "F \ (\C\\. {F. F face_of C})" if "G \ (\C\\. {F. F face_of C})" and "F face_of G" for F G using that face_of_trans by blast next fix F1 F2 assume "F1 \ (\C\\. {F. F face_of C})" and "F2 \ (\C\\. {F. F face_of C})" then obtain C1 C2 where "C1 \ \" "C2 \ \" and F: "F1 face_of C1" "F2 face_of C2" by auto show "F1 \ F2 face_of F1" using face_of_Int_subface [OF _ _ F] by (metis \C1 \ \\ \C2 \ \\ face inf_commute) qed moreover have "\(\C\\. {F. F face_of C}) = \\" using face_of_imp_subset face by blast ultimately show ?thesis apply clarify apply (rule that, assumption+) using n apply blast apply (simp_all add: poly face_of_refl polytope_imp_convex) using face_of_imp_subset by fastforce next case False then have m1: "\C. C \ \ \ aff_dim C = -1" by (metis aff aff_dim_empty_eq aff_dim_negative_iff dual_order.trans not_less) then have face\: "\F S. \S \ \; F face_of S\ \ F \ \" by (metis aff_dim_empty face_of_empty) show ?thesis proof have "\S. S \ \ \ \n. n simplex S" by (metis (no_types) m1 aff_dim_empty simplex_minus_1) then show "simplicial_complex \" by (auto simp: simplicial_complex_def \finite \\ face intro: face\) show "aff_dim K \ d" if "K \ \" for K by (simp add: that aff) show "\F. finite F \ F \ \ \ C = \F" if "C \ \" for C using \C \ \\ equals0I by auto show "\C. C \ \ \ K \ C" if "K \ \" for K using \K \ \\ by blast qed auto qed proposition simplicial_subdivision_of_cell_complex: assumes "finite \" and poly: "\C. C \ \ \ polytope C" and face: "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1" obtains \ where "simplicial_complex \" "\\ = \\" "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" "\K. K \ \ \ \C. C \ \ \ K \ C" by (blast intro: simplicial_subdivision_of_cell_complex_lowdim [OF assms aff_dim_le_DIM]) corollary fine_simplicial_subdivision_of_cell_complex: assumes "0 < e" "finite \" and poly: "\C. C \ \ \ polytope C" and face: "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1" obtains \ where "simplicial_complex \" "\K. K \ \ \ diameter K < e" "\\ = \\" "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" "\K. K \ \ \ \C. C \ \ \ K \ C" proof - obtain \ where \: "finite \" "\\ = \\" and diapoly: "\X. X \ \ \ diameter X < e" "\X. X \ \ \ polytope X" and "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X" and \covers: "\C x. C \ \ \ x \ C \ \D. D \ \ \ x \ D \ D \ C" and \covered: "\C. C \ \ \ \D. D \ \ \ C \ D" by (blast intro: cell_complex_subdivision_exists [OF \0 < e\ \finite \\ poly aff_dim_le_DIM face]) then obtain \ where \: "simplicial_complex \" "\\ = \\" and \covers: "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" and \covered: "\K. K \ \ \ \C. C \ \ \ K \ C" using simplicial_subdivision_of_cell_complex [OF \finite \\] by metis show ?thesis proof show "simplicial_complex \" by (rule \) show "diameter K < e" if "K \ \" for K by (metis le_less_trans diapoly \covered diameter_subset polytope_imp_bounded that) show "\\ = \\" by (simp add: \(2) \\\ = \\\) show "\F. finite F \ F \ \ \ C = \F" if "C \ \" for C proof - { fix x assume "x \ C" then obtain D where "D \ \" "x \ D" "D \ C" using \covers \C \ \\ \covers by force then have "\X\\ \ Pow C. x \ X" using \D \ \\ \D \ C\ \x \ D\ by blast } moreover have "finite (\ \ Pow C)" using \simplicial_complex \\ simplicial_complex_def by auto ultimately show ?thesis by (rule_tac x="(\ \ Pow C)" in exI) auto qed show "\C. C \ \ \ K \ C" if "K \ \" for K by (meson \covered \covered order_trans that) qed qed subsection\Some results on cell division with full-dimensional cells only\ lemma convex_Union_fulldim_cells: assumes "finite \" and clo: "\C. C \ \ \ closed C" and con: "\C. C \ \ \ convex C" and eq: "\\ = U"and "convex U" shows "\{C \ \. aff_dim C = aff_dim U} = U" (is "?lhs = U") proof - have "closed U" using \finite \\ clo eq by blast have "?lhs \ U" using eq by blast moreover have "U \ ?lhs" proof (cases "\C \ \. aff_dim C = aff_dim U") case True then show ?thesis using eq by blast next case False have "closed ?lhs" by (simp add: \finite \\ clo closed_Union) moreover have "U \ closure ?lhs" proof - have "U \ closure(\{U - C |C. C \ \ \ aff_dim C < aff_dim U})" proof (rule Baire [OF \closed U\]) show "countable {U - C |C. C \ \ \ aff_dim C < aff_dim U}" using \finite \\ uncountable_infinite by fastforce have "\C. C \ \ \ openin (top_of_set U) (U-C)" by (metis Sup_upper clo closed_limpt closedin_limpt eq openin_diff openin_subtopology_self) then show "openin (top_of_set U) T \ U \ closure T" if "T \ {U - C |C. C \ \ \ aff_dim C < aff_dim U}" for T using that dense_complement_convex_closed \closed U\ \convex U\ by auto qed also have "... \ closure ?lhs" proof - obtain C where "C \ \" "aff_dim C < aff_dim U" by (metis False Sup_upper aff_dim_subset eq eq_iff not_le) have "\X. X \ \ \ aff_dim X = aff_dim U \ x \ X" if "\V. (\C. V = U - C \ C \ \ \ aff_dim C < aff_dim U) \ x \ V" for x proof - have "x \ U \ x \ \\" using \C \ \\ \aff_dim C < aff_dim U\ eq that by blast then show ?thesis by (metis Diff_iff Sup_upper Union_iff aff_dim_subset dual_order.order_iff_strict eq that) qed then show ?thesis by (auto intro!: closure_mono) qed finally show ?thesis . qed ultimately show ?thesis using closure_subset_eq by blast qed ultimately show ?thesis by blast qed proposition fine_triangular_subdivision_of_cell_complex: assumes "0 < e" "finite \" and poly: "\C. C \ \ \ polytope C" and aff: "\C. C \ \ \ aff_dim C = d" and face: "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1" obtains \ where "triangulation \" "\k. k \ \ \ diameter k < e" "\k. k \ \ \ aff_dim k = d" "\\ = \\" "\C. C \ \ \ \f. finite f \ f \ \ \ C = \f" "\k. k \ \ \ \C. C \ \ \ k \ C" proof - obtain \ where "simplicial_complex \" and dia\: "\K. K \ \ \ diameter K < e" and "\\ = \\" and in\: "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" and in\: "\K. K \ \ \ \C. C \ \ \ K \ C" by (blast intro: fine_simplicial_subdivision_of_cell_complex [OF \e > 0\ \finite \\ poly face]) let ?\ = "{K \ \. aff_dim K = d}" show thesis proof show "triangulation ?\" using \simplicial_complex \\ by (auto simp: triangulation_def simplicial_complex_def) show "diameter L < e" if "L \ {K \ \. aff_dim K = d}" for L using that by (auto simp: dia\) show "aff_dim L = d" if "L \ {K \ \. aff_dim K = d}" for L using that by auto show "\F. finite F \ F \ {K \ \. aff_dim K = d} \ C = \F" if "C \ \" for C proof - obtain F where "finite F" "F \ \" "C = \F" using in\ [OF \C \ \\] by auto show ?thesis proof (intro exI conjI) show "finite {K \ F. aff_dim K = d}" by (simp add: \finite F\) show "{K \ F. aff_dim K = d} \ {K \ \. aff_dim K = d}" using \F \ \\ by blast have "d = aff_dim C" by (simp add: aff that) moreover have "\K. K \ F \ closed K \ convex K" using \simplicial_complex \\ \F \ \\ unfolding simplicial_complex_def by (metis subsetCE \F \ \\ closed_simplex convex_simplex) moreover have "convex (\F)" using \C = \F\ poly polytope_imp_convex that by blast ultimately show "C = \{K \ F. aff_dim K = d}" by (simp add: convex_Union_fulldim_cells \C = \F\ \finite F\) qed qed then show "\{K \ \. aff_dim K = d} = \\" by auto (meson in\ subsetCE) show "\C. C \ \ \ L \ C" if "L \ {K \ \. aff_dim K = d}" for L using that by (auto simp: in\) qed qed end diff --git a/src/HOL/Library/Infinite_Set.thy b/src/HOL/Library/Infinite_Set.thy --- a/src/HOL/Library/Infinite_Set.thy +++ b/src/HOL/Library/Infinite_Set.thy @@ -1,587 +1,588 @@ (* Title: HOL/Library/Infinite_Set.thy Author: Stephan Merz *) section \Infinite Sets and Related Concepts\ theory Infinite_Set imports Main begin subsection \The set of natural numbers is infinite\ lemma infinite_nat_iff_unbounded_le: "infinite S \ (\m. \n\m. n \ S)" for S :: "nat set" using frequently_cofinite[of "\x. x \ S"] by (simp add: cofinite_eq_sequentially frequently_def eventually_sequentially) lemma infinite_nat_iff_unbounded: "infinite S \ (\m. \n>m. n \ S)" for S :: "nat set" using frequently_cofinite[of "\x. x \ S"] by (simp add: cofinite_eq_sequentially frequently_def eventually_at_top_dense) lemma finite_nat_iff_bounded: "finite S \ (\k. S \ {.. (\k. S \ {.. k})" for S :: "nat set" using infinite_nat_iff_unbounded[of S] by (simp add: subset_eq) (metis not_le) lemma finite_nat_bounded: "finite S \ \k. S \ {.. For a set of natural numbers to be infinite, it is enough to know that for any number larger than some \k\, there is some larger number that is an element of the set. \ lemma unbounded_k_infinite: "\m>k. \n>m. n \ S \ infinite (S::nat set)" apply (clarsimp simp add: finite_nat_set_iff_bounded) apply (drule_tac x="Suc (max m k)" in spec) using less_Suc_eq apply fastforce done lemma nat_not_finite: "finite (UNIV::nat set) \ R" by simp lemma range_inj_infinite: fixes f :: "nat \ 'a" assumes "inj f" shows "infinite (range f)" proof assume "finite (range f)" from this assms have "finite (UNIV::nat set)" by (rule finite_imageD) then show False by simp qed subsection \The set of integers is also infinite\ lemma infinite_int_iff_infinite_nat_abs: "infinite S \ infinite ((nat \ abs) ` S)" for S :: "int set" proof (unfold Not_eq_iff, rule iffI) assume "finite ((nat \ abs) ` S)" then have "finite (nat ` (abs ` S))" by (simp add: image_image cong: image_cong) moreover have "inj_on nat (abs ` S)" by (rule inj_onI) auto ultimately have "finite (abs ` S)" by (rule finite_imageD) then show "finite S" by (rule finite_image_absD) qed simp proposition infinite_int_iff_unbounded_le: "infinite S \ (\m. \n. \n\ \ m \ n \ S)" for S :: "int set" by (simp add: infinite_int_iff_infinite_nat_abs infinite_nat_iff_unbounded_le o_def image_def) (metis abs_ge_zero nat_le_eq_zle le_nat_iff) proposition infinite_int_iff_unbounded: "infinite S \ (\m. \n. \n\ > m \ n \ S)" for S :: "int set" by (simp add: infinite_int_iff_infinite_nat_abs infinite_nat_iff_unbounded o_def image_def) (metis (full_types) nat_le_iff nat_mono not_le) proposition finite_int_iff_bounded: "finite S \ (\k. abs ` S \ {.. (\k. abs ` S \ {.. k})" for S :: "int set" using infinite_int_iff_unbounded[of S] by (simp add: subset_eq) (metis not_le) subsection \Infinitely Many and Almost All\ text \ We often need to reason about the existence of infinitely many (resp., all but finitely many) objects satisfying some predicate, so we introduce corresponding binders and their proof rules. \ lemma not_INFM [simp]: "\ (INFM x. P x) \ (MOST x. \ P x)" by (rule not_frequently) lemma not_MOST [simp]: "\ (MOST x. P x) \ (INFM x. \ P x)" by (rule not_eventually) lemma INFM_const [simp]: "(INFM x::'a. P) \ P \ infinite (UNIV::'a set)" by (simp add: frequently_const_iff) lemma MOST_const [simp]: "(MOST x::'a. P) \ P \ finite (UNIV::'a set)" by (simp add: eventually_const_iff) lemma INFM_imp_distrib: "(INFM x. P x \ Q x) \ ((MOST x. P x) \ (INFM x. Q x))" by (rule frequently_imp_iff) lemma MOST_imp_iff: "MOST x. P x \ (MOST x. P x \ Q x) \ (MOST x. Q x)" by (auto intro: eventually_rev_mp eventually_mono) lemma INFM_conjI: "INFM x. P x \ MOST x. Q x \ INFM x. P x \ Q x" by (rule frequently_rev_mp[of P]) (auto elim: eventually_mono) text \Properties of quantifiers with injective functions.\ lemma INFM_inj: "INFM x. P (f x) \ inj f \ INFM x. P x" using finite_vimageI[of "{x. P x}" f] by (auto simp: frequently_cofinite) lemma MOST_inj: "MOST x. P x \ inj f \ MOST x. P (f x)" using finite_vimageI[of "{x. \ P x}" f] by (auto simp: eventually_cofinite) text \Properties of quantifiers with singletons.\ lemma not_INFM_eq [simp]: "\ (INFM x. x = a)" "\ (INFM x. a = x)" unfolding frequently_cofinite by simp_all lemma MOST_neq [simp]: "MOST x. x \ a" "MOST x. a \ x" unfolding eventually_cofinite by simp_all lemma INFM_neq [simp]: "(INFM x::'a. x \ a) \ infinite (UNIV::'a set)" "(INFM x::'a. a \ x) \ infinite (UNIV::'a set)" unfolding frequently_cofinite by simp_all lemma MOST_eq [simp]: "(MOST x::'a. x = a) \ finite (UNIV::'a set)" "(MOST x::'a. a = x) \ finite (UNIV::'a set)" unfolding eventually_cofinite by simp_all lemma MOST_eq_imp: "MOST x. x = a \ P x" "MOST x. a = x \ P x" unfolding eventually_cofinite by simp_all text \Properties of quantifiers over the naturals.\ lemma MOST_nat: "(\\<^sub>\n. P n) \ (\m. \n>m. P n)" for P :: "nat \ bool" by (auto simp add: eventually_cofinite finite_nat_iff_bounded_le subset_eq simp flip: not_le) lemma MOST_nat_le: "(\\<^sub>\n. P n) \ (\m. \n\m. P n)" for P :: "nat \ bool" by (auto simp add: eventually_cofinite finite_nat_iff_bounded subset_eq simp flip: not_le) lemma INFM_nat: "(\\<^sub>\n. P n) \ (\m. \n>m. P n)" for P :: "nat \ bool" by (simp add: frequently_cofinite infinite_nat_iff_unbounded) lemma INFM_nat_le: "(\\<^sub>\n. P n) \ (\m. \n\m. P n)" for P :: "nat \ bool" by (simp add: frequently_cofinite infinite_nat_iff_unbounded_le) lemma MOST_INFM: "infinite (UNIV::'a set) \ MOST x::'a. P x \ INFM x::'a. P x" by (simp add: eventually_frequently) lemma MOST_Suc_iff: "(MOST n. P (Suc n)) \ (MOST n. P n)" by (simp add: cofinite_eq_sequentially) lemma MOST_SucI: "MOST n. P n \ MOST n. P (Suc n)" and MOST_SucD: "MOST n. P (Suc n) \ MOST n. P n" by (simp_all add: MOST_Suc_iff) lemma MOST_ge_nat: "MOST n::nat. m \ n" by (simp add: cofinite_eq_sequentially) \ \legacy names\ lemma Inf_many_def: "Inf_many P \ infinite {x. P x}" by (fact frequently_cofinite) lemma Alm_all_def: "Alm_all P \ \ (INFM x. \ P x)" by simp lemma INFM_iff_infinite: "(INFM x. P x) \ infinite {x. P x}" by (fact frequently_cofinite) lemma MOST_iff_cofinite: "(MOST x. P x) \ finite {x. \ P x}" by (fact eventually_cofinite) lemma INFM_EX: "(\\<^sub>\x. P x) \ (\x. P x)" by (fact frequently_ex) lemma ALL_MOST: "\x. P x \ \\<^sub>\x. P x" by (fact always_eventually) lemma INFM_mono: "\\<^sub>\x. P x \ (\x. P x \ Q x) \ \\<^sub>\x. Q x" by (fact frequently_elim1) lemma MOST_mono: "\\<^sub>\x. P x \ (\x. P x \ Q x) \ \\<^sub>\x. Q x" by (fact eventually_mono) lemma INFM_disj_distrib: "(\\<^sub>\x. P x \ Q x) \ (\\<^sub>\x. P x) \ (\\<^sub>\x. Q x)" by (fact frequently_disj_iff) lemma MOST_rev_mp: "\\<^sub>\x. P x \ \\<^sub>\x. P x \ Q x \ \\<^sub>\x. Q x" by (fact eventually_rev_mp) lemma MOST_conj_distrib: "(\\<^sub>\x. P x \ Q x) \ (\\<^sub>\x. P x) \ (\\<^sub>\x. Q x)" by (fact eventually_conj_iff) lemma MOST_conjI: "MOST x. P x \ MOST x. Q x \ MOST x. P x \ Q x" by (fact eventually_conj) lemma INFM_finite_Bex_distrib: "finite A \ (INFM y. \x\A. P x y) \ (\x\A. INFM y. P x y)" by (fact frequently_bex_finite_distrib) lemma MOST_finite_Ball_distrib: "finite A \ (MOST y. \x\A. P x y) \ (\x\A. MOST y. P x y)" by (fact eventually_ball_finite_distrib) lemma INFM_E: "INFM x. P x \ (\x. P x \ thesis) \ thesis" by (fact frequentlyE) lemma MOST_I: "(\x. P x) \ MOST x. P x" by (rule eventuallyI) lemmas MOST_iff_finiteNeg = MOST_iff_cofinite subsection \Enumeration of an Infinite Set\ text \The set's element type must be wellordered (e.g. the natural numbers).\ text \ Could be generalized to \<^prop>\enumerate' S n = (SOME t. t \ s \ finite {s\S. s < t} \ card {s\S. s < t} = n)\. \ primrec (in wellorder) enumerate :: "'a set \ nat \ 'a" where enumerate_0: "enumerate S 0 = (LEAST n. n \ S)" | enumerate_Suc: "enumerate S (Suc n) = enumerate (S - {LEAST n. n \ S}) n" lemma enumerate_Suc': "enumerate S (Suc n) = enumerate (S - {enumerate S 0}) n" by simp lemma enumerate_in_set: "infinite S \ enumerate S n \ S" proof (induct n arbitrary: S) case 0 then show ?case by (fastforce intro: LeastI dest!: infinite_imp_nonempty) next case (Suc n) then show ?case by simp (metis DiffE infinite_remove) qed declare enumerate_0 [simp del] enumerate_Suc [simp del] lemma enumerate_step: "infinite S \ enumerate S n < enumerate S (Suc n)" proof (induction n arbitrary: S) case 0 then have "enumerate S 0 \ enumerate S (Suc 0)" by (simp add: enumerate_0 Least_le enumerate_in_set) moreover have "enumerate (S - {enumerate S 0}) 0 \ S - {enumerate S 0}" by (meson "0.prems" enumerate_in_set infinite_remove) then have "enumerate S 0 \ enumerate (S - {enumerate S 0}) 0" by auto ultimately show ?case by (simp add: enumerate_Suc') next case (Suc n) then show ?case by (simp add: enumerate_Suc') qed lemma enumerate_mono: "m < n \ infinite S \ enumerate S m < enumerate S n" by (induct m n rule: less_Suc_induct) (auto intro: enumerate_step) lemma le_enumerate: assumes S: "infinite S" shows "n \ enumerate S n" using S proof (induct n) case 0 then show ?case by simp next case (Suc n) then have "n \ enumerate S n" by simp also note enumerate_mono[of n "Suc n", OF _ \infinite S\] finally show ?case by simp qed lemma infinite_enumerate: assumes fS: "infinite S" shows "\r::nat\nat. strict_mono r \ (\n. r n \ S)" unfolding strict_mono_def using enumerate_in_set[OF fS] enumerate_mono[of _ _ S] fS by auto lemma enumerate_Suc'': fixes S :: "'a::wellorder set" assumes "infinite S" shows "enumerate S (Suc n) = (LEAST s. s \ S \ enumerate S n < s)" using assms proof (induct n arbitrary: S) case 0 then have "\s \ S. enumerate S 0 \ s" by (auto simp: enumerate.simps intro: Least_le) then show ?case unfolding enumerate_Suc' enumerate_0[of "S - {enumerate S 0}"] by (intro arg_cong[where f = Least] ext) auto next case (Suc n S) show ?case using enumerate_mono[OF zero_less_Suc \infinite S\, of n] \infinite S\ apply (subst (1 2) enumerate_Suc') apply (subst Suc) apply (use \infinite S\ in simp) apply (intro arg_cong[where f = Least] ext) apply (auto simp flip: enumerate_Suc') done qed lemma enumerate_Ex: fixes S :: "nat set" assumes S: "infinite S" and s: "s \ S" shows "\n. enumerate S n = s" using s proof (induct s rule: less_induct) case (less s) show ?case proof (cases "\y\S. y < s") case True let ?y = "Max {s'\S. s' < s}" from True have y: "\x. ?y < x \ (\s'\S. s' < s \ s' < x)" by (subst Max_less_iff) auto then have y_in: "?y \ {s'\S. s' < s}" by (intro Max_in) auto with less.hyps[of ?y] obtain n where "enumerate S n = ?y" by auto with S have "enumerate S (Suc n) = s" by (auto simp: y less enumerate_Suc'' intro!: Least_equality) then show ?thesis by auto next case False then have "\t\S. s \ t" by auto with \s \ S\ show ?thesis by (auto intro!: exI[of _ 0] Least_equality simp: enumerate_0) qed qed lemma inj_enumerate: fixes S :: "'a::wellorder set" assumes S: "infinite S" shows "inj (enumerate S)" unfolding inj_on_def proof clarsimp show "\x y. enumerate S x = enumerate S y \ x = y" by (metis neq_iff enumerate_mono[OF _ \infinite S\]) qed text \To generalise this, we'd need a condition that all initial segments were finite\ lemma bij_enumerate: fixes S :: "nat set" assumes S: "infinite S" shows "bij_betw (enumerate S) UNIV S" proof - have "\s \ S. \i. enumerate S i = s" using enumerate_Ex[OF S] by auto moreover note \infinite S\ inj_enumerate ultimately show ?thesis unfolding bij_betw_def by (auto intro: enumerate_in_set) qed text \A pair of weird and wonderful lemmas from HOL Light.\ lemma finite_transitivity_chain: assumes "finite A" and R: "\x. \ R x x" "\x y z. \R x y; R y z\ \ R x z" and A: "\x. x \ A \ \y. y \ A \ R x y" shows "A = {}" using \finite A\ A proof (induct A) case empty then show ?case by simp next case (insert a A) - with R show ?case - by (metis empty_iff insert_iff) (* somewhat slow *) + have False + using R(1)[of a] R(2)[of _ a] insert(3,4) by blast + thus ?case .. qed corollary Union_maximal_sets: assumes "finite \" shows "\{T \ \. \U\\. \ T \ U} = \\" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by force show "?rhs \ ?lhs" proof (rule Union_subsetI) fix S assume "S \ \" have "{T \ \. S \ T} = {}" if "\ (\y. y \ {T \ \. \U\\. \ T \ U} \ S \ y)" proof - have \
: "\x. x \ \ \ S \ x \ \y. y \ \ \ S \ y \ x \ y" using that by (blast intro: dual_order.trans psubset_imp_subset) show ?thesis proof (rule finite_transitivity_chain [of _ "\T U. S \ T \ T \ U"]) qed (use assms in \auto intro: \
\) qed with \S \ \\ show "\y. y \ {T \ \. \U\\. \ T \ U} \ S \ y" by blast qed qed subsection \Properties of @{term enumerate} on finite sets\ lemma finite_enumerate_in_set: "\finite S; n < card S\ \ enumerate S n \ S" proof (induction n arbitrary: S) case 0 then show ?case by (metis all_not_in_conv card_empty enumerate.simps(1) not_less0 wellorder_Least_lemma(1)) next case (Suc n) show ?case using Suc.prems Suc.IH [of "S - {LEAST n. n \ S}"] apply (simp add: enumerate.simps) by (metis Diff_empty Diff_insert0 Suc_lessD card.remove less_Suc_eq) qed lemma finite_enumerate_step: "\finite S; Suc n < card S\ \ enumerate S n < enumerate S (Suc n)" proof (induction n arbitrary: S) case 0 then have "enumerate S 0 \ enumerate S (Suc 0)" by (simp add: Least_le enumerate.simps(1) finite_enumerate_in_set) moreover have "enumerate (S - {enumerate S 0}) 0 \ S - {enumerate S 0}" by (metis 0 Suc_lessD Suc_less_eq card_Suc_Diff1 enumerate_in_set finite_enumerate_in_set) then have "enumerate S 0 \ enumerate (S - {enumerate S 0}) 0" by auto ultimately show ?case by (simp add: enumerate_Suc') next case (Suc n) then show ?case by (simp add: enumerate_Suc' finite_enumerate_in_set) qed lemma finite_enumerate_mono: "\m < n; finite S; n < card S\ \ enumerate S m < enumerate S n" by (induct m n rule: less_Suc_induct) (auto intro: finite_enumerate_step) lemma finite_le_enumerate: assumes "finite S" "n < card S" shows "n \ enumerate S n" using assms proof (induction n) case 0 then show ?case by simp next case (Suc n) then have "n \ enumerate S n" by simp also note finite_enumerate_mono[of n "Suc n", OF _ \finite S\] finally show ?case using Suc.prems(2) Suc_leI by blast qed lemma finite_enumerate: assumes fS: "finite S" shows "\r::nat\nat. strict_mono_on r {.. (\n S)" unfolding strict_mono_def using finite_enumerate_in_set[OF fS] finite_enumerate_mono[of _ _ S] fS by (metis lessThan_iff strict_mono_on_def) lemma finite_enumerate_Suc'': fixes S :: "'a::wellorder set" assumes "finite S" "Suc n < card S" shows "enumerate S (Suc n) = (LEAST s. s \ S \ enumerate S n < s)" using assms proof (induction n arbitrary: S) case 0 then have "\s \ S. enumerate S 0 \ s" by (auto simp: enumerate.simps intro: Least_le) then show ?case unfolding enumerate_Suc' enumerate_0[of "S - {enumerate S 0}"] by (metis Diff_iff dual_order.strict_iff_order singletonD singletonI) next case (Suc n S) then have "Suc n < card (S - {enumerate S 0})" using Suc.prems(2) finite_enumerate_in_set by force then show ?case apply (subst (1 2) enumerate_Suc') apply (simp add: Suc) apply (intro arg_cong[where f = Least] HOL.ext) using finite_enumerate_mono[OF zero_less_Suc \finite S\, of n] Suc.prems by (auto simp flip: enumerate_Suc') qed lemma finite_enumerate_initial_segment: fixes S :: "'a::wellorder set" assumes "finite S" "s \ S" and n: "n < card (S \ {.. {.. S \ n < s) = (LEAST n. n \ S)" proof (rule Least_equality) have "\t. t \ S \ t < s" by (metis "0" card_gt_0_iff disjoint_iff_not_equal lessThan_iff) then show "(LEAST n. n \ S) \ S \ (LEAST n. n \ S) < s" by (meson LeastI Least_le le_less_trans) qed (simp add: Least_le) then show ?case by (auto simp: enumerate_0) next case (Suc n) then have less_card: "Suc n < card S" by (meson assms(1) card_mono inf_sup_ord(1) leD le_less_linear order.trans) obtain t where t: "t \ {s \ S. enumerate S n < s}" by (metis Infinite_Set.enumerate_step enumerate_in_set finite_enumerate_in_set finite_enumerate_step less_card mem_Collect_eq) have "(LEAST x. x \ S \ x < s \ enumerate S n < x) = (LEAST s. s \ S \ enumerate S n < s)" (is "_ = ?r") proof (intro Least_equality conjI) show "?r \ S" by (metis (mono_tags, lifting) LeastI mem_Collect_eq t) show "?r < s" using not_less_Least [of _ "\t. t \ S \ enumerate S n < t"] Suc assms by (metis (no_types, lifting) Int_Collect Suc_lessD finite_Int finite_enumerate_in_set finite_enumerate_step lessThan_def linorder_cases) show "enumerate S n < ?r" by (metis (no_types, lifting) LeastI mem_Collect_eq t) qed (auto simp: Least_le) then show ?case using Suc assms by (simp add: finite_enumerate_Suc'' less_card) qed lemma finite_enumerate_Ex: fixes S :: "'a::wellorder set" assumes S: "finite S" and s: "s \ S" shows "\ny\S. y < s") case True let ?T = "S \ {..finite S\]) from True have y: "\x. Max ?T < x \ (\s'\S. s' < s \ s' < x)" by (subst Max_less_iff) (auto simp: \finite ?T\) then have y_in: "Max ?T \ {s'\S. s' < s}" using Max_in \finite ?T\ by fastforce with less.IH[of "Max ?T" ?T] obtain n where n: "enumerate ?T n = Max ?T" "n < card ?T" using \finite ?T\ by blast then have "Suc n < card S" using TS less_trans_Suc by blast with S n have "enumerate S (Suc n) = s" by (subst finite_enumerate_Suc'') (auto simp: y finite_enumerate_initial_segment less finite_enumerate_Suc'' intro!: Least_equality) then show ?thesis using \Suc n < card S\ by blast next case False then have "\t\S. s \ t" by auto moreover have "0 < card S" using card_0_eq less.prems by blast ultimately show ?thesis using \s \ S\ by (auto intro!: exI[of _ 0] Least_equality simp: enumerate_0) qed qed lemma finite_bij_enumerate: fixes S :: "'a::wellorder set" assumes S: "finite S" shows "bij_betw (enumerate S) {..n m. \n \ m; n < card S; m < card S\ \ enumerate S n \ enumerate S m" using finite_enumerate_mono[OF _ \finite S\] by (auto simp: neq_iff) then have "inj_on (enumerate S) {..s \ S. \ifinite S\ ultimately show ?thesis unfolding bij_betw_def by (auto intro: finite_enumerate_in_set) qed lemma ex_bij_betw_strict_mono_card: fixes M :: "'a::wellorder set" assumes "finite M" obtains h where "bij_betw h {..