# HG changeset patch # User paulson # Date 1553012091 0 # Node ID 4a9167f377b0633e4714726cc36c5075294ff195 # Parent eddcc7c726f3e57e73c59d328e19a3455874f457 new material about topology, etc.; also fixes for yesterday's diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Abstract_Topology.thy --- a/src/HOL/Analysis/Abstract_Topology.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Abstract_Topology.thy Tue Mar 19 16:14:51 2019 +0000 @@ -8,7 +8,6 @@ Complex_Main "HOL-Library.Set_Idioms" "HOL-Library.FuncSet" - (* Path_Connected *) begin subsection \General notion of a topology as a value\ @@ -427,7 +426,7 @@ lemma topspace_euclidean [simp]: "topspace euclidean = UNIV" by (force simp: topspace_def) -lemma topspace_euclidean_subtopology[simp]: "topspace (subtopology euclidean S) = S" +lemma topspace_euclidean_subtopology[simp]: "topspace (top_of_set S) = S" by (simp add: topspace_subtopology) lemma closed_closedin: "closed S \ closedin euclidean S" @@ -435,7 +434,7 @@ declare closed_closedin [symmetric, simp] -lemma openin_subtopology_self [simp]: "openin (subtopology euclidean S) S" +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\ @@ -445,51 +444,51 @@ subsection \Basic "localization" results are handy for connectedness.\ -lemma openin_open: "openin (subtopology euclidean U) S \ (\T. open T \ (S = U \ T))" +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 (subtopology euclidean U) S; open T\ - \ openin (subtopology euclidean U) (S \ T)" + "\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 (subtopology euclidean U) (U \ S)" +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 (subtopology euclidean S) T" + "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 (subtopology euclidean T) S" +lemma open_subset: "S \ T \ open S \ openin (top_of_set T) S" by (auto simp: openin_open) -lemma closedin_closed: "closedin (subtopology euclidean U) S \ (\T. closed T \ S = U \ T)" +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 (subtopology euclidean U) (U \ S)" +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 (subtopology euclidean T) S" +lemma closed_subset: "S \ T \ closed S \ closedin (top_of_set T) S" by (auto simp: closedin_closed) lemma closedin_closed_subset: - "\closedin (subtopology euclidean U) V; T \ U; S = V \ T\ - \ closedin (subtopology euclidean T) S" + "\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 (subtopology euclidean T) S" + 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 (subtopology euclidean U) {a} \ a \ U" + 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 (subtopology euclidean U) S \ + 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 @@ -513,16 +512,16 @@ lemma connected_openin: "connected S \ - \(\E1 E2. openin (subtopology euclidean S) E1 \ - openin (subtopology euclidean S) E2 \ + \(\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 *) lemma connected_openin_eq: "connected S \ - \(\E1 E2. openin (subtopology euclidean S) E1 \ - openin (subtopology euclidean S) E2 \ + \(\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) @@ -531,8 +530,8 @@ lemma connected_closedin: "connected S \ (\E1 E2. - closedin (subtopology euclidean S) E1 \ - closedin (subtopology euclidean S) E2 \ + closedin (top_of_set S) E1 \ + closedin (top_of_set S) E2 \ S \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" (is "?lhs = ?rhs") proof @@ -561,8 +560,8 @@ lemma connected_closedin_eq: "connected S \ \(\E1 E2. - closedin (subtopology euclidean S) E1 \ - closedin (subtopology euclidean S) 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) @@ -571,26 +570,26 @@ text \These "transitivity" results are handy too\ lemma openin_trans[trans]: - "openin (subtopology euclidean T) S \ openin (subtopology euclidean U) T \ - openin (subtopology euclidean U) S" + "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 (subtopology euclidean T) S \ open T \ open S" +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 (subtopology euclidean T) S \ closedin (subtopology euclidean U) T \ - closedin (subtopology euclidean U) S" + "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 (subtopology euclidean T) S \ closed T \ closed S" +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 (subtopology euclidean u) (u \ S); v \ u\ \ openin (subtopology euclidean v) (v \ S)" + "\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 (subtopology euclidean s) t \ open t \ t \ s)" +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 @@ -1684,11 +1683,11 @@ "continuous_map (discrete_topology U) X f \ f ` U \ topspace X" by (auto simp: continuous_map_def) -lemma continuous_map_iff_continuous_real [simp]: "continuous_map (subtopology euclideanreal S) euclideanreal g = continuous_on S g" - by (force simp: continuous_map openin_subtopology continuous_on_open_invariant) - -lemma continuous_map_iff_continuous_real2 [simp]: "continuous_map euclideanreal euclideanreal g = continuous_on UNIV g" - by (metis continuous_map_iff_continuous_real subtopology_UNIV) +lemma continuous_map_iff_continuous [simp]: "continuous_map (subtopology euclidean 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 \ @@ -2844,9 +2843,8 @@ done qed -lemma connectedin_iff_connected_real [simp]: - "connectedin euclideanreal S \ connected S" - by (simp add: connected_def connectedin) +lemma connectedin_iff_connected [simp]: "connectedin euclidean S \ connected S" + by (simp add: connected_def connectedin) lemma connectedin_closedin: "connectedin X S \ @@ -3190,7 +3188,7 @@ unfolding compact_space_alt using openin_subset by fastforce -lemma compactin_euclideanreal_iff [simp]: "compactin euclideanreal S \ compact S" +lemma compactin_euclidean_iff [simp]: "compactin euclidean S \ compact S" by (simp add: compact_eq_Heine_Borel compactin_def) meson lemma compactin_absolute [simp]: @@ -3639,36 +3637,127 @@ unfolding embedding_map_def using homeomorphic_space by blast + +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 topspace_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 (subtopology euclidean (f ` S)) T \ - openin (subtopology euclidean S) (S \ f -` T))" + (\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 (subtopology euclidean (f ` S)) T \ - closedin (subtopology euclidean S) (S \ f -` T))" + (\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 (subtopology euclidean (f ` S)) T" - shows "closedin (subtopology euclidean S) (S \ f -` T)" + 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 subsection%unimportant \Half-global and completely global cases\ lemma continuous_openin_preimage_gen: assumes "continuous_on S f" "open T" - shows "openin (subtopology euclidean S) (S \ f -` T)" + shows "openin (top_of_set S) (S \ f -` T)" proof - have *: "(S \ f -` T) = (S \ f -` (T \ f ` S))" by auto - have "openin (subtopology euclidean (f ` S)) (T \ f ` S)" + 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"]] @@ -3677,11 +3766,11 @@ lemma continuous_closedin_preimage: assumes "continuous_on S f" and "closed T" - shows "closedin (subtopology euclidean S) (S \ f -` T)" + shows "closedin (top_of_set S) (S \ f -` T)" proof - have *: "(S \ f -` T) = (S \ f -` (T \ f ` S))" by auto - have "closedin (subtopology euclidean (f ` S)) (T \ f ` S)" + 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 @@ -3691,7 +3780,7 @@ lemma continuous_openin_preimage_eq: "continuous_on S f \ - (\T. open T \ openin (subtopology euclidean S) (S \ f -` T))" + (\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) @@ -3699,7 +3788,7 @@ lemma continuous_closedin_preimage_eq: "continuous_on S f \ - (\T. closed T \ closedin (subtopology euclidean S) (S \ f -` T))" + (\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) @@ -3733,9 +3822,9 @@ by (simp add: closed_vimage continuous_on_eq_continuous_within) lemma Times_in_interior_subtopology: - assumes "(x, y) \ U" "openin (subtopology euclidean (S \ T)) U" - obtains V W where "openin (subtopology euclidean S) V" "x \ V" - "openin (subtopology euclidean T) W" "y \ W" "(V \ W) \ U" + 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) @@ -3744,8 +3833,8 @@ by blast show ?thesis proof - show "openin (subtopology euclidean S) (E1 \ S)" - "openin (subtopology euclidean T) (E2 \ T)" + 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" @@ -3757,20 +3846,20 @@ qed lemma closedin_Times: - "closedin (subtopology euclidean S) S' \ closedin (subtopology euclidean T) T' \ - closedin (subtopology euclidean (S \ T)) (S' \ T')" + "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 (subtopology euclidean S) S' \ openin (subtopology euclidean T) T' \ - openin (subtopology euclidean (S \ T)) (S' \ T')" + "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 (subtopology euclidean (S \ T)) (S' \ T') \ - S' = {} \ T' = {} \ openin (subtopology euclidean S) S' \ openin (subtopology euclidean T) T'" + "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 @@ -3782,13 +3871,13 @@ show ?thesis proof assume ?lhs - have "openin (subtopology euclidean S) S'" + 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 (subtopology euclidean T) T'" + moreover have "openin (top_of_set T) T'" apply (subst openin_subopen, clarify) apply (rule Times_in_interior_subtopology [OF _ \?lhs\]) using \x \ S'\ @@ -3805,7 +3894,7 @@ lemma Lim_transform_within_openin: assumes f: "(f \ l) (at a within T)" - and "openin (subtopology euclidean T) S" "a \ S" + 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 - @@ -3827,8 +3916,8 @@ lemma continuous_on_open_gen: assumes "f ` S \ T" shows "continuous_on S f \ - (\U. openin (subtopology euclidean T) U - \ openin (subtopology euclidean S) (S \ f -` U))" + (\U. openin (top_of_set T) U + \ openin (top_of_set S) (S \ f -` U))" (is "?lhs = ?rhs") proof assume ?lhs @@ -3841,23 +3930,23 @@ proof (clarsimp simp add: continuous_openin_preimage_eq) fix U::"'a set" assume "open U" - then have "openin (subtopology euclidean S) (S \ f -` (U \ T))" + then have "openin (top_of_set S) (S \ f -` (U \ T))" by (metis R inf_commute openin_open) - then show "openin (subtopology euclidean S) (S \ f -` U)" + 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 (subtopology euclidean T) U\ - \ openin (subtopology euclidean S) (S \ f -` U)" + "\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 (subtopology euclidean T) U - \ closedin (subtopology euclidean S) (S \ f -` U))" + (\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 @@ -3868,8 +3957,8 @@ show ?rhs proof clarify fix U - assume "closedin (subtopology euclidean T) U" - then show "closedin (subtopology euclidean S) (S \ f -` 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 @@ -3882,16 +3971,376 @@ qed lemma continuous_closedin_preimage_gen: - assumes "continuous_on S f" "f ` S \ T" "closedin (subtopology euclidean T) U" - shows "closedin (subtopology euclidean S) (S \ f -` U)" + 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 (subtopology euclidean T) S" "a \ S" + 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%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%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: + "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]) + +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 + +subsubsection%important \Continuity\ + +text \We will need to deal with continuous maps in terms of topologies and not in terms +of type classes, as defined below.\ + +definition%important continuous_on_topo::"'a topology \ 'b topology \ ('a \ 'b) \ bool" + where "continuous_on_topo T1 T2 f = ((\ U. openin T2 U \ openin T1 (f-`U \ topspace(T1))) + \ (f`(topspace T1) \ (topspace T2)))" + +lemma continuous_on_continuous_on_topo: + "continuous_on s f \ continuous_on_topo (top_of_set s) euclidean f" + by (auto simp: continuous_on_topo_def Int_commute continuous_openin_preimage_eq) + +lemma continuous_on_topo_UNIV: + "continuous_on UNIV f \ continuous_on_topo euclidean euclidean f" +using continuous_on_continuous_on_topo[of UNIV f] subtopology_UNIV[of euclidean] by auto + +lemma continuous_on_topo_open [intro]: + "continuous_on_topo T1 T2 f \ openin T2 U \ openin T1 (f-`U \ topspace(T1))" + unfolding continuous_on_topo_def by auto + +lemma continuous_on_topo_topspace [intro]: + "continuous_on_topo T1 T2 f \ f`(topspace T1) \ (topspace T2)" +unfolding continuous_on_topo_def by auto + +lemma continuous_on_generated_topo_iff: + "continuous_on_topo T1 (topology_generated_by S) f \ + ((\U. U \ S \ openin T1 (f-`U \ topspace(T1))) \ (f`(topspace T1) \ (\ S)))" +unfolding continuous_on_topo_def 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_on_topo T1 (topology_generated_by S) f" + using assms continuous_on_generated_topo_iff by blast + +proposition continuous_on_topo_compose: + assumes "continuous_on_topo T1 T2 f" "continuous_on_topo T2 T3 g" + shows "continuous_on_topo T1 T3 (g o f)" + using assms unfolding continuous_on_topo_def +proof (auto) + fix U :: "'c set" + assume H: "openin T3 U" + have "openin T1 (f -` (g -` U \ topspace T2) \ topspace T1)" + using H assms by blast + moreover have "f -` (g -` U \ topspace T2) \ topspace T1 = (g \ f) -` U \ topspace T1" + using H assms continuous_on_topo_topspace by fastforce + ultimately show "openin T1 ((g \ f) -` U \ topspace T1)" + by simp +qed (blast) + +lemma continuous_on_topo_preimage_topspace [intro]: + assumes "continuous_on_topo T1 T2 f" + shows "f-`(topspace T2) \ topspace T1 = topspace T1" +using assms unfolding continuous_on_topo_def by auto + + +subsubsection%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%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_on_topo_pullback [intro]: + assumes "continuous_on_topo T1 T2 g" + shows "continuous_on_topo (pullback_topology A f T1) T2 (g o f)" +unfolding continuous_on_topo_def +proof (auto) + fix U::"'b set" assume "openin T2 U" + then have "openin T1 (g-`U \ topspace T1)" + using assms unfolding continuous_on_topo_def 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_on_topo_def by auto +qed + +proposition continuous_on_topo_pullback' [intro]: + assumes "continuous_on_topo T1 T2 (f o g)" "topspace T1 \ g-`A" + shows "continuous_on_topo T1 (pullback_topology A f T2) g" +unfolding continuous_on_topo_def +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_on_topo_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 + end diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Abstract_Topology_2.thy --- a/src/HOL/Analysis/Abstract_Topology_2.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Abstract_Topology_2.thy Tue Mar 19 16:14:51 2019 +0000 @@ -99,14 +99,14 @@ qed lemma closedin_limpt: - "closedin (subtopology euclidean T) S \ S \ T \ (\x. x islimpt S \ x \ T \ x \ S)" + "closedin (top_of_set T) S \ S \ T \ (\x. x islimpt S \ x \ T \ x \ S)" apply (simp add: closedin_closed, safe) apply (simp add: closed_limpt islimpt_subset) apply (rule_tac x="closure S" in exI, simp) apply (force simp: closure_def) done -lemma closedin_closed_eq: "closed S \ closedin (subtopology euclidean S) T \ closed T \ T \ S" +lemma closedin_closed_eq: "closed S \ closedin (top_of_set S) T \ closed T \ T \ S" by (meson closedin_limpt closed_subset closedin_closed_trans) lemma connected_closed_set: @@ -123,32 +123,35 @@ by (metis assms closed_Un connected_closed_set) lemma closedin_subset_trans: - "closedin (subtopology euclidean U) S \ S \ T \ T \ U \ - closedin (subtopology euclidean T) S" + "closedin (top_of_set U) S \ S \ T \ T \ U \ + closedin (top_of_set T) S" by (meson closedin_limpt subset_iff) lemma openin_subset_trans: - "openin (subtopology euclidean U) S \ S \ T \ T \ U \ - openin (subtopology euclidean T) S" + "openin (top_of_set U) S \ S \ T \ T \ U \ + openin (top_of_set T) S" by (auto simp: openin_open) lemma closedin_compact: - "\compact S; closedin (subtopology euclidean S) T\ \ compact T" + "\compact S; closedin (top_of_set S) T\ \ compact T" by (metis closedin_closed compact_Int_closed) lemma closedin_compact_eq: fixes S :: "'a::t2_space set" shows "compact S - \ (closedin (subtopology euclidean S) T \ + \ (closedin (top_of_set S) T \ compact T \ T \ S)" by (metis closedin_imp_subset closedin_compact closed_subset compact_imp_closed) subsection \Closure\ +lemma euclidean_closure_of [simp]: "euclidean closure_of S = closure S" + by (auto simp: closure_of_def closure_def islimpt_def) + lemma closure_openin_Int_closure: - assumes ope: "openin (subtopology euclidean U) S" and "T \ U" + assumes ope: "openin (top_of_set U) S" and "T \ U" shows "closure(S \ closure T) = closure(S \ T)" proof obtain V where "open V" and S: "S = U \ V" @@ -171,16 +174,16 @@ corollary infinite_openin: fixes S :: "'a :: t1_space set" - shows "\openin (subtopology euclidean U) S; x \ S; x islimpt U\ \ infinite S" + shows "\openin (top_of_set U) S; x \ S; x islimpt U\ \ infinite S" by (clarsimp simp add: openin_open islimpt_eq_acc_point inf_commute) lemma closure_Int_ballI: - assumes "\U. \openin (subtopology euclidean S) U; U \ {}\ \ T \ U \ {}" + assumes "\U. \openin (top_of_set S) U; U \ {}\ \ T \ U \ {}" shows "S \ closure T" proof (clarsimp simp: closure_iff_nhds_not_empty) fix x and A and V assume "x \ S" "V \ A" "open V" "x \ V" "T \ A = {}" - then have "openin (subtopology euclidean S) (A \ V \ S)" + then have "openin (top_of_set S) (A \ V \ S)" by (auto simp: openin_open intro!: exI[where x="V"]) moreover have "A \ V \ S \ {}" using \x \ V\ \V \ A\ \x \ S\ by auto @@ -192,6 +195,12 @@ subsection \Frontier\ +lemma euclidean_interior_of [simp]: "euclidean interior_of S = interior S" + by (auto simp: interior_of_def interior_def) + +lemma euclidean_frontier_of [simp]: "euclidean frontier_of S = frontier S" + by (auto simp: frontier_of_def frontier_def) + lemma connected_Int_frontier: "\connected s; s \ t \ {}; s - t \ {}\ \ (s \ frontier t \ {})" apply (simp add: frontier_interiors connected_openin, safe) @@ -204,17 +213,17 @@ lemma openin_delete: fixes a :: "'a :: t1_space" - shows "openin (subtopology euclidean u) s - \ openin (subtopology euclidean u) (s - {a})" + shows "openin (top_of_set u) s + \ openin (top_of_set u) (s - {a})" by (metis Int_Diff open_delete openin_open) lemma compact_eq_openin_cover: "compact S \ - (\C. (\c\C. openin (subtopology euclidean S) c) \ S \ \C \ + (\C. (\c\C. openin (top_of_set S) c) \ S \ \C \ (\D\C. finite D \ S \ \D))" proof safe fix C - assume "compact S" and "\c\C. openin (subtopology euclidean S) c" and "S \ \C" + assume "compact S" and "\c\C. openin (top_of_set S) c" and "S \ \C" then have "\c\{T. open T \ S \ T \ C}. open c" and "S \ \{T. open T \ S \ T \ C}" unfolding openin_open by force+ with \compact S\ obtain D where "D \ {T. open T \ S \ T \ C}" and "finite D" and "S \ \D" @@ -223,14 +232,14 @@ by auto then show "\D\C. finite D \ S \ \D" .. next - assume 1: "\C. (\c\C. openin (subtopology euclidean S) c) \ S \ \C \ + assume 1: "\C. (\c\C. openin (top_of_set S) c) \ S \ \C \ (\D\C. finite D \ S \ \D)" show "compact S" proof (rule compactI) fix C let ?C = "image (\T. S \ T) C" assume "\t\C. open t" and "S \ \C" - then have "(\c\?C. openin (subtopology euclidean S) c) \ S \ \?C" + then have "(\c\?C. openin (top_of_set S) c) \ S \ \?C" unfolding openin_open by auto with 1 obtain D where "D \ ?C" and "finite D" and "S \ \D" by metis @@ -276,7 +285,7 @@ lemma continuous_closedin_preimage_constant: fixes f :: "_ \ 'b::t1_space" - shows "continuous_on S f \ closedin (subtopology euclidean S) {x \ S. f x = a}" + shows "continuous_on S f \ closedin (top_of_set S) {x \ S. f x = a}" using continuous_closedin_preimage[of S f "{a}"] by (simp add: vimage_def Collect_conj_eq) lemma continuous_closed_preimage_constant: @@ -335,14 +344,14 @@ subsection%unimportant \Continuity relative to a union.\ lemma continuous_on_Un_local: - "\closedin (subtopology euclidean (s \ t)) s; closedin (subtopology euclidean (s \ t)) t; + "\closedin (top_of_set (s \ t)) s; closedin (top_of_set (s \ t)) t; continuous_on s f; continuous_on t f\ \ continuous_on (s \ t) f" unfolding continuous_on closedin_limpt by (metis Lim_trivial_limit Lim_within_union Un_iff trivial_limit_within) lemma continuous_on_cases_local: - "\closedin (subtopology euclidean (s \ t)) s; closedin (subtopology euclidean (s \ t)) t; + "\closedin (top_of_set (s \ t)) s; closedin (top_of_set (s \ t)) t; continuous_on s f; continuous_on t g; \x. \x \ s \ \P x \ x \ t \ P x\ \ f x = g x\ \ continuous_on (s \ t) (\x. if P x then f x else g x)" @@ -358,9 +367,9 @@ proof - have s: "s = (s \ h -` atMost a) \ (s \ h -` atLeast a)" by force - have 1: "closedin (subtopology euclidean s) (s \ h -` atMost a)" + have 1: "closedin (top_of_set s) (s \ h -` atMost a)" by (rule continuous_closedin_preimage [OF h closed_atMost]) - have 2: "closedin (subtopology euclidean s) (s \ h -` atLeast a)" + have 2: "closedin (top_of_set s) (s \ h -` atLeast a)" by (rule continuous_closedin_preimage [OF h closed_atLeast]) have eq: "s \ h -` {..a} = {t \ s. h t \ a}" "s \ h -` {a..} = {t \ s. a \ h t}" by auto @@ -388,7 +397,7 @@ assumes contf: "continuous_on S f" and imf: "f ` S = T" and injf: "\x. x \ S \ g (f x) = x" - and oo: "\U. openin (subtopology euclidean S) U \ openin (subtopology euclidean T) (f ` U)" + and oo: "\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)" shows "continuous_on T g" proof - from imf injf have gTS: "g ` T = S" @@ -403,7 +412,7 @@ assumes contf: "continuous_on S f" and imf: "f ` S = T" and injf: "\x. x \ S \ g(f x) = x" - and oo: "\U. closedin (subtopology euclidean S) U \ closedin (subtopology euclidean T) (f ` U)" + and oo: "\U. closedin (top_of_set S) U \ closedin (top_of_set T) (f ` U)" shows "continuous_on T g" proof - from imf injf have gTS: "g ` T = S" @@ -418,7 +427,7 @@ assumes contf: "continuous_on S f" and imf: "f ` S = T" and injf: "inj_on f S" - and oo: "\U. openin (subtopology euclidean S) U \ openin (subtopology euclidean T) (f ` U)" + and oo: "\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)" obtains g where "homeomorphism S T f g" proof have "continuous_on T (inv_into S f)" @@ -431,7 +440,7 @@ assumes contf: "continuous_on S f" and imf: "f ` S = T" and injf: "inj_on f S" - and oo: "\U. closedin (subtopology euclidean S) U \ closedin (subtopology euclidean T) (f ` U)" + and oo: "\U. closedin (top_of_set S) U \ closedin (top_of_set T) (f ` U)" obtains g where "homeomorphism S T f g" proof have "continuous_on T (inv_into S f)" @@ -442,8 +451,8 @@ lemma homeomorphism_imp_open_map: assumes hom: "homeomorphism S T f g" - and oo: "openin (subtopology euclidean S) U" - shows "openin (subtopology euclidean T) (f ` U)" + and oo: "openin (top_of_set S) U" + shows "openin (top_of_set T) (f ` U)" proof - from hom oo have [simp]: "f ` U = T \ g -` U" using openin_subset by (fastforce simp: homeomorphism_def rev_image_eqI) @@ -457,8 +466,8 @@ lemma homeomorphism_imp_closed_map: assumes hom: "homeomorphism S T f g" - and oo: "closedin (subtopology euclidean S) U" - shows "closedin (subtopology euclidean T) (f ` U)" + and oo: "closedin (top_of_set S) U" + shows "closedin (top_of_set T) (f ` U)" proof - from hom oo have [simp]: "f ` U = T \ g -` U" using closedin_subset by (fastforce simp: homeomorphism_def rev_image_eqI) @@ -476,13 +485,13 @@ obtains \ :: "'a:: second_countable_topology set set" where "countable \" "{} \ \" - "\C. C \ \ \ openin(subtopology euclidean S) C" - "\T. openin(subtopology euclidean S) T \ \\. \ \ \ \ T = \\" + "\C. C \ \ \ openin(top_of_set S) C" + "\T. openin(top_of_set S) T \ \\. \ \ \ \ T = \\" proof - obtain \ :: "'a set set" where "countable \" - and opeB: "\C. C \ \ \ openin(subtopology euclidean S) C" - and \: "\T. openin(subtopology euclidean S) T \ \\. \ \ \ \ T = \\" + and opeB: "\C. C \ \ \ openin(top_of_set S) C" + and \: "\T. openin(top_of_set S) T \ \\. \ \ \ \ T = \\" proof - obtain \ :: "'a set set" where "countable \" and ope: "\C. C \ \ \ open C" @@ -492,9 +501,9 @@ proof show "countable ((\C. S \ C) ` \)" by (simp add: \countable \\) - show "\C. C \ (\) S ` \ \ openin (subtopology euclidean S) C" + show "\C. C \ (\) S ` \ \ openin (top_of_set S) C" using ope by auto - show "\T. openin (subtopology euclidean S) T \ \\\(\) S ` \. T = \\" + show "\T. openin (top_of_set S) T \ \\\(\) S ` \. T = \\" by (metis \ image_mono inf_Sup openin_open) qed qed @@ -502,9 +511,9 @@ proof show "countable (\ - {{}})" using \countable \\ by blast - show "\C. \C \ \ - {{}}\ \ openin (subtopology euclidean S) C" - by (simp add: \\C. C \ \ \ openin (subtopology euclidean S) C\) - show "\\\\ - {{}}. T = \\" if "openin (subtopology euclidean S) T" for T + show "\C. \C \ \ - {{}}\ \ openin (top_of_set S) C" + by (simp add: \\C. C \ \ \ openin (top_of_set S) C\) + show "\\\\ - {{}}. T = \\" if "openin (top_of_set S) T" for T using \ [OF that] apply clarify apply (rule_tac x="\ - {{}}" in exI, auto) @@ -514,7 +523,7 @@ lemma Lindelof_openin: fixes \ :: "'a::second_countable_topology set set" - assumes "\S. S \ \ \ openin (subtopology euclidean U) S" + assumes "\S. S \ \ \ openin (top_of_set U) S" obtains \' where "\' \ \" "countable \'" "\\' = \\" proof - have "\S. S \ \ \ \T. open T \ S = U \ T" @@ -535,16 +544,16 @@ lemma continuous_imp_closed_map: fixes f :: "'a::t2_space \ 'b::t2_space" - assumes "closedin (subtopology euclidean S) U" + assumes "closedin (top_of_set S) U" "continuous_on S f" "f ` S = T" "compact S" - shows "closedin (subtopology euclidean T) (f ` U)" + shows "closedin (top_of_set T) (f ` U)" by (metis assms closedin_compact_eq compact_continuous_image continuous_on_subset subset_image_iff) lemma closed_map_restrict: - assumes cloU: "closedin (subtopology euclidean (S \ f -` T')) U" - and cc: "\U. closedin (subtopology euclidean S) U \ closedin (subtopology euclidean T) (f ` U)" + assumes cloU: "closedin (top_of_set (S \ f -` T')) U" + and cc: "\U. closedin (top_of_set S) U \ closedin (top_of_set T) (f ` U)" and "T' \ T" - shows "closedin (subtopology euclidean T') (f ` U)" + shows "closedin (top_of_set T') (f ` U)" proof - obtain V where "closed V" "U = S \ f -` T' \ V" using cloU by (auto simp: closedin_closed) @@ -555,10 +564,10 @@ subsection%unimportant\Open Maps\ lemma open_map_restrict: - assumes opeU: "openin (subtopology euclidean (S \ f -` T')) U" - and oo: "\U. openin (subtopology euclidean S) U \ openin (subtopology euclidean T) (f ` U)" + assumes opeU: "openin (top_of_set (S \ f -` T')) U" + and oo: "\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)" and "T' \ T" - shows "openin (subtopology euclidean T') (f ` U)" + shows "openin (top_of_set T') (f ` U)" proof - obtain V where "open V" "U = S \ f -` T' \ V" using opeU by (auto simp: openin_open) @@ -572,8 +581,8 @@ lemma quotient_map_imp_continuous_open: assumes T: "f ` S \ T" and ope: "\U. U \ T - \ (openin (subtopology euclidean S) (S \ f -` U) \ - openin (subtopology euclidean T) U)" + \ (openin (top_of_set S) (S \ f -` U) \ + openin (top_of_set T) U)" shows "continuous_on S f" proof - have [simp]: "S \ f -` f ` S = S" by auto @@ -586,8 +595,8 @@ lemma quotient_map_imp_continuous_closed: assumes T: "f ` S \ T" and ope: "\U. U \ T - \ (closedin (subtopology euclidean S) (S \ f -` U) \ - closedin (subtopology euclidean T) U)" + \ (closedin (top_of_set S) (S \ f -` U) \ + closedin (top_of_set T) U)" shows "continuous_on S f" proof - have [simp]: "S \ f -` f ` S = S" by auto @@ -600,10 +609,10 @@ lemma open_map_imp_quotient_map: assumes contf: "continuous_on S f" and T: "T \ f ` S" - and ope: "\T. openin (subtopology euclidean S) T - \ openin (subtopology euclidean (f ` S)) (f ` T)" - shows "openin (subtopology euclidean S) (S \ f -` T) = - openin (subtopology euclidean (f ` S)) T" + and ope: "\T. openin (top_of_set S) T + \ openin (top_of_set (f ` S)) (f ` T)" + shows "openin (top_of_set S) (S \ f -` T) = + openin (top_of_set (f ` S)) T" proof - have "T = f ` (S \ f -` T)" using T by blast @@ -614,14 +623,14 @@ lemma closed_map_imp_quotient_map: assumes contf: "continuous_on S f" and T: "T \ f ` S" - and ope: "\T. closedin (subtopology euclidean S) T - \ closedin (subtopology euclidean (f ` S)) (f ` T)" - shows "openin (subtopology euclidean S) (S \ f -` T) \ - openin (subtopology euclidean (f ` S)) T" + and ope: "\T. closedin (top_of_set S) T + \ closedin (top_of_set (f ` S)) (f ` T)" + shows "openin (top_of_set S) (S \ f -` T) \ + openin (top_of_set (f ` S)) T" (is "?lhs = ?rhs") proof assume ?lhs - then have *: "closedin (subtopology euclidean S) (S - (S \ f -` T))" + then have *: "closedin (top_of_set S) (S - (S \ f -` T))" using closedin_diff by fastforce have [simp]: "(f ` S - f ` (S - (S \ f -` T))) = T" using T by blast @@ -638,14 +647,14 @@ and contg: "continuous_on T g" and img: "g ` T \ S" and fg [simp]: "\y. y \ T \ f(g y) = y" and U: "U \ T" - shows "openin (subtopology euclidean S) (S \ f -` U) \ - openin (subtopology euclidean T) U" + shows "openin (top_of_set S) (S \ f -` U) \ + openin (top_of_set T) U" (is "?lhs = ?rhs") proof - - have f: "\Z. openin (subtopology euclidean (f ` S)) Z \ - openin (subtopology euclidean S) (S \ f -` Z)" - and g: "\Z. openin (subtopology euclidean (g ` T)) Z \ - openin (subtopology euclidean T) (T \ g -` Z)" + have f: "\Z. openin (top_of_set (f ` S)) Z \ + openin (top_of_set S) (S \ f -` Z)" + and g: "\Z. openin (top_of_set (g ` T)) Z \ + openin (top_of_set T) (T \ g -` Z)" using contf contg by (auto simp: continuous_on_open) show ?thesis proof @@ -655,7 +664,7 @@ using U by auto finally have eq: "T \ g -` (g ` T \ (S \ f -` U)) = U" . assume ?lhs - then have *: "openin (subtopology euclidean (g ` T)) (g ` T \ (S \ f -` U))" + then have *: "openin (top_of_set (g ` T)) (g ` T \ (S \ f -` U))" by (meson img openin_Int openin_subtopology_Int_subset openin_subtopology_self) show ?rhs using g [OF *] eq by auto @@ -671,8 +680,8 @@ and "continuous_on (f ` S) g" and "\x. x \ S \ g(f x) = x" and "U \ f ` S" - shows "openin (subtopology euclidean S) (S \ f -` U) \ - openin (subtopology euclidean (f ` S)) U" + shows "openin (top_of_set S) (S \ f -` U) \ + openin (top_of_set (f ` S)) U" apply (rule continuous_right_inverse_imp_quotient_map) using assms apply force+ done @@ -680,101 +689,220 @@ lemma continuous_imp_quotient_map: fixes f :: "'a::t2_space \ 'b::t2_space" assumes "continuous_on S f" "f ` S = T" "compact S" "U \ T" - shows "openin (subtopology euclidean S) (S \ f -` U) \ - openin (subtopology euclidean T) U" + shows "openin (top_of_set S) (S \ f -` U) \ + openin (top_of_set T) U" by (metis (no_types, lifting) assms closed_map_imp_quotient_map continuous_imp_closed_map) -subsection%unimportant\Pasting functions together\ +subsection%unimportant\Pasting lemmas for functions, for of casewise definitions\ -text\on open sets\ +subsubsection\on open sets\ lemma pasting_lemma: - fixes f :: "'i \ 'a::topological_space \ 'b::topological_space" - assumes clo: "\i. i \ I \ openin (subtopology euclidean S) (T i)" - and cont: "\i. i \ I \ continuous_on (T i) (f i)" - and f: "\i j x. \i \ I; j \ I; x \ S \ T i \ T j\ \ f i x = f j x" - and g: "\x. x \ S \ \j. j \ I \ x \ T j \ g x = f j x" - shows "continuous_on S g" -proof (clarsimp simp: continuous_openin_preimage_eq) - fix U :: "'b set" - assume "open U" - have S: "\i. i \ I \ (T i) \ S" - using clo openin_imp_subset by blast - have *: "(S \ g -` U) = (\i \ I. T i \ f i -` U)" - using S f g by fastforce - show "openin (subtopology euclidean S) (S \ g -` U)" - apply (subst *) - apply (rule openin_Union, clarify) - using \open U\ clo cont continuous_openin_preimage_gen openin_trans by blast + assumes ope: "\i. i \ I \ openin X (T i)" + and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" + and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" + and g: "\x. x \ topspace X \ \j. j \ I \ x \ T j \ g x = f j x" + shows "continuous_map X Y g" + unfolding continuous_map_openin_preimage_eq +proof (intro conjI allI impI) + show "g ` topspace X \ topspace Y" + using g cont continuous_map_image_subset_topspace topspace_subtopology by fastforce +next + fix U + assume Y: "openin Y U" + have T: "T i \ topspace X" if "i \ I" for i + using ope by (simp add: openin_subset that) + have *: "topspace X \ g -` U = (\i \ I. T i \ f i -` U)" + using f g T by fastforce + have "\i. i \ I \ openin X (T i \ f i -` U)" + using cont unfolding continuous_map_openin_preimage_eq + by (metis Y T inf.commute inf_absorb1 ope topspace_subtopology openin_trans_full) + then show "openin X (topspace X \ g -` U)" + by (auto simp: *) qed lemma pasting_lemma_exists: - fixes f :: "'i \ 'a::topological_space \ 'b::topological_space" - assumes S: "S \ (\i \ I. T i)" - and clo: "\i. i \ I \ openin (subtopology euclidean S) (T i)" - and cont: "\i. i \ I \ continuous_on (T i) (f i)" - and f: "\i j x. \i \ I; j \ I; x \ S \ T i \ T j\ \ f i x = f j x" - obtains g where "continuous_on S g" "\x i. \i \ I; x \ S \ T i\ \ g x = f i x" + assumes X: "topspace X \ (\i \ I. T i)" + and ope: "\i. i \ I \ openin X (T i)" + and cont: "\i. i \ I \ continuous_map (subtopology X (T i)) Y (f i)" + and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" + obtains g where "continuous_map X Y g" "\x i. \i \ I; x \ topspace X \ T i\ \ g x = f i x" proof - show "continuous_on S (\x. f (SOME i. i \ I \ x \ T i) x)" - apply (rule pasting_lemma [OF clo cont]) + let ?h = "\x. f (SOME i. i \ I \ x \ T i) x" + show "continuous_map X Y ?h" + apply (rule pasting_lemma [OF ope cont]) apply (blast intro: f)+ - apply (metis (mono_tags, lifting) S UN_iff subsetCE someI) + by (metis (no_types, lifting) UN_E X subsetD someI_ex) + show "f (SOME i. i \ I \ x \ T i) x = f i x" if "i \ I" "x \ topspace X \ T i" for i x + by (metis (no_types, lifting) IntD2 IntI f someI_ex that) +qed + +lemma pasting_lemma_locally_finite: + assumes fin: "\x. x \ topspace X \ \V. openin X V \ x \ V \ finite {i \ I. T i \ V \ {}}" + and clo: "\i. i \ I \ closedin X (T i)" + and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" + and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" + and g: "\x. x \ topspace X \ \j. j \ I \ x \ T j \ g x = f j x" + shows "continuous_map X Y g" + unfolding continuous_map_closedin_preimage_eq +proof (intro conjI allI impI) + show "g ` topspace X \ topspace Y" + using g cont continuous_map_image_subset_topspace topspace_subtopology by fastforce +next + fix U + assume Y: "closedin Y U" + have T: "T i \ topspace X" if "i \ I" for i + using clo by (simp add: closedin_subset that) + have *: "topspace X \ g -` U = (\i \ I. T i \ f i -` U)" + using f g T by fastforce + have cTf: "\i. i \ I \ closedin X (T i \ f i -` U)" + using cont unfolding continuous_map_closedin_preimage_eq topspace_subtopology + by (simp add: Int_absorb1 T Y clo closedin_closed_subtopology) + have sub: "{Z \ (\i. T i \ f i -` U) ` I. Z \ V \ {}} + \ (\i. T i \ f i -` U) ` {i \ I. T i \ V \ {}}" for V + by auto + have 1: "(\i\I. T i \ f i -` U) \ topspace X" + using T by blast + then have lf: "locally_finite_in X ((\i. T i \ f i -` U) ` I)" + unfolding locally_finite_in_def + using finite_subset [OF sub] fin by force + show "closedin X (topspace X \ g -` U)" + apply (subst *) + apply (rule closedin_locally_finite_Union) + apply (auto intro: cTf lf) done +qed + +subsubsection\Likewise on closed sets, with a finiteness assumption\ + +lemma pasting_lemma_closed: + assumes fin: "finite I" + and clo: "\i. i \ I \ closedin X (T i)" + and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" + and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" + and g: "\x. x \ topspace X \ \j. j \ I \ x \ T j \ g x = f j x" + shows "continuous_map X Y g" + using pasting_lemma_locally_finite [OF _ clo cont f g] fin by auto + +lemma pasting_lemma_exists_locally_finite: + assumes fin: "\x. x \ topspace X \ \V. openin X V \ x \ V \ finite {i \ I. T i \ V \ {}}" + and X: "topspace X \ \(T ` I)" + and clo: "\i. i \ I \ closedin X (T i)" + and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" + and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" + and g: "\x. x \ topspace X \ \j. j \ I \ x \ T j \ g x = f j x" + obtains g where "continuous_map X Y g" "\x i. \i \ I; x \ topspace X \ T i\ \ g x = f i x" +proof + show "continuous_map X Y (\x. f(@i. i \ I \ x \ T i) x)" + apply (rule pasting_lemma_locally_finite [OF fin]) + apply (blast intro: assms)+ + by (metis (no_types, lifting) UN_E X set_rev_mp someI_ex) next fix x i - assume "i \ I" "x \ S \ T i" + assume "i \ I" and "x \ topspace X \ T i" + show "f (SOME i. i \ I \ x \ T i) x = f i x" + apply (rule someI2_ex) + using \i \ I\ \x \ topspace X \ T i\ apply blast + by (meson Int_iff \i \ I\ \x \ topspace X \ T i\ f) +qed + +lemma pasting_lemma_exists_closed: + assumes fin: "finite I" + and X: "topspace X \ \(T ` I)" + and clo: "\i. i \ I \ closedin X (T i)" + and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" + and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" + obtains g where "continuous_map X Y g" "\x i. \i \ I; x \ topspace X \ T i\ \ g x = f i x" +proof + show "continuous_map X Y (\x. f (SOME i. i \ I \ x \ T i) x)" + apply (rule pasting_lemma_closed [OF \finite I\ clo cont]) + apply (blast intro: f)+ + by (metis (mono_tags, lifting) UN_iff X someI_ex subset_iff) +next + fix x i + assume "i \ I" "x \ topspace X \ T i" then show "f (SOME i. i \ I \ x \ T i) x = f i x" by (metis (no_types, lifting) IntD2 IntI f someI_ex) qed -text\Likewise on closed sets, with a finiteness assumption\ +lemma continuous_map_cases: + assumes f: "continuous_map (subtopology X (X closure_of {x. P x})) Y f" + and g: "continuous_map (subtopology X (X closure_of {x. \ P x})) Y g" + and fg: "\x. x \ X frontier_of {x. P x} \ f x = g x" + shows "continuous_map X Y (\x. if P x then f x else g x)" +proof (rule pasting_lemma_closed) + let ?f = "\b. if b then f else g" + let ?g = "\x. if P x then f x else g x" + let ?T = "\b. if b then X closure_of {x. P x} else X closure_of {x. ~P x}" + show "finite {True,False}" by auto + have eq: "topspace X - Collect P = topspace X \ {x. \ P x}" + by blast + show "?f i x = ?f j x" + if "i \ {True,False}" "j \ {True,False}" and x: "x \ topspace X \ ?T i \ ?T j" for i j x + proof - + have "f x = g x" + if "i" "\ j" + apply (rule fg) + unfolding frontier_of_closures eq + using x that closure_of_restrict by fastforce + moreover + have "g x = f x" + if "x \ X closure_of {x. \ P x}" "x \ X closure_of Collect P" "\ i" "j" for x + apply (rule fg [symmetric]) + unfolding frontier_of_closures eq + using x that closure_of_restrict by fastforce + ultimately show ?thesis + using that by (auto simp flip: closure_of_restrict) + qed + show "\j. j \ {True,False} \ x \ ?T j \ (if P x then f x else g x) = ?f j x" + if "x \ topspace X" for x + apply simp + apply safe + apply (metis Int_iff closure_of inf_sup_absorb mem_Collect_eq that) + by (metis DiffI eq closure_of_subset_Int contra_subsetD mem_Collect_eq that) +qed (auto simp: f g) -lemma pasting_lemma_closed: - fixes f :: "'i \ 'a::topological_space \ 'b::topological_space" - assumes "finite I" - and clo: "\i. i \ I \ closedin (subtopology euclidean S) (T i)" - and cont: "\i. i \ I \ continuous_on (T i) (f i)" - and f: "\i j x. \i \ I; j \ I; x \ S \ T i \ T j\ \ f i x = f j x" - and g: "\x. x \ S \ \j. j \ I \ x \ T j \ g x = f j x" - shows "continuous_on S g" -proof (clarsimp simp: continuous_closedin_preimage_eq) - fix U :: "'b set" - assume "closed U" - have S: "\i. i \ I \ (T i) \ S" - using clo closedin_imp_subset by blast - have *: "(S \ g -` U) = (\i \ I. T i \ f i -` U)" - using S f g by fastforce - show "closedin (subtopology euclidean S) (S \ g -` U)" - apply (subst *) - apply (rule closedin_Union) - using \finite I\ apply simp - apply (blast intro: \closed U\ continuous_closedin_preimage cont clo closedin_trans) - done +lemma continuous_map_cases_alt: + assumes f: "continuous_map (subtopology X (X closure_of {x \ topspace X. P x})) Y f" + and g: "continuous_map (subtopology X (X closure_of {x \ topspace X. ~P x})) Y g" + and fg: "\x. x \ X frontier_of {x \ topspace X. P x} \ f x = g x" + shows "continuous_map X Y (\x. if P x then f x else g x)" + apply (rule continuous_map_cases) + using assms + apply (simp_all add: Collect_conj_eq closure_of_restrict [symmetric] frontier_of_restrict [symmetric]) + done + +lemma continuous_map_cases_function: + assumes contp: "continuous_map X Z p" + and contf: "continuous_map (subtopology X {x \ topspace X. p x \ Z closure_of U}) Y f" + and contg: "continuous_map (subtopology X {x \ topspace X. p x \ Z closure_of (topspace Z - U)}) Y g" + and fg: "\x. \x \ topspace X; p x \ Z frontier_of U\ \ f x = g x" + shows "continuous_map X Y (\x. if p x \ U then f x else g x)" +proof (rule continuous_map_cases_alt) + show "continuous_map (subtopology X (X closure_of {x \ topspace X. p x \ U})) Y f" + proof (rule continuous_map_from_subtopology_mono) + let ?T = "{x \ topspace X. p x \ Z closure_of U}" + show "continuous_map (subtopology X ?T) Y f" + by (simp add: contf) + show "X closure_of {x \ topspace X. p x \ U} \ ?T" + by (rule continuous_map_closure_preimage_subset [OF contp]) + qed + show "continuous_map (subtopology X (X closure_of {x \ topspace X. p x \ U})) Y g" + proof (rule continuous_map_from_subtopology_mono) + let ?T = "{x \ topspace X. p x \ Z closure_of (topspace Z - U)}" + show "continuous_map (subtopology X ?T) Y g" + by (simp add: contg) + have "X closure_of {x \ topspace X. p x \ U} \ X closure_of {x \ topspace X. p x \ topspace Z - U}" + apply (rule closure_of_mono) + using continuous_map_closedin contp by fastforce + then show "X closure_of {x \ topspace X. p x \ U} \ ?T" + by (rule order_trans [OF _ continuous_map_closure_preimage_subset [OF contp]]) + qed +next + show "f x = g x" if "x \ X frontier_of {x \ topspace X. p x \ U}" for x + using that continuous_map_frontier_frontier_preimage_subset [OF contp, of U] fg by blast qed -lemma pasting_lemma_exists_closed: - fixes f :: "'i \ 'a::topological_space \ 'b::topological_space" - assumes "finite I" - and S: "S \ (\i \ I. T i)" - and clo: "\i. i \ I \ closedin (subtopology euclidean S) (T i)" - and cont: "\i. i \ I \ continuous_on (T i) (f i)" - and f: "\i j x. \i \ I; j \ I; x \ S \ T i \ T j\ \ f i x = f j x" - obtains g where "continuous_on S g" "\x i. \i \ I; x \ S \ T i\ \ g x = f i x" -proof - show "continuous_on S (\x. f (SOME i. i \ I \ x \ T i) x)" - apply (rule pasting_lemma_closed [OF \finite I\ clo cont]) - apply (blast intro: f)+ - apply (metis (mono_tags, lifting) S UN_iff subsetCE someI) - done -next - fix x i - assume "i \ I" "x \ S \ T i" - then show "f (SOME i. i \ I \ x \ T i) x = f i x" - by (metis (no_types, lifting) IntD2 IntI f someI_ex) -qed - - subsection \Retractions\ definition%important retraction :: "('a::topological_space) set \ 'a set \ ('a \ 'a) \ bool" @@ -948,7 +1076,7 @@ lemma closedin_retract: fixes S :: "'a :: t2_space set" assumes "S retract_of T" - shows "closedin (subtopology euclidean T) S" + shows "closedin (top_of_set T) S" proof - obtain r where r: "S \ T" "continuous_on T r" "r ` T \ S" "\x. x \ S \ r x = x" using assms by (auto simp: retract_of_def retraction_def) @@ -963,7 +1091,7 @@ finally show ?thesis . qed -lemma closedin_self [simp]: "closedin (subtopology euclidean S) S" +lemma closedin_self [simp]: "closedin (top_of_set S) S" by simp lemma retract_of_closed: @@ -980,7 +1108,7 @@ by (metis Topological_Spaces.connected_continuous_image retract_of_def retraction) lemma retraction_imp_quotient_map: - "openin (subtopology euclidean S) (S \ r -` U) \ openin (subtopology euclidean T) U" + "openin (top_of_set S) (S \ r -` U) \ openin (top_of_set T) U" if retraction: "retraction S T r" and "U \ T" using retraction apply (rule retractionE) apply (rule continuous_right_inverse_imp_quotient_map [where g=r]) @@ -995,4 +1123,472 @@ apply (rule conjI continuous_intros | erule continuous_on_subset | force)+ done +subsection\Retractions on a topological space\ + +definition retract_of_space :: "'a set \ 'a topology \ bool" (infix "retract'_of'_space" 50) + where "S retract_of_space X + \ S \ topspace X \ (\r. continuous_map X (subtopology X S) r \ (\x \ S. r x = x))" + +lemma retract_of_space_retraction_maps: + "S retract_of_space X \ S \ topspace X \ (\r. retraction_maps X (subtopology X S) r id)" + by (auto simp: retract_of_space_def retraction_maps_def) + +lemma retract_of_space_section_map: + "S retract_of_space X \ S \ topspace X \ section_map (subtopology X S) X id" + unfolding retract_of_space_def retraction_maps_def section_map_def + by (auto simp: continuous_map_from_subtopology) + +lemma retract_of_space_imp_subset: + "S retract_of_space X \ S \ topspace X" + by (simp add: retract_of_space_def) + +lemma retract_of_space_topspace: + "topspace X retract_of_space X" + using retract_of_space_def by force + +lemma retract_of_space_empty [simp]: + "{} retract_of_space X \ topspace X = {}" + by (auto simp: continuous_map_def retract_of_space_def) + +lemma retract_of_space_singleton [simp]: + "{a} retract_of_space X \ a \ topspace X" +proof - + have "continuous_map X (subtopology X {a}) (\x. a) \ (\x. a) a = a" if "a \ topspace X" + using that by simp + then show ?thesis + by (force simp: retract_of_space_def) +qed + +lemma retract_of_space_clopen: + assumes "openin X S" "closedin X S" "S = {} \ topspace X = {}" + shows "S retract_of_space X" +proof (cases "S = {}") + case False + then obtain a where "a \ S" + by blast + show ?thesis + unfolding retract_of_space_def + proof (intro exI conjI) + show "S \ topspace X" + by (simp add: assms closedin_subset) + have "continuous_map X X (\x. if x \ S then x else a)" + proof (rule continuous_map_cases) + show "continuous_map (subtopology X (X closure_of {x. x \ S})) X (\x. x)" + by (simp add: continuous_map_from_subtopology) + show "continuous_map (subtopology X (X closure_of {x. x \ S})) X (\x. a)" + using \S \ topspace X\ \a \ S\ by force + show "x = a" if "x \ X frontier_of {x. x \ S}" for x + using assms that clopenin_eq_frontier_of by fastforce + qed + then show "continuous_map X (subtopology X S) (\x. if x \ S then x else a)" + using \S \ topspace X\ \a \ S\ by (auto simp: continuous_map_in_subtopology) + qed auto +qed (use assms in auto) + +lemma retract_of_space_disjoint_union: + assumes "openin X S" "openin X T" and ST: "disjnt S T" "S \ T = topspace X" and "S = {} \ topspace X = {}" + shows "S retract_of_space X" +proof (rule retract_of_space_clopen) + have "S \ T = {}" + by (meson ST disjnt_def) + then have "S = topspace X - T" + using ST by auto + then show "closedin X S" + using \openin X T\ by blast +qed (auto simp: assms) + +lemma retraction_maps_section_image1: + assumes "retraction_maps X Y r s" + shows "s ` (topspace Y) retract_of_space X" + unfolding retract_of_space_section_map +proof + show "s ` topspace Y \ topspace X" + using assms continuous_map_image_subset_topspace retraction_maps_def by blast + show "section_map (subtopology X (s ` topspace Y)) X id" + unfolding section_map_def + using assms retraction_maps_to_retract_maps by blast +qed + +lemma retraction_maps_section_image2: + "retraction_maps X Y r s + \ subtopology X (s ` (topspace Y)) homeomorphic_space Y" + using embedding_map_imp_homeomorphic_space homeomorphic_space_sym section_imp_embedding_map + section_map_def by blast + +subsection\Paths and path-connectedness\ + +definition pathin :: "'a topology \ (real \ 'a) \ bool" where + "pathin X g \ continuous_map (subtopology euclideanreal {0..1}) X g" + +lemma pathin_compose: + "\pathin X g; continuous_map X Y f\ \ pathin Y (f \ g)" + by (simp add: continuous_map_compose pathin_def) + +lemma pathin_subtopology: + "pathin (subtopology X S) g \ pathin X g \ (\x \ {0..1}. g x \ S)" + by (auto simp: pathin_def continuous_map_in_subtopology) + +lemma pathin_const: + "pathin X (\x. a) \ a \ topspace X" + by (simp add: pathin_def) + +definition path_connected_space :: "'a topology \ bool" + where "path_connected_space X \ \x \ topspace X. \ y \ topspace X. \g. pathin X g \ g 0 = x \ g 1 = y" + +definition path_connectedin :: "'a topology \ 'a set \ bool" + where "path_connectedin X S \ S \ topspace X \ path_connected_space(subtopology X S)" + +lemma path_connectedin_absolute [simp]: + "path_connectedin (subtopology X S) S \ path_connectedin X S" + by (simp add: path_connectedin_def subtopology_subtopology topspace_subtopology) + +lemma path_connectedin_subset_topspace: + "path_connectedin X S \ S \ topspace X" + by (simp add: path_connectedin_def) + +lemma path_connectedin_subtopology: + "path_connectedin (subtopology X S) T \ path_connectedin X T \ T \ S" + by (auto simp: path_connectedin_def subtopology_subtopology topspace_subtopology inf.absorb2) + +lemma path_connectedin: + "path_connectedin X S \ + S \ topspace X \ + (\x \ S. \y \ S. \g. pathin X g \ g ` {0..1} \ S \ g 0 = x \ g 1 = y)" + unfolding path_connectedin_def path_connected_space_def pathin_def continuous_map_in_subtopology + by (intro conj_cong refl ball_cong) (simp_all add: inf.absorb_iff2 topspace_subtopology) + +lemma path_connectedin_topspace: + "path_connectedin X (topspace X) \ path_connected_space X" + by (simp add: path_connectedin_def) + +lemma path_connected_imp_connected_space: + assumes "path_connected_space X" + shows "connected_space X" +proof - + have *: "\S. connectedin X S \ g 0 \ S \ g 1 \ S" if "pathin X g" for g + proof (intro exI conjI) + have "continuous_map (subtopology euclideanreal {0..1}) X g" + using connectedin_absolute that by (simp add: pathin_def) + then show "connectedin X (g ` {0..1})" + by (rule connectedin_continuous_map_image) auto + qed auto + show ?thesis + using assms + by (auto intro: * simp add: path_connected_space_def connected_space_subconnected Ball_def) +qed + +lemma path_connectedin_imp_connectedin: + "path_connectedin X S \ connectedin X S" + by (simp add: connectedin_def path_connected_imp_connected_space path_connectedin_def) + +lemma path_connected_space_topspace_empty: + "topspace X = {} \ path_connected_space X" + by (simp add: path_connected_space_def) + +lemma path_connectedin_empty [simp]: "path_connectedin X {}" + by (simp add: path_connectedin) + +lemma path_connectedin_singleton [simp]: "path_connectedin X {a} \ a \ topspace X" +proof + show "path_connectedin X {a} \ a \ topspace X" + by (simp add: path_connectedin) + show "a \ topspace X \ path_connectedin X {a}" + unfolding path_connectedin + using pathin_const by fastforce +qed + +lemma path_connectedin_continuous_map_image: + assumes f: "continuous_map X Y f" and S: "path_connectedin X S" + shows "path_connectedin Y (f ` S)" +proof - + have fX: "f ` (topspace X) \ topspace Y" + by (metis f continuous_map_image_subset_topspace) + show ?thesis + unfolding path_connectedin + proof (intro conjI ballI; clarify?) + fix x + assume "x \ S" + show "f x \ topspace Y" + by (meson S fX \x \ S\ image_subset_iff path_connectedin_subset_topspace set_mp) + next + fix x y + assume "x \ S" and "y \ S" + then obtain g where g: "pathin X g" "g ` {0..1} \ S" "g 0 = x" "g 1 = y" + using S by (force simp: path_connectedin) + show "\g. pathin Y g \ g ` {0..1} \ f ` S \ g 0 = f x \ g 1 = f y" + proof (intro exI conjI) + show "pathin Y (f \ g)" + using \pathin X g\ f pathin_compose by auto + qed (use g in auto) + qed +qed + +lemma homeomorphic_path_connected_space_imp: + "\path_connected_space X; X homeomorphic_space Y\ \ path_connected_space Y" + unfolding homeomorphic_space_def homeomorphic_maps_def + by (metis (no_types, hide_lams) continuous_map_closedin continuous_map_image_subset_topspace imageI order_class.order.antisym path_connectedin_continuous_map_image path_connectedin_topspace subsetI) + +lemma homeomorphic_path_connected_space: + "X homeomorphic_space Y \ path_connected_space X \ path_connected_space Y" + by (meson homeomorphic_path_connected_space_imp homeomorphic_space_sym) + +lemma homeomorphic_map_path_connectedness: + assumes "homeomorphic_map X Y f" "U \ topspace X" + shows "path_connectedin Y (f ` U) \ path_connectedin X U" + unfolding path_connectedin_def +proof (intro conj_cong homeomorphic_path_connected_space) + show "(f ` U \ topspace Y) = (U \ topspace X)" + using assms homeomorphic_imp_surjective_map by blast +next + assume "U \ topspace X" + show "subtopology Y (f ` U) homeomorphic_space subtopology X U" + using assms unfolding homeomorphic_eq_everything_map + by (metis (no_types, hide_lams) assms homeomorphic_map_subtopologies homeomorphic_space homeomorphic_space_sym image_mono inf.absorb_iff2) +qed + +lemma homeomorphic_map_path_connectedness_eq: + "homeomorphic_map X Y f \ path_connectedin X U \ U \ topspace X \ path_connectedin Y (f ` U)" + by (meson homeomorphic_map_path_connectedness path_connectedin_def) + +subsection\Connected components\ + +definition connected_component_of :: "'a topology \ 'a \ 'a \ bool" + where "connected_component_of X x y \ + \T. connectedin X T \ x \ T \ y \ T" + +abbreviation connected_component_of_set + where "connected_component_of_set X x \ Collect (connected_component_of X x)" + +definition connected_components_of :: "'a topology \ ('a set) set" + where "connected_components_of X \ connected_component_of_set X ` topspace X" + +lemma connected_component_in_topspace: + "connected_component_of X x y \ x \ topspace X \ y \ topspace X" + by (meson connected_component_of_def connectedin_subset_topspace in_mono) + +lemma connected_component_of_refl: + "connected_component_of X x x \ x \ topspace X" + by (meson connected_component_in_topspace connected_component_of_def connectedin_sing insertI1) + +lemma connected_component_of_sym: + "connected_component_of X x y \ connected_component_of X y x" + by (meson connected_component_of_def) + +lemma connected_component_of_trans: + "\connected_component_of X x y; connected_component_of X y z\ + \ connected_component_of X x z" + unfolding connected_component_of_def + using connectedin_Un by fastforce + +lemma connected_component_of_mono: + "\connected_component_of (subtopology X S) x y; S \ T\ + \ connected_component_of (subtopology X T) x y" + by (metis connected_component_of_def connectedin_subtopology inf.absorb_iff2 subtopology_subtopology) + +lemma connected_component_of_set: + "connected_component_of_set X x = {y. \T. connectedin X T \ x \ T \ y \ T}" + by (meson connected_component_of_def) + +lemma connected_component_of_subset_topspace: + "connected_component_of_set X x \ topspace X" + using connected_component_in_topspace by force + +lemma connected_component_of_eq_empty: + "connected_component_of_set X x = {} \ (x \ topspace X)" + using connected_component_in_topspace connected_component_of_refl by fastforce + +lemma connected_space_iff_connected_component: + "connected_space X \ (\x \ topspace X. \y \ topspace X. connected_component_of X x y)" + by (simp add: connected_component_of_def connected_space_subconnected) + +lemma connected_space_imp_connected_component_of: + "\connected_space X; a \ topspace X; b \ topspace X\ + \ connected_component_of X a b" + by (simp add: connected_space_iff_connected_component) + +lemma connected_space_connected_component_set: + "connected_space X \ (\x \ topspace X. connected_component_of_set X x = topspace X)" + using connected_component_of_subset_topspace connected_space_iff_connected_component by fastforce + +lemma connected_component_of_maximal: + "\connectedin X S; x \ S\ \ S \ connected_component_of_set X x" + by (meson Ball_Collect connected_component_of_def) + +lemma connected_component_of_equiv: + "connected_component_of X x y \ + x \ topspace X \ y \ topspace X \ connected_component_of X x = connected_component_of X y" + apply (simp add: connected_component_in_topspace fun_eq_iff) + by (meson connected_component_of_refl connected_component_of_sym connected_component_of_trans) + +lemma connected_component_of_disjoint: + "disjnt (connected_component_of_set X x) (connected_component_of_set X y) + \ ~(connected_component_of X x y)" + using connected_component_of_equiv unfolding disjnt_iff by force + +lemma connected_component_of_eq: + "connected_component_of X x = connected_component_of X y \ + (x \ topspace X) \ (y \ topspace X) \ + x \ topspace X \ y \ topspace X \ + connected_component_of X x y" + by (metis Collect_empty_eq_bot connected_component_of_eq_empty connected_component_of_equiv) + +lemma connectedin_connected_component_of: + "connectedin X (connected_component_of_set X x)" +proof - + have "connected_component_of_set X x = \ {T. connectedin X T \ x \ T}" + by (auto simp: connected_component_of_def) + then show ?thesis + apply (rule ssubst) + by (blast intro: connectedin_Union) +qed + + +lemma Union_connected_components_of: + "\(connected_components_of X) = topspace X" + unfolding connected_components_of_def + apply (rule equalityI) + apply (simp add: SUP_least connected_component_of_subset_topspace) + using connected_component_of_refl by fastforce + +lemma connected_components_of_maximal: + "\C \ connected_components_of X; connectedin X S; ~disjnt C S\ \ S \ C" + unfolding connected_components_of_def disjnt_def + apply clarify + by (metis Int_emptyI connected_component_of_def connected_component_of_trans mem_Collect_eq) + +lemma pairwise_disjoint_connected_components_of: + "pairwise disjnt (connected_components_of X)" + unfolding connected_components_of_def pairwise_def + apply clarify + by (metis connected_component_of_disjoint connected_component_of_equiv) + +lemma complement_connected_components_of_Union: + "C \ connected_components_of X + \ topspace X - C = \ (connected_components_of X - {C})" + apply (rule equalityI) + using Union_connected_components_of apply fastforce + by (metis Diff_cancel Diff_subset Union_connected_components_of cSup_singleton diff_Union_pairwise_disjoint equalityE insert_subsetI pairwise_disjoint_connected_components_of) + +lemma nonempty_connected_components_of: + "C \ connected_components_of X \ C \ {}" + unfolding connected_components_of_def + by (metis (no_types, lifting) connected_component_of_eq_empty imageE) + +lemma connected_components_of_subset: + "C \ connected_components_of X \ C \ topspace X" + using Union_connected_components_of by fastforce + +lemma connectedin_connected_components_of: + assumes "C \ connected_components_of X" + shows "connectedin X C" +proof - + have "C \ connected_component_of_set X ` topspace X" + using assms connected_components_of_def by blast +then show ?thesis + using connectedin_connected_component_of by fastforce +qed + +lemma connected_component_in_connected_components_of: + "connected_component_of_set X a \ connected_components_of X \ a \ topspace X" + apply (rule iffI) + using connected_component_of_eq_empty nonempty_connected_components_of apply fastforce + by (simp add: connected_components_of_def) + +lemma connected_space_iff_components_eq: + "connected_space X \ (\C \ connected_components_of X. \C' \ connected_components_of X. C = C')" + apply (rule iffI) + apply (force simp: connected_components_of_def connected_space_connected_component_set image_iff) + by (metis connected_component_in_connected_components_of connected_component_of_refl connected_space_iff_connected_component mem_Collect_eq) + +lemma connected_components_of_eq_empty: + "connected_components_of X = {} \ topspace X = {}" + by (simp add: connected_components_of_def) + +lemma connected_components_of_empty_space: + "topspace X = {} \ connected_components_of X = {}" + by (simp add: connected_components_of_eq_empty) + +lemma connected_components_of_subset_sing: + "connected_components_of X \ {S} \ connected_space X \ (topspace X = {} \ topspace X = S)" +proof (cases "topspace X = {}") + case True + then show ?thesis + by (simp add: connected_components_of_empty_space connected_space_topspace_empty) +next + case False + then show ?thesis + by (metis (no_types, hide_lams) Union_connected_components_of ccpo_Sup_singleton + connected_components_of_eq_empty connected_space_iff_components_eq insertI1 singletonD + subsetI subset_singleton_iff) +qed + +lemma connected_space_iff_components_subset_singleton: + "connected_space X \ (\a. connected_components_of X \ {a})" + by (simp add: connected_components_of_subset_sing) + +lemma connected_components_of_eq_singleton: + "connected_components_of X = {S} +\ connected_space X \ topspace X \ {} \ S = topspace X" + by (metis ccpo_Sup_singleton connected_components_of_subset_sing insert_not_empty subset_singleton_iff) + +lemma connected_components_of_connected_space: + "connected_space X \ connected_components_of X = (if topspace X = {} then {} else {topspace X})" + by (simp add: connected_components_of_eq_empty connected_components_of_eq_singleton) + +lemma exists_connected_component_of_superset: + assumes "connectedin X S" and ne: "topspace X \ {}" + shows "\C. C \ connected_components_of X \ S \ C" +proof (cases "S = {}") + case True + then show ?thesis + using ne connected_components_of_def by blast +next + case False + then show ?thesis + by (meson all_not_in_conv assms(1) connected_component_in_connected_components_of connected_component_of_maximal connectedin_subset_topspace in_mono) +qed + +lemma closedin_connected_components_of: + assumes "C \ connected_components_of X" + shows "closedin X C" +proof - + obtain x where "x \ topspace X" and x: "C = connected_component_of_set X x" + using assms by (auto simp: connected_components_of_def) + have "connected_component_of_set X x \ topspace X" + by (simp add: connected_component_of_subset_topspace) + moreover have "X closure_of connected_component_of_set X x \ connected_component_of_set X x" + proof (rule connected_component_of_maximal) + show "connectedin X (X closure_of connected_component_of_set X x)" + by (simp add: connectedin_closure_of connectedin_connected_component_of) + show "x \ X closure_of connected_component_of_set X x" + by (simp add: \x \ topspace X\ closure_of connected_component_of_refl) + qed + ultimately + show ?thesis + using closure_of_subset_eq x by auto +qed + +lemma closedin_connected_component_of: + "closedin X (connected_component_of_set X x)" + by (metis closedin_connected_components_of closedin_empty connected_component_in_connected_components_of connected_component_of_eq_empty) + +lemma connected_component_of_eq_overlap: + "connected_component_of_set X x = connected_component_of_set X y \ + (x \ topspace X) \ (y \ topspace X) \ + ~(connected_component_of_set X x \ connected_component_of_set X y = {})" + using connected_component_of_equiv by fastforce + +lemma connected_component_of_nonoverlap: + "connected_component_of_set X x \ connected_component_of_set X y = {} \ + (x \ topspace X) \ (y \ topspace X) \ + ~(connected_component_of_set X x = connected_component_of_set X y)" + by (metis connected_component_of_eq_empty connected_component_of_eq_overlap inf.idem) + +lemma connected_component_of_overlap: + "~(connected_component_of_set X x \ connected_component_of_set X y = {}) \ + x \ topspace X \ y \ topspace X \ + connected_component_of_set X x = connected_component_of_set X y" + by (meson connected_component_of_nonoverlap) + + end \ No newline at end of file diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Analysis.thy --- a/src/HOL/Analysis/Analysis.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Analysis.thy Tue Mar 19 16:14:51 2019 +0000 @@ -27,7 +27,7 @@ Homeomorphism Bounded_Continuous_Function Abstract_Topology - Function_Topology + Product_Topology T1_Spaces Infinite_Products Infinite_Set_Sum diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Arcwise_Connected.thy --- a/src/HOL/Analysis/Arcwise_Connected.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Arcwise_Connected.thy Tue Mar 19 16:14:51 2019 +0000 @@ -2005,7 +2005,7 @@ lemma dense_accessible_frontier_points: fixes S :: "'a::{complete_space,real_normed_vector} set" - assumes "open S" and opeSV: "openin (subtopology euclidean (frontier S)) V" and "V \ {}" + assumes "open S" and opeSV: "openin (top_of_set (frontier S)) V" and "V \ {}" obtains g where "arc g" "g ` {0..<1} \ S" "pathstart g \ S" "pathfinish g \ V" proof - obtain z where "z \ V" @@ -2101,7 +2101,7 @@ lemma dense_accessible_frontier_points_connected: fixes S :: "'a::{complete_space,real_normed_vector} set" assumes "open S" "connected S" "x \ S" "V \ {}" - and ope: "openin (subtopology euclidean (frontier S)) V" + and ope: "openin (top_of_set (frontier S)) V" obtains g where "arc g" "g ` {0..<1} \ S" "pathstart g = x" "pathfinish g \ V" proof - have "V \ frontier S" @@ -2136,8 +2136,8 @@ lemma dense_access_fp_aux: fixes S :: "'a::{complete_space,real_normed_vector} set" assumes S: "open S" "connected S" - and opeSU: "openin (subtopology euclidean (frontier S)) U" - and opeSV: "openin (subtopology euclidean (frontier S)) V" + and opeSU: "openin (top_of_set (frontier S)) U" + and opeSV: "openin (top_of_set (frontier S)) V" and "V \ {}" "\ U \ V" obtains g where "arc g" "pathstart g \ U" "pathfinish g \ V" "g ` {0<..<1} \ S" proof - @@ -2150,7 +2150,7 @@ proof (rule dense_accessible_frontier_points_connected [OF S \x \ S\]) show "U - {pathfinish g} \ {}" using \pathfinish g \ V\ \\ U \ V\ by blast - show "openin (subtopology euclidean (frontier S)) (U - {pathfinish g})" + show "openin (top_of_set (frontier S)) (U - {pathfinish g})" by (simp add: opeSU openin_delete) qed auto obtain \ where "arc \" @@ -2183,8 +2183,8 @@ lemma dense_accessible_frontier_point_pairs: fixes S :: "'a::{complete_space,real_normed_vector} set" assumes S: "open S" "connected S" - and opeSU: "openin (subtopology euclidean (frontier S)) U" - and opeSV: "openin (subtopology euclidean (frontier S)) V" + and opeSU: "openin (top_of_set (frontier S)) U" + and opeSV: "openin (top_of_set (frontier S)) V" and "U \ {}" "V \ {}" "U \ V" obtains g where "arc g" "pathstart g \ U" "pathfinish g \ V" "g ` {0<..<1} \ S" proof - diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Brouwer_Fixpoint.thy --- a/src/HOL/Analysis/Brouwer_Fixpoint.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Brouwer_Fixpoint.thy Tue Mar 19 16:14:51 2019 +0000 @@ -85,7 +85,7 @@ qed lemma retraction_imp_quotient_map: - "openin (subtopology euclidean S) (S \ r -` U) \ openin (subtopology euclidean T) U" + "openin (top_of_set S) (S \ r -` U) \ openin (top_of_set T) U" if retraction: "retraction S T r" and "U \ T" using retraction apply (rule retractionE) apply (rule continuous_right_inverse_imp_quotient_map [where g=r]) @@ -233,12 +233,12 @@ definition%important AR :: "'a::topological_space set \ bool" where "AR S \ \U. \S'::('a * real) set. - S homeomorphic S' \ closedin (subtopology euclidean U) S' \ S' retract_of U" + S homeomorphic S' \ closedin (top_of_set U) S' \ S' retract_of U" definition%important ANR :: "'a::topological_space set \ bool" where "ANR S \ \U. \S'::('a * real) set. - S homeomorphic S' \ closedin (subtopology euclidean U) S' - \ (\T. openin (subtopology euclidean U) T \ S' retract_of T)" + S homeomorphic S' \ closedin (top_of_set U) S' + \ (\T. openin (top_of_set U) T \ S' retract_of T)" definition%important ENR :: "'a::topological_space set \ bool" where "ENR S \ \U. open U \ S retract_of U" @@ -248,14 +248,14 @@ lemma AR_imp_absolute_extensor: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "AR S" and contf: "continuous_on T f" and "f ` T \ S" - and cloUT: "closedin (subtopology euclidean U) T" + and cloUT: "closedin (top_of_set U) T" obtains g where "continuous_on U g" "g ` U \ S" "\x. x \ T \ g x = f x" proof - have "aff_dim S < int (DIM('b \ real))" using aff_dim_le_DIM [of S] by simp then obtain C and S' :: "('b * real) set" where C: "convex C" "C \ {}" - and cloCS: "closedin (subtopology euclidean C) S'" + and cloCS: "closedin (top_of_set C) S'" and hom: "S homeomorphic S'" by (metis that homeomorphic_closedin_convex) then have "S' retract_of C" @@ -296,7 +296,7 @@ lemma AR_imp_absolute_retract: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" assumes "AR S" "S homeomorphic S'" - and clo: "closedin (subtopology euclidean U) S'" + and clo: "closedin (top_of_set U) S'" shows "S' retract_of U" proof - obtain g h where hom: "homeomorphism S S' g h" @@ -334,12 +334,12 @@ fixes S :: "'a::euclidean_space set" assumes "\f :: 'a * real \ 'a. \U T. \continuous_on T f; f ` T \ S; - closedin (subtopology euclidean U) T\ + closedin (top_of_set U) T\ \ \g. continuous_on U g \ g ` U \ S \ (\x \ T. g x = f x)" shows "AR S" proof (clarsimp simp: AR_def) fix U and T :: "('a * real) set" - assume "S homeomorphic T" and clo: "closedin (subtopology euclidean U) T" + assume "S homeomorphic T" and clo: "closedin (top_of_set U) T" then obtain g h where hom: "homeomorphism S T g h" by (force simp: homeomorphic_def) have h: "continuous_on T h" " h ` T \ S" @@ -369,7 +369,7 @@ shows "AR S \ (\f :: 'a * real \ 'a. \U T. continuous_on T f \ f ` T \ S \ - closedin (subtopology euclidean U) T \ + closedin (top_of_set U) T \ (\g. continuous_on U g \ g ` U \ S \ (\x \ T. g x = f x)))" apply (rule iffI) apply (metis AR_imp_absolute_extensor) @@ -378,7 +378,7 @@ lemma AR_imp_retract: fixes S :: "'a::euclidean_space set" - assumes "AR S \ closedin (subtopology euclidean U) S" + assumes "AR S \ closedin (top_of_set U) S" shows "S retract_of U" using AR_imp_absolute_retract assms homeomorphic_refl by blast @@ -398,8 +398,8 @@ lemma ANR_imp_absolute_neighbourhood_extensor: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "ANR S" and contf: "continuous_on T f" and "f ` T \ S" - and cloUT: "closedin (subtopology euclidean U) T" - obtains V g where "T \ V" "openin (subtopology euclidean U) V" + and cloUT: "closedin (top_of_set U) T" + obtains V g where "T \ V" "openin (top_of_set U) V" "continuous_on V g" "g ` V \ S" "\x. x \ T \ g x = f x" proof - @@ -407,10 +407,10 @@ using aff_dim_le_DIM [of S] by simp then obtain C and S' :: "('b * real) set" where C: "convex C" "C \ {}" - and cloCS: "closedin (subtopology euclidean C) S'" + and cloCS: "closedin (top_of_set C) S'" and hom: "S homeomorphic S'" by (metis that homeomorphic_closedin_convex) - then obtain D where opD: "openin (subtopology euclidean C) D" and "S' retract_of D" + then obtain D where opD: "openin (top_of_set C) D" and "S' retract_of D" using \ANR S\ by (auto simp: ANR_def) then obtain r where "S' \ D" and contr: "continuous_on D r" and "r ` D \ S'" and rid: "\x. x \ S' \ r x = x" @@ -437,7 +437,7 @@ show "T \ U \ f' -` D" using cloUT closedin_imp_subset \S' \ D\ \f ` T \ S\ eq homeomorphism_image1 homgh by fastforce - show ope: "openin (subtopology euclidean U) (U \ f' -` D)" + show ope: "openin (top_of_set U) (U \ f' -` D)" using \f' ` U \ C\ by (auto simp: opD contf' continuous_openin_preimage) have conth: "continuous_on (r ` f' ` (U \ f' -` D)) h" apply (rule continuous_on_subset [of S']) @@ -460,8 +460,8 @@ corollary ANR_imp_absolute_neighbourhood_retract: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" assumes "ANR S" "S homeomorphic S'" - and clo: "closedin (subtopology euclidean U) S'" - obtains V where "openin (subtopology euclidean U) V" "S' retract_of V" + and clo: "closedin (top_of_set U) S'" + obtains V where "openin (top_of_set U) V" "S' retract_of V" proof - obtain g h where hom: "homeomorphism S S' g h" using assms by (force simp: homeomorphic_def) @@ -470,7 +470,7 @@ apply (metis hom equalityE homeomorphism_def) done from ANR_imp_absolute_neighbourhood_extensor [OF \ANR S\ h clo] - obtain V h' where "S' \ V" and opUV: "openin (subtopology euclidean U) V" + obtain V h' where "S' \ V" and opUV: "openin (top_of_set U) V" and h': "continuous_on V h'" "h' ` V \ S" and h'h:"\x. x \ S' \ h' x = h x" by (blast intro: ANR_imp_absolute_neighbourhood_extensor [OF \ANR S\ h clo]) @@ -508,20 +508,20 @@ fixes S :: "'a::euclidean_space set" assumes "\f :: 'a * real \ 'a. \U T. \continuous_on T f; f ` T \ S; - closedin (subtopology euclidean U) T\ - \ \V g. T \ V \ openin (subtopology euclidean U) V \ + closedin (top_of_set U) T\ + \ \V g. T \ V \ openin (top_of_set U) V \ continuous_on V g \ g ` V \ S \ (\x \ T. g x = f x)" shows "ANR S" proof (clarsimp simp: ANR_def) fix U and T :: "('a * real) set" - assume "S homeomorphic T" and clo: "closedin (subtopology euclidean U) T" + assume "S homeomorphic T" and clo: "closedin (top_of_set U) T" then obtain g h where hom: "homeomorphism S T g h" by (force simp: homeomorphic_def) have h: "continuous_on T h" " h ` T \ S" using hom homeomorphism_def apply blast apply (metis hom equalityE homeomorphism_def) done - obtain V h' where "T \ V" and opV: "openin (subtopology euclidean U) V" + obtain V h' where "T \ V" and opV: "openin (top_of_set U) V" and h': "continuous_on V h'" "h' ` V \ S" and h'h: "\x\T. h' x = h x" using assms [OF h clo] by blast @@ -538,7 +538,7 @@ show "\x\T. (g \ h') x = x" by clarsimp (metis h'h hom homeomorphism_def) qed - then show "\V. openin (subtopology euclidean U) V \ T retract_of V" + then show "\V. openin (top_of_set U) V \ T retract_of V" using opV by blast qed @@ -547,8 +547,8 @@ shows "ANR S \ (\f :: 'a * real \ 'a. \U T. continuous_on T f \ f ` T \ S \ - closedin (subtopology euclidean U) T \ - (\V g. T \ V \ openin (subtopology euclidean U) V \ + closedin (top_of_set U) T \ + (\V g. T \ V \ openin (top_of_set U) V \ continuous_on V g \ g ` V \ S \ (\x \ T. g x = f x)))" apply (rule iffI) apply (metis ANR_imp_absolute_neighbourhood_extensor) @@ -557,27 +557,27 @@ lemma ANR_imp_neighbourhood_retract: fixes S :: "'a::euclidean_space set" - assumes "ANR S" "closedin (subtopology euclidean U) S" - obtains V where "openin (subtopology euclidean U) V" "S retract_of V" + assumes "ANR S" "closedin (top_of_set U) S" + obtains V where "openin (top_of_set U) V" "S retract_of V" using ANR_imp_absolute_neighbourhood_retract assms homeomorphic_refl by blast lemma ANR_imp_absolute_closed_neighbourhood_retract: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" - assumes "ANR S" "S homeomorphic S'" and US': "closedin (subtopology euclidean U) S'" + assumes "ANR S" "S homeomorphic S'" and US': "closedin (top_of_set U) S'" obtains V W - where "openin (subtopology euclidean U) V" - "closedin (subtopology euclidean U) W" + where "openin (top_of_set U) V" + "closedin (top_of_set U) W" "S' \ V" "V \ W" "S' retract_of W" proof - - obtain Z where "openin (subtopology euclidean U) Z" and S'Z: "S' retract_of Z" + obtain Z where "openin (top_of_set U) Z" and S'Z: "S' retract_of Z" by (blast intro: assms ANR_imp_absolute_neighbourhood_retract) - then have UUZ: "closedin (subtopology euclidean U) (U - Z)" + then have UUZ: "closedin (top_of_set U) (U - Z)" by auto have "S' \ (U - Z) = {}" using \S' retract_of Z\ closedin_retract closedin_subtopology by fastforce then obtain V W - where "openin (subtopology euclidean U) V" - and "openin (subtopology euclidean U) W" + where "openin (top_of_set U) V" + and "openin (top_of_set U) W" and "S' \ V" "U - Z \ W" "V \ W = {}" using separation_normal_local [OF US' UUZ] by auto moreover have "S' retract_of U - W" @@ -592,9 +592,9 @@ lemma ANR_imp_closed_neighbourhood_retract: fixes S :: "'a::euclidean_space set" - assumes "ANR S" "closedin (subtopology euclidean U) S" - obtains V W where "openin (subtopology euclidean U) V" - "closedin (subtopology euclidean U) W" + assumes "ANR S" "closedin (top_of_set U) S" + obtains V W where "openin (top_of_set U) V" + "closedin (top_of_set U) W" "S \ V" "V \ W" "S retract_of W" by (meson ANR_imp_absolute_closed_neighbourhood_retract assms homeomorphic_refl) @@ -616,7 +616,7 @@ fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" assumes "ENR S" and hom: "S homeomorphic S'" and "S' \ U" - obtains V where "openin (subtopology euclidean U) V" "S' retract_of V" + obtains V where "openin (top_of_set U) V" "S' retract_of V" proof - obtain X where "open X" "S retract_of X" using \ENR S\ by (auto simp: ENR_def) @@ -625,8 +625,8 @@ have "locally compact S'" using retract_of_locally_compact open_imp_locally_compact homeomorphic_local_compactness \S retract_of X\ \open X\ hom by blast - then obtain W where UW: "openin (subtopology euclidean U) W" - and WS': "closedin (subtopology euclidean W) S'" + then obtain W where UW: "openin (top_of_set U) W" + and WS': "closedin (top_of_set W) S'" apply (rule locally_compact_closedin_open) apply (rename_tac W) apply (rule_tac W = "U \ W" in that, blast) @@ -752,7 +752,7 @@ proof assume ?lhs obtain C and S' :: "('a * real) set" - where "convex C" "C \ {}" "closedin (subtopology euclidean C) S'" "S homeomorphic S'" + where "convex C" "C \ {}" "closedin (top_of_set C) S'" "S homeomorphic S'" apply (rule homeomorphic_closedin_convex [of S, where 'n = "'a * real"]) using aff_dim_le_DIM [of S] by auto with \AR S\ have "contractible S" @@ -777,27 +777,27 @@ by (metis all_not_in_conv atLeastAtMost_iff image_subset_iff mem_Sigma_iff order_refl zero_le_one) have "\g. continuous_on W g \ g ` W \ S \ (\x\T. g x = f x)" if f: "continuous_on T f" "f ` T \ S" - and WT: "closedin (subtopology euclidean W) T" + and WT: "closedin (top_of_set W) T" for W T and f :: "'a \ real \ 'a" proof - obtain U g - where "T \ U" and WU: "openin (subtopology euclidean W) U" + where "T \ U" and WU: "openin (top_of_set W) U" and contg: "continuous_on U g" and "g ` U \ S" and gf: "\x. x \ T \ g x = f x" using iffD1 [OF ANR_eq_absolute_neighbourhood_extensor \ANR S\, rule_format, OF f WT] by auto - have WWU: "closedin (subtopology euclidean W) (W - U)" + have WWU: "closedin (top_of_set W) (W - U)" using WU closedin_diff by fastforce moreover have "(W - U) \ T = {}" using \T \ U\ by auto ultimately obtain V V' - where WV': "openin (subtopology euclidean W) V'" - and WV: "openin (subtopology euclidean W) V" + where WV': "openin (top_of_set W) V'" + and WV: "openin (top_of_set W) V" and "W - U \ V'" "T \ V" "V' \ V = {}" using separation_normal_local [of W "W-U" T] WT by blast then have WVT: "T \ (W - V) = {}" by auto - have WWV: "closedin (subtopology euclidean W) (W - V)" + have WWV: "closedin (top_of_set W) (W - V)" using WV closedin_diff by fastforce obtain j :: " 'a \ real \ real" where contj: "continuous_on W j" @@ -931,8 +931,8 @@ lemma AR_closed_Un_local_aux: fixes U :: "'a::euclidean_space set" - assumes "closedin (subtopology euclidean U) S" - "closedin (subtopology euclidean U) T" + assumes "closedin (top_of_set U) S" + "closedin (top_of_set U) T" "AR S" "AR T" "AR(S \ T)" shows "(S \ T) retract_of U" proof - @@ -943,10 +943,10 @@ define S' where "S' \ {x \ U. setdist {x} S \ setdist {x} T}" define T' where "T' \ {x \ U. setdist {x} T \ setdist {x} S}" define W where "W \ {x \ U. setdist {x} S = setdist {x} T}" - have US': "closedin (subtopology euclidean U) S'" + have US': "closedin (top_of_set U) S'" using continuous_closedin_preimage [of U "\x. setdist {x} S - setdist {x} T" "{..0}"] by (simp add: S'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) - have UT': "closedin (subtopology euclidean U) T'" + have UT': "closedin (top_of_set U) T'" using continuous_closedin_preimage [of U "\x. setdist {x} T - setdist {x} S" "{..0}"] by (simp add: T'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) have "S \ S'" @@ -971,7 +971,7 @@ by (force simp: W_def setdist_sing_in_set) have "S' \ T' = W" by (auto simp: S'_def T'_def W_def) - then have cloUW: "closedin (subtopology euclidean U) W" + then have cloUW: "closedin (top_of_set U) W" using closedin_Int US' UT' by blast define r where "r \ \x. if x \ W then r0 x else x" have "r ` (W \ S) \ S" "r ` (W \ T) \ T" @@ -979,14 +979,14 @@ have contr: "continuous_on (W \ (S \ T)) r" unfolding r_def proof (rule continuous_on_cases_local [OF _ _ contr0 continuous_on_id]) - show "closedin (subtopology euclidean (W \ (S \ T))) W" - using \S \ U\ \T \ U\ \W \ U\ \closedin (subtopology euclidean U) W\ closedin_subset_trans by fastforce - show "closedin (subtopology euclidean (W \ (S \ T))) (S \ T)" + show "closedin (top_of_set (W \ (S \ T))) W" + using \S \ U\ \T \ U\ \W \ U\ \closedin (top_of_set U) W\ closedin_subset_trans by fastforce + show "closedin (top_of_set (W \ (S \ T))) (S \ T)" by (meson \S \ U\ \T \ U\ \W \ U\ assms closedin_Un closedin_subset_trans sup.bounded_iff sup.cobounded2) show "\x. x \ W \ x \ W \ x \ S \ T \ x \ W \ r0 x = x" by (auto simp: ST) qed - have cloUWS: "closedin (subtopology euclidean U) (W \ S)" + have cloUWS: "closedin (top_of_set U) (W \ S)" by (simp add: cloUW assms closedin_Un) obtain g where contg: "continuous_on U g" and "g ` U \ S" and geqr: "\x. x \ W \ S \ g x = r x" @@ -994,7 +994,7 @@ apply (rule continuous_on_subset [OF contr]) using \r ` (W \ S) \ S\ apply auto done - have cloUWT: "closedin (subtopology euclidean U) (W \ T)" + have cloUWT: "closedin (top_of_set U) (W \ T)" by (simp add: cloUW assms closedin_Un) obtain h where conth: "continuous_on U h" and "h ` U \ T" and heqr: "\x. x \ W \ T \ h x = r x" @@ -1022,24 +1022,24 @@ lemma AR_closed_Un_local: fixes S :: "'a::euclidean_space set" - assumes STS: "closedin (subtopology euclidean (S \ T)) S" - and STT: "closedin (subtopology euclidean (S \ T)) T" + assumes STS: "closedin (top_of_set (S \ T)) S" + and STT: "closedin (top_of_set (S \ T)) T" and "AR S" "AR T" "AR(S \ T)" shows "AR(S \ T)" proof - have "C retract_of U" - if hom: "S \ T homeomorphic C" and UC: "closedin (subtopology euclidean U) C" + if hom: "S \ T homeomorphic C" and UC: "closedin (top_of_set U) C" for U and C :: "('a * real) set" proof - obtain f g where hom: "homeomorphism (S \ T) C f g" using hom by (force simp: homeomorphic_def) - have US: "closedin (subtopology euclidean U) (C \ g -` S)" + have US: "closedin (top_of_set U) (C \ g -` S)" apply (rule closedin_trans [OF _ UC]) apply (rule continuous_closedin_preimage_gen [OF _ _ STS]) using hom homeomorphism_def apply blast apply (metis hom homeomorphism_def set_eq_subset) done - have UT: "closedin (subtopology euclidean U) (C \ g -` T)" + have UT: "closedin (top_of_set U) (C \ g -` T)" apply (rule closedin_trans [OF _ UC]) apply (rule continuous_closedin_preimage_gen [OF _ _ STT]) using hom homeomorphism_def apply blast @@ -1088,10 +1088,10 @@ lemma ANR_closed_Un_local_aux: fixes U :: "'a::euclidean_space set" - assumes US: "closedin (subtopology euclidean U) S" - and UT: "closedin (subtopology euclidean U) T" + assumes US: "closedin (top_of_set U) S" + and UT: "closedin (top_of_set U) T" and "ANR S" "ANR T" "ANR(S \ T)" - obtains V where "openin (subtopology euclidean U) V" "(S \ T) retract_of V" + obtains V where "openin (top_of_set U) V" "(S \ T) retract_of V" proof (cases "S = {} \ T = {}") case True with assms that show ?thesis by (metis ANR_imp_neighbourhood_retract Un_commute inf_bot_right sup_inf_absorb) @@ -1103,10 +1103,10 @@ define S' where "S' \ {x \ U. setdist {x} S \ setdist {x} T}" define T' where "T' \ {x \ U. setdist {x} T \ setdist {x} S}" define W where "W \ {x \ U. setdist {x} S = setdist {x} T}" - have cloUS': "closedin (subtopology euclidean U) S'" + have cloUS': "closedin (top_of_set U) S'" using continuous_closedin_preimage [of U "\x. setdist {x} S - setdist {x} T" "{..0}"] by (simp add: S'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) - have cloUT': "closedin (subtopology euclidean U) T'" + have cloUT': "closedin (top_of_set U) T'" using continuous_closedin_preimage [of U "\x. setdist {x} T - setdist {x} S" "{..0}"] by (simp add: T'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) have "S \ S'" @@ -1123,17 +1123,17 @@ using \S \ U\ by (force simp: W_def setdist_sing_in_set)+ have "S' \ T' = W" by (auto simp: S'_def T'_def W_def) - then have cloUW: "closedin (subtopology euclidean U) W" + then have cloUW: "closedin (top_of_set U) W" using closedin_Int cloUS' cloUT' by blast - obtain W' W0 where "openin (subtopology euclidean W) W'" - and cloWW0: "closedin (subtopology euclidean W) W0" + obtain W' W0 where "openin (top_of_set W) W'" + and cloWW0: "closedin (top_of_set W) W0" and "S \ T \ W'" "W' \ W0" and ret: "(S \ T) retract_of W0" apply (rule ANR_imp_closed_neighbourhood_retract [OF \ANR(S \ T)\]) apply (rule closedin_subset_trans [of U, OF _ ST_W \W \ U\]) apply (blast intro: assms)+ done - then obtain U0 where opeUU0: "openin (subtopology euclidean U) U0" + then obtain U0 where opeUU0: "openin (top_of_set U) U0" and U0: "S \ T \ U0" "U0 \ W \ W0" unfolding openin_open using \W \ U\ by blast have "W0 \ U" @@ -1150,29 +1150,29 @@ have contr: "continuous_on (W0 \ (S \ T)) r" unfolding r_def proof (rule continuous_on_cases_local [OF _ _ contr0 continuous_on_id]) - show "closedin (subtopology euclidean (W0 \ (S \ T))) W0" + show "closedin (top_of_set (W0 \ (S \ T))) W0" apply (rule closedin_subset_trans [of U]) using cloWW0 cloUW closedin_trans \W0 \ U\ \S \ U\ \T \ U\ apply blast+ done - show "closedin (subtopology euclidean (W0 \ (S \ T))) (S \ T)" + show "closedin (top_of_set (W0 \ (S \ T))) (S \ T)" by (meson \S \ U\ \T \ U\ \W0 \ U\ assms closedin_Un closedin_subset_trans sup.bounded_iff sup.cobounded2) show "\x. x \ W0 \ x \ W0 \ x \ S \ T \ x \ W0 \ r0 x = x" using ST cloWW0 closedin_subset by fastforce qed - have cloS'WS: "closedin (subtopology euclidean S') (W0 \ S)" + have cloS'WS: "closedin (top_of_set S') (W0 \ S)" by (meson closedin_subset_trans US cloUS' \S \ S'\ \W \ S'\ cloUW cloWW0 closedin_Un closedin_imp_subset closedin_trans) obtain W1 g where "W0 \ S \ W1" and contg: "continuous_on W1 g" - and opeSW1: "openin (subtopology euclidean S') W1" + and opeSW1: "openin (top_of_set S') W1" and "g ` W1 \ S" and geqr: "\x. x \ W0 \ S \ g x = r x" apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \ANR S\ _ \r ` (W0 \ S) \ S\ cloS'WS]) apply (rule continuous_on_subset [OF contr], blast+) done - have cloT'WT: "closedin (subtopology euclidean T') (W0 \ T)" + have cloT'WT: "closedin (top_of_set T') (W0 \ T)" by (meson closedin_subset_trans UT cloUT' \T \ T'\ \W \ T'\ cloUW cloWW0 closedin_Un closedin_imp_subset closedin_trans) obtain W2 h where "W0 \ T \ W2" and conth: "continuous_on W2 h" - and opeSW2: "openin (subtopology euclidean T') W2" + and opeSW2: "openin (top_of_set T') W2" and "h ` W2 \ T" and heqr: "\x. x \ W0 \ T \ h x = r x" apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \ANR T\ _ \r ` (W0 \ T) \ T\ cloT'WT]) apply (rule continuous_on_subset [OF contr], blast+) @@ -1187,17 +1187,17 @@ ((U - T') \ O1 \ (U - S') \ O2 \ U \ O1 \ O2) - (W - U0)" using \U0 \ W \ W0\ \W0 \ S \ W1\ \W0 \ T \ W2\ by (auto simp: \S' \ T' = U\ [symmetric] \S' \ T' = W\ [symmetric] \W1 = S' \ O1\ \W2 = T' \ O2\) - show "openin (subtopology euclidean U) (W1 - (W - U0) \ (W2 - (W - U0)))" + show "openin (top_of_set U) (W1 - (W - U0) \ (W2 - (W - U0)))" apply (subst eq) apply (intro openin_Un openin_Int_open openin_diff closedin_diff cloUW opeUU0 cloUS' cloUT' \open O1\ \open O2\, simp_all) done - have cloW1: "closedin (subtopology euclidean (W1 - (W - U0) \ (W2 - (W - U0)))) (W1 - (W - U0))" + have cloW1: "closedin (top_of_set (W1 - (W - U0) \ (W2 - (W - U0)))) (W1 - (W - U0))" using cloUS' apply (simp add: closedin_closed) apply (erule ex_forward) using U0 \W0 \ S \ W1\ apply (auto simp: \W1 = S' \ O1\ \W2 = T' \ O2\ \S' \ T' = U\ [symmetric]\S' \ T' = W\ [symmetric]) done - have cloW2: "closedin (subtopology euclidean (W1 - (W - U0) \ (W2 - (W - U0)))) (W2 - (W - U0))" + have cloW2: "closedin (top_of_set (W1 - (W - U0) \ (W2 - (W - U0)))) (W2 - (W - U0))" using cloUT' apply (simp add: closedin_closed) apply (erule ex_forward) using U0 \W0 \ T \ W2\ @@ -1227,24 +1227,24 @@ lemma ANR_closed_Un_local: fixes S :: "'a::euclidean_space set" - assumes STS: "closedin (subtopology euclidean (S \ T)) S" - and STT: "closedin (subtopology euclidean (S \ T)) T" + assumes STS: "closedin (top_of_set (S \ T)) S" + and STT: "closedin (top_of_set (S \ T)) T" and "ANR S" "ANR T" "ANR(S \ T)" shows "ANR(S \ T)" proof - - have "\T. openin (subtopology euclidean U) T \ C retract_of T" - if hom: "S \ T homeomorphic C" and UC: "closedin (subtopology euclidean U) C" + have "\T. openin (top_of_set U) T \ C retract_of T" + if hom: "S \ T homeomorphic C" and UC: "closedin (top_of_set U) C" for U and C :: "('a * real) set" proof - obtain f g where hom: "homeomorphism (S \ T) C f g" using hom by (force simp: homeomorphic_def) - have US: "closedin (subtopology euclidean U) (C \ g -` S)" + have US: "closedin (top_of_set U) (C \ g -` S)" apply (rule closedin_trans [OF _ UC]) apply (rule continuous_closedin_preimage_gen [OF _ _ STS]) using hom [unfolded homeomorphism_def] apply blast apply (metis hom homeomorphism_def set_eq_subset) done - have UT: "closedin (subtopology euclidean U) (C \ g -` T)" + have UT: "closedin (top_of_set U) (C \ g -` T)" apply (rule closedin_trans [OF _ UC]) apply (rule continuous_closedin_preimage_gen [OF _ _ STT]) using hom [unfolded homeomorphism_def] apply blast @@ -1291,26 +1291,26 @@ lemma ANR_openin: fixes S :: "'a::euclidean_space set" - assumes "ANR T" and opeTS: "openin (subtopology euclidean T) S" + assumes "ANR T" and opeTS: "openin (top_of_set T) S" shows "ANR S" proof (clarsimp simp only: ANR_eq_absolute_neighbourhood_extensor) fix f :: "'a \ real \ 'a" and U C assume contf: "continuous_on C f" and fim: "f ` C \ S" - and cloUC: "closedin (subtopology euclidean U) C" + and cloUC: "closedin (top_of_set U) C" have "f ` C \ T" using fim opeTS openin_imp_subset by blast obtain W g where "C \ W" - and UW: "openin (subtopology euclidean U) W" + and UW: "openin (top_of_set U) W" and contg: "continuous_on W g" and gim: "g ` W \ T" and geq: "\x. x \ C \ g x = f x" apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \ANR T\ contf \f ` C \ T\ cloUC]) using fim by auto - show "\V g. C \ V \ openin (subtopology euclidean U) V \ continuous_on V g \ g ` V \ S \ (\x\C. g x = f x)" + show "\V g. C \ V \ openin (top_of_set U) V \ continuous_on V g \ g ` V \ S \ (\x\C. g x = f x)" proof (intro exI conjI) show "C \ W \ g -` S" using \C \ W\ fim geq by blast - show "openin (subtopology euclidean U) (W \ g -` S)" + show "openin (top_of_set U) (W \ g -` S)" by (metis (mono_tags, lifting) UW contg continuous_openin_preimage gim opeTS openin_trans) show "continuous_on (W \ g -` S) g" by (blast intro: continuous_on_subset [OF contg]) @@ -1323,20 +1323,20 @@ lemma ENR_openin: fixes S :: "'a::euclidean_space set" - assumes "ENR T" and opeTS: "openin (subtopology euclidean T) S" + assumes "ENR T" and opeTS: "openin (top_of_set T) S" shows "ENR S" using assms apply (simp add: ENR_ANR) using ANR_openin locally_open_subset by blast lemma ANR_neighborhood_retract: fixes S :: "'a::euclidean_space set" - assumes "ANR U" "S retract_of T" "openin (subtopology euclidean U) T" + assumes "ANR U" "S retract_of T" "openin (top_of_set U) T" shows "ANR S" using ANR_openin ANR_retract_of_ANR assms by blast lemma ENR_neighborhood_retract: fixes S :: "'a::euclidean_space set" - assumes "ENR U" "S retract_of T" "openin (subtopology euclidean U) T" + assumes "ENR U" "S retract_of T" "openin (top_of_set U) T" shows "ENR S" using ENR_openin ENR_retract_of_ENR assms by blast @@ -1438,7 +1438,7 @@ proof - obtain U and T :: "('a \ real) set" where "convex U" "U \ {}" - and UT: "closedin (subtopology euclidean U) T" + and UT: "closedin (top_of_set U) T" and "S homeomorphic T" apply (rule homeomorphic_closedin_convex [of S]) using aff_dim_le_DIM [of S] apply auto @@ -1447,7 +1447,7 @@ by (meson ANR_imp_absolute_neighbourhood_retract assms convex_imp_locally_path_connected locally_open_subset retract_of_locally_path_connected) then have S: "locally path_connected S" - if "openin (subtopology euclidean U) V" "T retract_of V" "U \ {}" for V + if "openin (top_of_set U) V" "T retract_of V" "U \ {}" for V using \S homeomorphic T\ homeomorphic_locally homeomorphic_path_connectedness by blast show ?thesis using assms @@ -1494,11 +1494,11 @@ proof (clarsimp simp only: ANR_eq_absolute_neighbourhood_extensor) fix f :: " ('a \ 'b) \ real \ 'a \ 'b" and U C assume "continuous_on C f" and fim: "f ` C \ S \ T" - and cloUC: "closedin (subtopology euclidean U) C" + and cloUC: "closedin (top_of_set U) C" have contf1: "continuous_on C (fst \ f)" by (simp add: \continuous_on C f\ continuous_on_fst) obtain W1 g where "C \ W1" - and UW1: "openin (subtopology euclidean U) W1" + and UW1: "openin (top_of_set U) W1" and contg: "continuous_on W1 g" and gim: "g ` W1 \ S" and geq: "\x. x \ C \ g x = (fst \ f) x" @@ -1508,7 +1508,7 @@ have contf2: "continuous_on C (snd \ f)" by (simp add: \continuous_on C f\ continuous_on_snd) obtain W2 h where "C \ W2" - and UW2: "openin (subtopology euclidean U) W2" + and UW2: "openin (top_of_set U) W2" and conth: "continuous_on W2 h" and him: "h ` W2 \ T" and heq: "\x. x \ C \ h x = (snd \ f) x" @@ -1516,12 +1516,12 @@ using fim apply auto done show "\V g. C \ V \ - openin (subtopology euclidean U) V \ + openin (top_of_set U) V \ continuous_on V g \ g ` V \ S \ T \ (\x\C. g x = f x)" proof (intro exI conjI) show "C \ W1 \ W2" by (simp add: \C \ W1\ \C \ W2\) - show "openin (subtopology euclidean U) (W1 \ W2)" + show "openin (top_of_set U) (W1 \ W2)" by (simp add: UW1 UW2 openin_Int) show "continuous_on (W1 \ W2) (\x. (g x, h x))" by (metis (no_types) contg conth continuous_on_Pair continuous_on_subset inf_commute inf_le1) @@ -3706,7 +3706,7 @@ apply (auto simp: False closest_point_self affine_imp_convex closest_point_in_set continuous_on_closest_point) done finally have "rel_frontier S retract_of {x. closest_point (affine hull S) x \ a}" . - moreover have "openin (subtopology euclidean UNIV) (UNIV \ closest_point (affine hull S) -` (- {a}))" + moreover have "openin (top_of_set UNIV) (UNIV \ closest_point (affine hull S) -` (- {a}))" apply (rule continuous_openin_preimage_gen) apply (auto simp: False affine_imp_convex continuous_on_closest_point) done @@ -3726,7 +3726,7 @@ lemma ENR_closedin_Un_local: fixes S :: "'a::euclidean_space set" shows "\ENR S; ENR T; ENR(S \ T); - closedin (subtopology euclidean (S \ T)) S; closedin (subtopology euclidean (S \ T)) T\ + closedin (top_of_set (S \ T)) S; closedin (top_of_set (S \ T)) T\ \ ENR(S \ T)" by (simp add: ENR_ANR ANR_closed_Un_local locally_compact_closedin_Un) @@ -3743,8 +3743,8 @@ lemma retract_from_Un_Int: fixes S :: "'a::euclidean_space set" - assumes clS: "closedin (subtopology euclidean (S \ T)) S" - and clT: "closedin (subtopology euclidean (S \ T)) T" + assumes clS: "closedin (top_of_set (S \ T)) S" + and clT: "closedin (top_of_set (S \ T)) T" and Un: "(S \ T) retract_of U" and Int: "(S \ T) retract_of T" shows "S retract_of U" proof - @@ -3764,8 +3764,8 @@ lemma AR_from_Un_Int_local: fixes S :: "'a::euclidean_space set" - assumes clS: "closedin (subtopology euclidean (S \ T)) S" - and clT: "closedin (subtopology euclidean (S \ T)) T" + assumes clS: "closedin (top_of_set (S \ T)) S" + and clT: "closedin (top_of_set (S \ T)) T" and Un: "AR(S \ T)" and Int: "AR(S \ T)" shows "AR S" apply (rule AR_retract_of_AR [OF Un]) @@ -3773,8 +3773,8 @@ lemma AR_from_Un_Int_local': fixes S :: "'a::euclidean_space set" - assumes "closedin (subtopology euclidean (S \ T)) S" - and "closedin (subtopology euclidean (S \ T)) T" + assumes "closedin (top_of_set (S \ T)) S" + and "closedin (top_of_set (S \ T)) T" and "AR(S \ T)" "AR(S \ T)" shows "AR T" using AR_from_Un_Int_local [of T S] assms by (simp add: Un_commute Int_commute) @@ -3787,13 +3787,13 @@ lemma ANR_from_Un_Int_local: fixes S :: "'a::euclidean_space set" - assumes clS: "closedin (subtopology euclidean (S \ T)) S" - and clT: "closedin (subtopology euclidean (S \ T)) T" + assumes clS: "closedin (top_of_set (S \ T)) S" + and clT: "closedin (top_of_set (S \ T)) T" and Un: "ANR(S \ T)" and Int: "ANR(S \ T)" shows "ANR S" proof - - obtain V where clo: "closedin (subtopology euclidean (S \ T)) (S \ T)" - and ope: "openin (subtopology euclidean (S \ T)) V" + obtain V where clo: "closedin (top_of_set (S \ T)) (S \ T)" + and ope: "openin (top_of_set (S \ T)) V" and ret: "S \ T retract_of V" using ANR_imp_neighbourhood_retract [OF Int] by (metis clS clT closedin_Int) then obtain r where r: "continuous_on V r" and rim: "r ` V \ S \ T" and req: "\x\S \ T. r x = x" @@ -3808,9 +3808,9 @@ using Vsub by blast have "continuous_on (S \ V \ T) (\x. if x \ S then x else r x)" proof (rule continuous_on_cases_local) - show "closedin (subtopology euclidean (S \ V \ T)) S" + show "closedin (top_of_set (S \ V \ T)) S" using clS closedin_subset_trans inf.boundedE by blast - show "closedin (subtopology euclidean (S \ V \ T)) (V \ T)" + show "closedin (top_of_set (S \ V \ T)) (V \ T)" using clT Vsup by (auto simp: closedin_closed) show "continuous_on (V \ T) r" by (meson Int_lower1 continuous_on_subset r) @@ -3938,7 +3938,7 @@ lemma absolute_retract_imp_AR_gen: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" - assumes "S retract_of T" "convex T" "T \ {}" "S homeomorphic S'" "closedin (subtopology euclidean U) S'" + assumes "S retract_of T" "convex T" "T \ {}" "S homeomorphic S'" "closedin (top_of_set U) S'" shows "S' retract_of U" proof - have "AR T" @@ -3972,7 +3972,7 @@ lemma ENR_from_Un_Int_gen: fixes S :: "'a::euclidean_space set" - assumes "closedin (subtopology euclidean (S \ T)) S" "closedin (subtopology euclidean (S \ T)) T" "ENR(S \ T)" "ENR(S \ T)" + assumes "closedin (top_of_set (S \ T)) S" "closedin (top_of_set (S \ T)) T" "ENR(S \ T)" "ENR(S \ T)" shows "ENR S" apply (simp add: ENR_ANR) using ANR_from_Un_Int_local ENR_ANR assms locally_compact_closedin by blast @@ -4124,7 +4124,7 @@ theorem Borsuk_homotopy_extension_homotopic: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes cloTS: "closedin (subtopology euclidean T) S" + assumes cloTS: "closedin (top_of_set T) S" and anr: "(ANR S \ ANR T) \ ANR U" and contf: "continuous_on T f" and "f ` T \ U" @@ -4140,15 +4140,15 @@ using assms by (auto simp: homotopic_with_def) define h' where "h' \ \z. if snd z \ S then h z else (f \ snd) z" define B where "B \ {0::real} \ T \ {0..1} \ S" - have clo0T: "closedin (subtopology euclidean ({0..1} \ T)) ({0::real} \ T)" + have clo0T: "closedin (top_of_set ({0..1} \ T)) ({0::real} \ T)" by (simp add: closedin_subtopology_refl closedin_Times) - moreover have cloT1S: "closedin (subtopology euclidean ({0..1} \ T)) ({0..1} \ S)" + moreover have cloT1S: "closedin (top_of_set ({0..1} \ T)) ({0..1} \ S)" by (simp add: closedin_subtopology_refl closedin_Times assms) - ultimately have clo0TB:"closedin (subtopology euclidean ({0..1} \ T)) B" + ultimately have clo0TB:"closedin (top_of_set ({0..1} \ T)) B" by (auto simp: B_def) - have cloBS: "closedin (subtopology euclidean B) ({0..1} \ S)" + have cloBS: "closedin (top_of_set B) ({0..1} \ S)" by (metis (no_types) Un_subset_iff B_def closedin_subset_trans [OF cloT1S] clo0TB closedin_imp_subset closedin_self) - moreover have cloBT: "closedin (subtopology euclidean B) ({0} \ T)" + moreover have cloBT: "closedin (top_of_set B) ({0} \ T)" using \S \ T\ closedin_subset_trans [OF clo0T] by (metis B_def Un_upper1 clo0TB closedin_closed inf_le1) moreover have "continuous_on ({0} \ T) (f \ snd)" @@ -4161,7 +4161,7 @@ done have "image h' B \ U" using \f ` T \ U\ him by (auto simp: h'_def B_def) - obtain V k where "B \ V" and opeTV: "openin (subtopology euclidean ({0..1} \ T)) V" + obtain V k where "B \ V" and opeTV: "openin (top_of_set ({0..1} \ T)) V" and contk: "continuous_on V k" and kim: "k ` V \ U" and keq: "\x. x \ B \ k x = h' x" using anr @@ -4177,7 +4177,7 @@ apply (simp_all add: ANR_Times convex_imp_ANR ANR_singleton ST eq) done note Vk = that - have *: thesis if "openin (subtopology euclidean ({0..1::real} \ T)) V" + have *: thesis if "openin (top_of_set ({0..1::real} \ T)) V" "retraction V B r" for V r using that apply (clarsimp simp add: retraction_def) @@ -4195,14 +4195,14 @@ show ?thesis by blast qed define S' where "S' \ {x. \u::real. u \ {0..1} \ (u, x::'a) \ {0..1} \ T - V}" - have "closedin (subtopology euclidean T) S'" + have "closedin (top_of_set T) S'" unfolding S'_def apply (rule closedin_compact_projection, blast) using closedin_self opeTV by blast have S'_def: "S' = {x. \u::real. (u, x::'a) \ {0..1} \ T - V}" by (auto simp: S'_def) - have cloTS': "closedin (subtopology euclidean T) S'" - using S'_def \closedin (subtopology euclidean T) S'\ by blast + have cloTS': "closedin (top_of_set T) S'" + using S'_def \closedin (top_of_set T) S'\ by blast have "S \ S' = {}" using S'_def B_def \B \ V\ by force obtain a :: "'a \ real" where conta: "continuous_on T a" @@ -4269,7 +4269,7 @@ assume ?lhs then obtain c where c: "homotopic_with (\x. True) S T (\x. c) f" by (blast intro: homotopic_with_symD) - have "closedin (subtopology euclidean UNIV) S" + have "closedin (top_of_set UNIV) S" using \closed S\ closed_closedin by fastforce then obtain g where "continuous_on UNIV g" "range g \ T" "\x. x \ S \ g x = f x" @@ -4412,7 +4412,7 @@ "g ` (S \ connected_component_set (- S) a) \ sphere 0 1" "\x. x \ S \ g x = (x - a) /\<^sub>R norm (x - a)" proof (rule Borsuk_homotopy_extension_homotopic) - show "closedin (subtopology euclidean ?T) S" + show "closedin (top_of_set ?T) S" by (simp add: \compact S\ closed_subset compact_imp_closed) show "continuous_on ?T (\x. (x - b) /\<^sub>R norm (x - b))" by (simp add: \b \ S\ notcc continuous_on_Borsuk_map) @@ -4467,8 +4467,8 @@ subsubsection\More extension theorems\ lemma extension_from_clopen: - assumes ope: "openin (subtopology euclidean S) T" - and clo: "closedin (subtopology euclidean S) T" + assumes ope: "openin (top_of_set S) T" + and clo: "closedin (top_of_set S) T" and contf: "continuous_on T f" and fim: "f ` T \ U" and null: "U = {} \ S = {}" obtains g where "continuous_on S g" "g ` S \ U" "\x. x \ T \ g x = f x" proof (cases "U = {}") @@ -4503,8 +4503,8 @@ and C: "C \ components S" and contf: "continuous_on C f" and fim: "f ` C \ U" obtains g where "continuous_on S g" "g ` S \ U" "\x. x \ C \ g x = f x" proof - - obtain T g where ope: "openin (subtopology euclidean S) T" - and clo: "closedin (subtopology euclidean S) T" + obtain T g where ope: "openin (top_of_set S) T" + and clo: "closedin (top_of_set S) T" and "C \ T" and contg: "continuous_on T g" and gim: "g ` T \ U" and gf: "\x. x \ C \ g x = f x" using S @@ -4514,7 +4514,7 @@ by (metis C \locally connected S\ openin_components_locally_connected closedin_component contf fim order_refl that) next assume "compact S" - then obtain W g where "C \ W" and opeW: "openin (subtopology euclidean S) W" + then obtain W g where "C \ W" and opeW: "openin (top_of_set S) W" and contg: "continuous_on W g" and gim: "g ` W \ U" and gf: "\x. x \ C \ g x = f x" using ANR_imp_absolute_neighbourhood_extensor [of U C f S] C \ANR U\ closedin_component contf fim by blast @@ -4522,11 +4522,11 @@ by (auto simp: openin_open) moreover have "locally compact S" by (simp add: \compact S\ closed_imp_locally_compact compact_imp_closed) - ultimately obtain K where opeK: "openin (subtopology euclidean S) K" and "compact K" "C \ K" "K \ V" + ultimately obtain K where opeK: "openin (top_of_set S) K" and "compact K" "C \ K" "K \ V" by (metis C Int_subset_iff \C \ W\ \compact S\ compact_components Sura_Bura_clopen_subset) show ?thesis proof - show "closedin (subtopology euclidean S) K" + show "closedin (top_of_set S) K" by (meson \compact K\ \compact S\ closedin_compact_eq opeK openin_imp_subset) show "continuous_on K g" by (metis Int_subset_iff V \K \ V\ contg continuous_on_subset opeK openin_subtopology subset_eq) @@ -4545,13 +4545,13 @@ lemma tube_lemma: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact S" and S: "S \ {}" "(\x. (x,a)) ` S \ U" - and ope: "openin (subtopology euclidean (S \ T)) U" - obtains V where "openin (subtopology euclidean T) V" "a \ V" "S \ V \ U" + and ope: "openin (top_of_set (S \ T)) U" + obtains V where "openin (top_of_set T) V" "a \ V" "S \ V \ U" proof - let ?W = "{y. \x. x \ S \ (x, y) \ (S \ T - U)}" - have "U \ S \ T" "closedin (subtopology euclidean (S \ T)) (S \ T - U)" + have "U \ S \ T" "closedin (top_of_set (S \ T)) (S \ T - U)" using ope by (auto simp: openin_closedin_eq) - then have "closedin (subtopology euclidean T) ?W" + then have "closedin (top_of_set T) ?W" using \compact S\ closedin_compact_projection by blast moreover have "a \ T - ?W" using \U \ S \ T\ S by auto @@ -4564,16 +4564,16 @@ lemma tube_lemma_gen: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact S" "S \ {}" "T \ T'" "S \ T \ U" - and ope: "openin (subtopology euclidean (S \ T')) U" - obtains V where "openin (subtopology euclidean T') V" "T \ V" "S \ V \ U" + and ope: "openin (top_of_set (S \ T')) U" + obtains V where "openin (top_of_set T') V" "T \ V" "S \ V \ U" proof - - have "\x. x \ T \ \V. openin (subtopology euclidean T') V \ x \ V \ S \ V \ U" + have "\x. x \ T \ \V. openin (top_of_set T') V \ x \ V \ S \ V \ U" using assms by (auto intro: tube_lemma [OF \compact S\]) - then obtain F where F: "\x. x \ T \ openin (subtopology euclidean T') (F x) \ x \ F x \ S \ F x \ U" + then obtain F where F: "\x. x \ T \ openin (top_of_set T') (F x) \ x \ F x \ S \ F x \ U" by metis show ?thesis proof - show "openin (subtopology euclidean T') (\(F ` T))" + show "openin (top_of_set T') (\(F ` T))" using F by blast show "T \ \(F ` T)" using F by blast @@ -4586,9 +4586,9 @@ fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and fim: "f ` S \ U" and contg: "continuous_on S g" and gim: "g ` S \ U" - and clo: "closedin (subtopology euclidean S) T" + and clo: "closedin (top_of_set S) T" and "ANR U" and hom: "homotopic_with (\x. True) T U f g" - obtains V where "T \ V" "openin (subtopology euclidean S) V" + obtains V where "T \ V" "openin (top_of_set S) V" "homotopic_with (\x. True) V U f g" proof - have "T \ S" @@ -4604,11 +4604,11 @@ have "continuous_on(?S0 \ (?S1 \ {0..1} \ T)) h'" unfolding h'_def proof (intro continuous_on_cases_local) - show "closedin (subtopology euclidean (?S0 \ (?S1 \ {0..1} \ T))) ?S0" - "closedin (subtopology euclidean (?S1 \ {0..1} \ T)) ?S1" + show "closedin (top_of_set (?S0 \ (?S1 \ {0..1} \ T))) ?S0" + "closedin (top_of_set (?S1 \ {0..1} \ T)) ?S1" using \T \ S\ by (force intro: closedin_Times closedin_subset_trans [of "{0..1} \ S"])+ - show "closedin (subtopology euclidean (?S0 \ (?S1 \ {0..1} \ T))) (?S1 \ {0..1} \ T)" - "closedin (subtopology euclidean (?S1 \ {0..1} \ T)) ({0..1} \ T)" + show "closedin (top_of_set (?S0 \ (?S1 \ {0..1} \ T))) (?S1 \ {0..1} \ T)" + "closedin (top_of_set (?S1 \ {0..1} \ T)) ({0..1} \ T)" using \T \ S\ by (force intro: clo closedin_Times closedin_subset_trans [of "{0..1} \ S"])+ show "continuous_on (?S0) (\x. f (snd x))" by (intro continuous_intros continuous_on_compose2 [OF contf]) auto @@ -4619,16 +4619,16 @@ by (metis Sigma_Un_distrib1 Un_assoc insert_is_Un) moreover have "h' ` ({0,1} \ S \ {0..1} \ T) \ U" using fim gim him \T \ S\ unfolding h'_def by force - moreover have "closedin (subtopology euclidean ({0..1::real} \ S)) ({0,1} \ S \ {0..1::real} \ T)" + moreover have "closedin (top_of_set ({0..1::real} \ S)) ({0,1} \ S \ {0..1::real} \ T)" by (intro closedin_Times closedin_Un clo) (simp_all add: closed_subset) ultimately obtain W k where W: "({0,1} \ S) \ ({0..1} \ T) \ W" - and opeW: "openin (subtopology euclidean ({0..1} \ S)) W" + and opeW: "openin (top_of_set ({0..1} \ S)) W" and contk: "continuous_on W k" and kim: "k ` W \ U" and kh': "\x. x \ ({0,1} \ S) \ ({0..1} \ T) \ k x = h' x" by (metis ANR_imp_absolute_neighbourhood_extensor [OF \ANR U\, of "({0,1} \ S) \ ({0..1} \ T)" h' "{0..1} \ S"]) - obtain T' where opeT': "openin (subtopology euclidean S) T'" + obtain T' where opeT': "openin (top_of_set S) T'" and "T \ T'" and TW: "{0..1} \ T' \ W" using tube_lemma_gen [of "{0..1::real}" T S W] W \T \ S\ opeW by auto moreover have "homotopic_with (\x. True) T' U f g" @@ -4649,8 +4649,8 @@ text\ Homotopy on a union of closed-open sets.\ proposition%unimportant homotopic_on_clopen_Union: fixes \ :: "'a::euclidean_space set set" - assumes "\S. S \ \ \ closedin (subtopology euclidean (\\)) S" - and "\S. S \ \ \ openin (subtopology euclidean (\\)) S" + assumes "\S. S \ \ \ closedin (top_of_set (\\)) S" + and "\S. S \ \ \ openin (top_of_set (\\)) S" and "\S. S \ \ \ homotopic_with (\x. True) S T f g" shows "homotopic_with (\x. True) (\\) T f g" proof - @@ -4665,8 +4665,8 @@ case False then obtain V :: "nat \ 'a set" where V: "range V = \" using range_from_nat_into \countable \\ by metis - with \\ \ \\ have clo: "\n. closedin (subtopology euclidean (\\)) (V n)" - and ope: "\n. openin (subtopology euclidean (\\)) (V n)" + with \\ \ \\ have clo: "\n. closedin (top_of_set (\\)) (V n)" + and ope: "\n. openin (top_of_set (\\)) (V n)" and hom: "\n. homotopic_with (\x. True) (V n) T f g" using assms by auto then obtain h where conth: "\n. continuous_on ({0..1::real} \ V n) (h n)" @@ -4680,20 +4680,20 @@ and eq: "\x i. \x \ {0..1} \ \(V ` UNIV) \ {0..1} \ (V i - (\m \ \ x = h i x" proof (rule pasting_lemma_exists) - show "{0..1} \ \(V ` UNIV) \ (\i. {0..1::real} \ (V i - (\m (\i. {0..1::real} \ (V i - (\m \(V ` UNIV))) + show "openin (top_of_set ({0..1} \ \(V ` UNIV))) ({0..1::real} \ (V i - (\m(V ` UNIV))) (V i)" + show "openin (top_of_set (\(V ` UNIV))) (V i)" using ope V eqU by auto - show "closedin (subtopology euclidean (\(V ` UNIV))) (\m(V ` UNIV))) (\m (V i - (\mi j x. x \ {0..1} \ \(V ` UNIV) \ - {0..1} \ (V i - (\m {0..1} \ (V j - (\m (V i - \ (V ` {..i j x. x \ topspace ?X \ {0..1} \ (V i - (\m {0..1} \ (V j - (\m h i x = h j x" by clarsimp (metis lessThan_iff linorder_neqE_nat) qed auto @@ -4738,8 +4738,8 @@ by (simp add: homotopic_with_subset_left in_components_subset) next assume R: ?rhs - have "\U. C \ U \ closedin (subtopology euclidean S) U \ - openin (subtopology euclidean S) U \ + have "\U. C \ U \ closedin (top_of_set S) U \ + openin (top_of_set S) U \ homotopic_with (\x. True) U T f g" if C: "C \ components S" for C proof - have "C \ S" @@ -4750,9 +4750,9 @@ assume "locally connected S" show ?thesis proof (intro exI conjI) - show "closedin (subtopology euclidean S) C" + show "closedin (top_of_set S) C" by (simp add: closedin_component that) - show "openin (subtopology euclidean S) C" + show "openin (top_of_set S) C" by (simp add: \locally connected S\ openin_components_locally_connected that) show "homotopic_with (\x. True) C T f g" by (simp add: R that) @@ -4761,7 +4761,7 @@ assume "compact S" have hom: "homotopic_with (\x. True) C T f g" using R that by blast - obtain U where "C \ U" and opeU: "openin (subtopology euclidean S) U" + obtain U where "C \ U" and opeU: "openin (top_of_set S) U" and hom: "homotopic_with (\x. True) U T f g" using homotopic_neighbourhood_extension [OF contf fim contg gim _ \ANR T\ hom] \C \ components S\ closedin_component by blast @@ -4769,11 +4769,11 @@ by (auto simp: openin_open) moreover have "locally compact S" by (simp add: \compact S\ closed_imp_locally_compact compact_imp_closed) - ultimately obtain K where opeK: "openin (subtopology euclidean S) K" and "compact K" "C \ K" "K \ V" + ultimately obtain K where opeK: "openin (top_of_set S) K" and "compact K" "C \ K" "K \ V" by (metis C Int_subset_iff Sura_Bura_clopen_subset \C \ U\ \compact S\ compact_components) show ?thesis proof (intro exI conjI) - show "closedin (subtopology euclidean S) K" + show "closedin (top_of_set S) K" by (meson \compact K\ \compact S\ closedin_compact_eq opeK openin_imp_subset) show "homotopic_with (\x. True) K T f g" using V \K \ V\ hom homotopic_with_subset_left opeK openin_imp_subset by fastforce @@ -4781,8 +4781,8 @@ qed qed then obtain \ where \: "\C. C \ components S \ C \ \ C" - and clo\: "\C. C \ components S \ closedin (subtopology euclidean S) (\ C)" - and ope\: "\C. C \ components S \ openin (subtopology euclidean S) (\ C)" + and clo\: "\C. C \ components S \ closedin (top_of_set S) (\ C)" + and ope\: "\C. C \ components S \ openin (top_of_set S) (\ C)" and hom\: "\C. C \ components S \ homotopic_with (\x. True) (\ C) T f g" by metis have Seq: "S = \ (\ ` components S)" diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Cauchy_Integral_Theorem.thy --- a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Tue Mar 19 16:14:51 2019 +0000 @@ -7001,13 +7001,13 @@ by (metis funpow_add o_apply) with that show ?thesis by auto qed - have 1: "openin (subtopology euclidean S) (\n. {w \ S. (deriv ^^ n) f w = 0})" + have 1: "openin (top_of_set S) (\n. {w \ S. (deriv ^^ n) f w = 0})" apply (rule open_subset, force) using \open S\ apply (simp add: open_contains_ball Ball_def) apply (erule all_forward) using "*" by auto blast+ - have 2: "closedin (subtopology euclidean S) (\n. {w \ S. (deriv ^^ n) f w = 0})" + have 2: "closedin (top_of_set S) (\n. {w \ S. (deriv ^^ n) f w = 0})" using assms by (auto intro: continuous_closedin_preimage_constant holomorphic_on_imp_continuous_on holomorphic_higher_deriv) obtain e where "e>0" and e: "ball w e \ S" using openE [OF \open S\ \w \ S\] . diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Change_Of_Vars.thy --- a/src/HOL/Analysis/Change_Of_Vars.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Change_Of_Vars.thy Tue Mar 19 16:14:51 2019 +0000 @@ -1567,24 +1567,24 @@ show "countable {A. \i j. A $ i $ j \ \}" using countable_vector [OF countable_vector, of "\i j. \"] by (simp add: countable_rat) qed blast - have *: "\U \ {} \ closedin (subtopology euclidean S) (S \ \ U)\ - \ closedin (subtopology euclidean S) (S \ \ U)" for U + have *: "\U \ {} \ closedin (top_of_set S) (S \ \ U)\ + \ closedin (top_of_set S) (S \ \ U)" for U by fastforce have eq: "{x::(real,'m)vec. P x \ (Q x \ R x)} = {x. P x \ \ Q x} \ {x. P x \ R x}" for P Q R by auto have sets: "S \ (\y\S. {x \ S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x)}) \ sets lebesgue" for e A d proof - - have clo: "closedin (subtopology euclidean S) + have clo: "closedin (top_of_set S) {x \ S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x)}" for y proof - have cont1: "continuous_on S (\x. norm (y - x))" and cont2: "continuous_on S (\x. e * norm (y - x) - norm (f y - f x - (A *v y - A *v x)))" by (force intro: contf continuous_intros)+ - have clo1: "closedin (subtopology euclidean S) {x \ S. d \ norm (y - x)}" + have clo1: "closedin (top_of_set S) {x \ S. d \ norm (y - x)}" using continuous_closedin_preimage [OF cont1, of "{d..}"] by (simp add: vimage_def Int_def) - have clo2: "closedin (subtopology euclidean S) + have clo2: "closedin (top_of_set S) {x \ S. norm (f y - f x - (A *v y - A *v x)) \ e * norm (y - x)}" using continuous_closedin_preimage [OF cont2, of "{0..}"] by (simp add: vimage_def Int_def) show ?thesis diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Connected.thy --- a/src/HOL/Analysis/Connected.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Connected.thy Tue Mar 19 16:14:51 2019 +0000 @@ -14,8 +14,8 @@ lemma connected_local: "connected S \ \ (\e1 e2. - openin (subtopology euclidean S) e1 \ - openin (subtopology euclidean S) e2 \ + openin (top_of_set S) e1 \ + openin (top_of_set S) e2 \ S \ e1 \ e2 \ e1 \ e2 = {} \ e1 \ {} \ @@ -39,8 +39,8 @@ qed lemma connected_clopen: "connected S \ - (\T. openin (subtopology euclidean S) T \ - closedin (subtopology euclidean S) T \ T = {} \ T = S)" (is "?lhs \ ?rhs") + (\T. openin (top_of_set S) T \ + closedin (top_of_set S) T \ T = {} \ T = S)" (is "?lhs \ ?rhs") proof - have "\ connected S \ (\e1 e2. open e1 \ open (- e2) \ S \ e1 \ (- e2) \ e1 \ (- e2) \ S = {} \ e1 \ S \ {} \ (- e2) \ S \ {})" @@ -402,7 +402,7 @@ by blast qed -lemma closedin_connected_component: "closedin (subtopology euclidean s) (connected_component_set s x)" +lemma closedin_connected_component: "closedin (top_of_set s) (connected_component_set s x)" proof (cases "connected_component_set s x = {}") case True then show ?thesis @@ -423,7 +423,7 @@ qed lemma closedin_component: - "C \ components s \ closedin (subtopology euclidean s) C" + "C \ components s \ closedin (top_of_set s) C" using closedin_connected_component componentsE by blast @@ -433,7 +433,7 @@ lemma continuous_levelset_openin_cases: fixes f :: "_ \ 'b::t1_space" shows "connected s \ continuous_on s f \ - openin (subtopology euclidean s) {x \ s. f x = a} + openin (top_of_set s) {x \ s. f x = a} \ (\x \ s. f x \ a) \ (\x \ s. f x = a)" unfolding connected_clopen using continuous_closedin_preimage_constant by auto @@ -441,7 +441,7 @@ lemma continuous_levelset_openin: fixes f :: "_ \ 'b::t1_space" shows "connected s \ continuous_on s f \ - openin (subtopology euclidean s) {x \ s. f x = a} \ + openin (top_of_set s) {x \ s. f x = a} \ (\x \ s. f x = a) \ (\x \ s. f x = a)" using continuous_levelset_openin_cases[of s f ] by meson @@ -469,8 +469,8 @@ assumes "connected T" and contf: "continuous_on S f" and fim: "f ` S = T" and opT: "\U. U \ T - \ openin (subtopology euclidean S) (S \ f -` U) \ - openin (subtopology euclidean T) U" + \ openin (top_of_set S) (S \ f -` U) \ + openin (top_of_set T) U" and connT: "\y. y \ T \ connected (S \ f -` {y})" shows "connected S" proof (rule connectedI) @@ -492,9 +492,9 @@ qed ultimately have UU: "(S \ f -` f ` (S \ U)) = S \ U" and VV: "(S \ f -` f ` (S \ V)) = S \ V" by auto - have opeU: "openin (subtopology euclidean T) (f ` (S \ U))" + have opeU: "openin (top_of_set T) (f ` (S \ U))" by (metis UU \open U\ fim image_Int_subset le_inf_iff opT openin_open_Int) - have opeV: "openin (subtopology euclidean T) (f ` (S \ V))" + have opeV: "openin (top_of_set T) (f ` (S \ V))" by (metis opT fim VV \open V\ openin_open_Int image_Int_subset inf.bounded_iff) have "T \ f ` (S \ U) \ f ` (S \ V)" using \S \ U \ V\ fim by auto @@ -505,7 +505,7 @@ lemma connected_open_monotone_preimage: assumes contf: "continuous_on S f" and fim: "f ` S = T" - and ST: "\C. openin (subtopology euclidean S) C \ openin (subtopology euclidean T) (f ` C)" + and ST: "\C. openin (top_of_set S) C \ openin (top_of_set T) (f ` C)" and connT: "\y. y \ T \ connected (S \ f -` {y})" and "connected C" "C \ T" shows "connected (S \ f -` C)" @@ -525,12 +525,12 @@ ultimately show ?thesis by metis qed - have "\U. openin (subtopology euclidean (S \ f -` C)) U - \ openin (subtopology euclidean C) (f ` U)" + have "\U. openin (top_of_set (S \ f -` C)) U + \ openin (top_of_set C) (f ` U)" using open_map_restrict [OF _ ST \C \ T\] by metis then show "\D. D \ C - \ openin (subtopology euclidean (S \ f -` C)) (S \ f -` C \ f -` D) = - openin (subtopology euclidean C) D" + \ openin (top_of_set (S \ f -` C)) (S \ f -` C \ f -` D) = + openin (top_of_set C) D" using open_map_imp_quotient_map [of "(S \ f -` C)" f] contf' by (simp add: eqC) qed qed @@ -538,7 +538,7 @@ lemma connected_closed_monotone_preimage: assumes contf: "continuous_on S f" and fim: "f ` S = T" - and ST: "\C. closedin (subtopology euclidean S) C \ closedin (subtopology euclidean T) (f ` C)" + and ST: "\C. closedin (top_of_set S) C \ closedin (top_of_set T) (f ` C)" and connT: "\y. y \ T \ connected (S \ f -` {y})" and "connected C" "C \ T" shows "connected (S \ f -` C)" @@ -558,12 +558,12 @@ ultimately show ?thesis by metis qed - have "\U. closedin (subtopology euclidean (S \ f -` C)) U - \ closedin (subtopology euclidean C) (f ` U)" + have "\U. closedin (top_of_set (S \ f -` C)) U + \ closedin (top_of_set C) (f ` U)" using closed_map_restrict [OF _ ST \C \ T\] by metis then show "\D. D \ C - \ openin (subtopology euclidean (S \ f -` C)) (S \ f -` C \ f -` D) = - openin (subtopology euclidean C) D" + \ openin (top_of_set (S \ f -` C)) (S \ f -` C \ f -` D) = + openin (top_of_set C) D" using closed_map_imp_quotient_map [of "(S \ f -` C)" f] contf' by (simp add: eqC) qed qed @@ -576,8 +576,8 @@ lemma connected_Un_clopen_in_complement: fixes S U :: "'a::metric_space set" assumes "connected S" "connected U" "S \ U" - and opeT: "openin (subtopology euclidean (U - S)) T" - and cloT: "closedin (subtopology euclidean (U - S)) T" + and opeT: "openin (top_of_set (U - S)) T" + and cloT: "closedin (top_of_set (U - S)) T" shows "connected (S \ T)" proof - have *: "\\x y. P x y \ P y x; \x y. P x y \ S \ x \ S \ y; @@ -587,11 +587,11 @@ unfolding connected_closedin_eq proof (rule *) fix H1 H2 - assume H: "closedin (subtopology euclidean (S \ T)) H1 \ - closedin (subtopology euclidean (S \ T)) H2 \ + assume H: "closedin (top_of_set (S \ T)) H1 \ + closedin (top_of_set (S \ T)) H2 \ H1 \ H2 = S \ T \ H1 \ H2 = {} \ H1 \ {} \ H2 \ {}" - then have clo: "closedin (subtopology euclidean S) (S \ H1)" - "closedin (subtopology euclidean S) (S \ H2)" + then have clo: "closedin (top_of_set S) (S \ H1)" + "closedin (top_of_set S) (S \ H2)" by (metis Un_upper1 closedin_closed_subset inf_commute)+ have Seq: "S \ (H1 \ H2) = S" by (simp add: H) @@ -606,8 +606,8 @@ using H \connected S\ unfolding connected_closedin by blast next fix H1 H2 - assume H: "closedin (subtopology euclidean (S \ T)) H1 \ - closedin (subtopology euclidean (S \ T)) H2 \ + assume H: "closedin (top_of_set (S \ T)) H1 \ + closedin (top_of_set (S \ T)) H2 \ H1 \ H2 = S \ T \ H1 \ H2 = {} \ H1 \ {} \ H2 \ {}" and "S \ H1" then have H2T: "H2 \ T" @@ -616,17 +616,17 @@ using Diff_iff opeT openin_imp_subset by auto with \S \ U\ have Ueq: "U = (U - S) \ (S \ T)" by auto - have "openin (subtopology euclidean ((U - S) \ (S \ T))) H2" + have "openin (top_of_set ((U - S) \ (S \ T))) H2" proof (rule openin_subtopology_Un) - show "openin (subtopology euclidean (S \ T)) H2" + show "openin (top_of_set (S \ T)) H2" using \H2 \ T\ apply (auto simp: openin_closedin_eq) by (metis Diff_Diff_Int Diff_disjoint Diff_partition Diff_subset H Int_absorb1 Un_Diff) - then show "openin (subtopology euclidean (U - S)) H2" + then show "openin (top_of_set (U - S)) H2" by (meson H2T Un_upper2 opeT openin_subset_trans openin_trans) qed - moreover have "closedin (subtopology euclidean ((U - S) \ (S \ T))) H2" + moreover have "closedin (top_of_set ((U - S) \ (S \ T))) H2" proof (rule closedin_subtopology_Un) - show "closedin (subtopology euclidean (U - S)) H2" + show "closedin (top_of_set (U - S)) H2" using H H2T cloT closedin_subset_trans by (blast intro: closedin_subtopology_Un closedin_trans) qed (simp add: H) @@ -650,33 +650,33 @@ using \connected S\ unfolding connected_closedin_eq not_ex de_Morgan_conj proof clarify fix H3 H4 - assume clo3: "closedin (subtopology euclidean (U - C)) H3" - and clo4: "closedin (subtopology euclidean (U - C)) H4" + assume clo3: "closedin (top_of_set (U - C)) H3" + and clo4: "closedin (top_of_set (U - C)) H4" and "H3 \ H4 = U - C" and "H3 \ H4 = {}" and "H3 \ {}" and "H4 \ {}" and * [rule_format]: - "\H1 H2. \ closedin (subtopology euclidean S) H1 \ - \ closedin (subtopology euclidean S) H2 \ + "\H1 H2. \ closedin (top_of_set S) H1 \ + \ closedin (top_of_set S) H2 \ H1 \ H2 \ S \ H1 \ H2 \ {} \ \ H1 \ {} \ \ H2 \ {}" - then have "H3 \ U-C" and ope3: "openin (subtopology euclidean (U - C)) (U - C - H3)" - and "H4 \ U-C" and ope4: "openin (subtopology euclidean (U - C)) (U - C - H4)" + then have "H3 \ U-C" and ope3: "openin (top_of_set (U - C)) (U - C - H3)" + and "H4 \ U-C" and ope4: "openin (top_of_set (U - C)) (U - C - H4)" by (auto simp: closedin_def) have "C \ {}" "C \ U-S" "connected C" using C in_components_nonempty in_components_subset in_components_maximal by blast+ have cCH3: "connected (C \ H3)" proof (rule connected_Un_clopen_in_complement [OF \connected C\ \connected U\ _ _ clo3]) - show "openin (subtopology euclidean (U - C)) H3" + show "openin (top_of_set (U - C)) H3" apply (simp add: openin_closedin_eq \H3 \ U - C\) apply (simp add: closedin_subtopology) by (metis Diff_cancel Diff_triv Un_Diff clo4 \H3 \ H4 = {}\ \H3 \ H4 = U - C\ closedin_closed inf_commute sup_bot.left_neutral) qed (use clo3 \C \ U - S\ in auto) have cCH4: "connected (C \ H4)" proof (rule connected_Un_clopen_in_complement [OF \connected C\ \connected U\ _ _ clo4]) - show "openin (subtopology euclidean (U - C)) H4" + show "openin (top_of_set (U - C)) H4" apply (simp add: openin_closedin_eq \H4 \ U - C\) apply (simp add: closedin_subtopology) by (metis Diff_cancel Int_commute Un_Diff Un_Diff_Int \H3 \ H4 = {}\ \H3 \ H4 = U - C\ clo3 closedin_closed) qed (use clo4 \C \ U - S\ in auto) - have "closedin (subtopology euclidean S) (S \ H3)" "closedin (subtopology euclidean S) (S \ H4)" + have "closedin (top_of_set S) (S \ H3)" "closedin (top_of_set S) (S \ H4)" using clo3 clo4 \S \ U\ \C \ U - S\ by (auto simp: closedin_closed) moreover have "S \ H3 \ {}" using components_maximal [OF C cCH3] \C \ {}\ \C \ U - S\ \H3 \ {}\ \H3 \ U - C\ by auto @@ -719,8 +719,8 @@ shows "connected S" proof - { fix t u - assume clt: "closedin (subtopology euclidean S) t" - and clu: "closedin (subtopology euclidean S) u" + assume clt: "closedin (top_of_set S) t" + and clu: "closedin (top_of_set S) u" and tue: "t \ u = {}" and tus: "t \ u = S" have conif: "continuous_on S (\x. if x \ t then 0 else 1)" apply (subst tus [symmetric]) diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Continuous_Extension.thy --- a/src/HOL/Analysis/Continuous_Extension.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Continuous_Extension.thy Tue Mar 19 16:14:51 2019 +0000 @@ -33,7 +33,7 @@ by (simp add: supp_sum_def sum_nonneg) have sd_pos: "0 < setdist {x} (S - V)" if "V \ \" "x \ S" "x \ V" for V x proof - - have "closedin (subtopology euclidean S) (S - V)" + have "closedin (top_of_set S) (S - V)" by (simp add: Diff_Diff_Int closedin_def opC openin_open_Int \V \ \\) with that False setdist_pos_le [of "{x}" "S - V"] show ?thesis @@ -67,7 +67,7 @@ fix x assume "x \ S" then obtain X where "open X" and x: "x \ S \ X" and finX: "finite {U \ \. U \ X \ {}}" using assms by blast - then have OSX: "openin (subtopology euclidean S) (S \ X)" by blast + then have OSX: "openin (top_of_set S) (S \ X)" by blast have sumeq: "\x. x \ S \ X \ (\V | V \ \ \ V \ X \ {}. setdist {x} (S - V)) = supp_sum (\V. setdist {x} (S - V)) \" @@ -114,8 +114,8 @@ text \For Euclidean spaces the proof is easy using distances.\ lemma Urysohn_both_ne: - assumes US: "closedin (subtopology euclidean U) S" - and UT: "closedin (subtopology euclidean U) T" + assumes US: "closedin (top_of_set U) S" + and UT: "closedin (top_of_set U) T" and "S \ T = {}" "S \ {}" "T \ {}" "a \ b" obtains f :: "'a::euclidean_space \ 'b::real_normed_vector" where "continuous_on U f" @@ -167,8 +167,8 @@ qed proposition Urysohn_local_strong: - assumes US: "closedin (subtopology euclidean U) S" - and UT: "closedin (subtopology euclidean U) T" + assumes US: "closedin (top_of_set U) S" + and UT: "closedin (top_of_set U) T" and "S \ T = {}" "a \ b" obtains f :: "'a::euclidean_space \ 'b::euclidean_space" where "continuous_on U f" @@ -250,8 +250,8 @@ qed lemma Urysohn_local: - assumes US: "closedin (subtopology euclidean U) S" - and UT: "closedin (subtopology euclidean U) T" + assumes US: "closedin (top_of_set U) S" + and UT: "closedin (top_of_set U) T" and "S \ T = {}" obtains f :: "'a::euclidean_space \ 'b::euclidean_space" where "continuous_on U f" @@ -301,7 +301,7 @@ theorem Dugundji: fixes f :: "'a::{metric_space,second_countable_topology} \ 'b::real_inner" assumes "convex C" "C \ {}" - and cloin: "closedin (subtopology euclidean U) S" + and cloin: "closedin (top_of_set U) S" and contf: "continuous_on S f" and "f ` S \ C" obtains g where "continuous_on U g" "g ` U \ C" "\x. x \ S \ g x = f x" @@ -427,7 +427,7 @@ obtain N where N: "open N" "a \ N" and finN: "finite {U \ \. \a\N. H U a \ 0}" using Hfin False \a \ U\ by auto - have oUS: "openin (subtopology euclidean U) (U - S)" + have oUS: "openin (top_of_set U) (U - S)" using cloin by (simp add: openin_diff) have HcontU: "continuous (at a within U) (H T)" if "T \ \" for T using Hcont [OF \T \ \\] False \a \ U\ \T \ \\ @@ -444,7 +444,7 @@ (\x. \T\{U \ \. \x\N. H U x \ 0}. H T x *\<^sub>R f (\ T))" by (force intro: continuous_intros HcontU)+ next - show "openin (subtopology euclidean U) ((U - S) \ N)" + show "openin (top_of_set U) ((U - S) \ N)" using N oUS openin_trans by blast next show "a \ (U - S) \ N" using False \a \ U\ N by blast @@ -475,7 +475,7 @@ corollary Tietze: fixes f :: "'a::{metric_space,second_countable_topology} \ 'b::real_inner" assumes "continuous_on S f" - and "closedin (subtopology euclidean U) S" + and "closedin (top_of_set U) S" and "0 \ B" and "\x. x \ S \ norm(f x) \ B" obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" @@ -485,7 +485,7 @@ corollary%unimportant Tietze_closed_interval: fixes f :: "'a::{metric_space,second_countable_topology} \ 'b::euclidean_space" assumes "continuous_on S f" - and "closedin (subtopology euclidean U) S" + and "closedin (top_of_set U) S" and "cbox a b \ {}" and "\x. x \ S \ f x \ cbox a b" obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" @@ -496,7 +496,7 @@ corollary%unimportant Tietze_closed_interval_1: fixes f :: "'a::{metric_space,second_countable_topology} \ real" assumes "continuous_on S f" - and "closedin (subtopology euclidean U) S" + and "closedin (top_of_set U) S" and "a \ b" and "\x. x \ S \ f x \ cbox a b" obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" @@ -507,7 +507,7 @@ corollary%unimportant Tietze_open_interval: fixes f :: "'a::{metric_space,second_countable_topology} \ 'b::euclidean_space" assumes "continuous_on S f" - and "closedin (subtopology euclidean U) S" + and "closedin (top_of_set U) S" and "box a b \ {}" and "\x. x \ S \ f x \ box a b" obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" @@ -518,7 +518,7 @@ corollary%unimportant Tietze_open_interval_1: fixes f :: "'a::{metric_space,second_countable_topology} \ real" assumes "continuous_on S f" - and "closedin (subtopology euclidean U) S" + and "closedin (top_of_set U) S" and "a < b" and no: "\x. x \ S \ f x \ box a b" obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" @@ -529,7 +529,7 @@ corollary%unimportant Tietze_unbounded: fixes f :: "'a::{metric_space,second_countable_topology} \ 'b::real_inner" assumes "continuous_on S f" - and "closedin (subtopology euclidean U) S" + and "closedin (top_of_set U) S" obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" apply (rule Dugundji [of UNIV U S f]) using assms by auto diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Analysis/Convex_Euclidean_Space.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy Tue Mar 19 16:14:51 2019 +0000 @@ -62,7 +62,7 @@ using g unfolding o_def id_def image_def by auto metis+ show ?thesis proof (rule closedin_closed_trans [of "range f"]) - show "closedin (subtopology euclidean (range f)) (f ` S)" + show "closedin (top_of_set (range f)) (f ` S)" using continuous_closedin_preimage [OF confg cgf] by simp show "closed (range f)" apply (rule closed_injective_image_subspace) @@ -319,7 +319,7 @@ subsection \Relative interior of a set\ definition%important "rel_interior S = - {x. \T. openin (subtopology euclidean (affine hull S)) T \ x \ T \ T \ S}" + {x. \T. openin (top_of_set (affine hull S)) T \ x \ T \ T \ S}" lemma rel_interior_mono: "\S \ T; affine hull S = affine hull T\ @@ -327,7 +327,7 @@ by (auto simp: rel_interior_def) lemma rel_interior_maximal: - "\T \ S; openin(subtopology euclidean (affine hull S)) T\ \ T \ (rel_interior S)" + "\T \ S; openin(top_of_set (affine hull S)) T\ \ T \ (rel_interior S)" by (auto simp: rel_interior_def) lemma rel_interior: @@ -651,20 +651,20 @@ definition%important "rel_open S \ rel_interior S = S" -lemma rel_open: "rel_open S \ openin (subtopology euclidean (affine hull S)) S" +lemma rel_open: "rel_open S \ openin (top_of_set (affine hull S)) S" unfolding rel_open_def rel_interior_def apply auto - using openin_subopen[of "subtopology euclidean (affine hull S)" S] + using openin_subopen[of "top_of_set (affine hull S)" S] apply auto done -lemma openin_rel_interior: "openin (subtopology euclidean (affine hull S)) (rel_interior S)" +lemma openin_rel_interior: "openin (top_of_set (affine hull S)) (rel_interior S)" apply (simp add: rel_interior_def) apply (subst openin_subopen, blast) done lemma openin_set_rel_interior: - "openin (subtopology euclidean S) (rel_interior S)" + "openin (top_of_set S) (rel_interior S)" by (rule openin_subset_trans [OF openin_rel_interior rel_interior_subset hull_subset]) lemma affine_rel_open: @@ -806,11 +806,11 @@ qed lemma rel_interior_eq: - "rel_interior s = s \ openin(subtopology euclidean (affine hull s)) s" + "rel_interior s = s \ openin(top_of_set (affine hull s)) s" using rel_open rel_open_def by blast lemma rel_interior_openin: - "openin(subtopology euclidean (affine hull s)) s \ rel_interior s = s" + "openin(top_of_set (affine hull s)) s \ rel_interior s = s" by (simp add: rel_interior_eq) lemma rel_interior_affine: @@ -1992,9 +1992,9 @@ proof - obtain e where "e > 0" and e: "cball a e \ T" using \open T\ \a \ T\ by (auto simp: open_contains_cball) - have "openin (subtopology euclidean S) {a}" + have "openin (top_of_set S) {a}" unfolding openin_open using that \a \ S\ by blast - moreover have "closedin (subtopology euclidean S) {a}" + moreover have "closedin (top_of_set S) {a}" by (simp add: assms) ultimately show "False" using \connected S\ connected_clopen S by blast diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Elementary_Metric_Spaces.thy --- a/src/HOL/Analysis/Elementary_Metric_Spaces.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Elementary_Metric_Spaces.thy Tue Mar 19 16:14:51 2019 +0000 @@ -2718,7 +2718,7 @@ subsection \With Abstract Topology (TODO: move and remove dependency?)\ lemma openin_contains_ball: - "openin (subtopology euclidean t) s \ + "openin (top_of_set t) s \ s \ t \ (\x \ s. \e. 0 < e \ ball x e \ t \ s)" (is "?lhs = ?rhs") proof @@ -2735,7 +2735,7 @@ qed lemma openin_contains_cball: - "openin (subtopology euclidean t) s \ + "openin (top_of_set t) s \ s \ t \ (\x \ s. \e. 0 < e \ cball x e \ t \ s)" apply (simp add: openin_contains_ball) diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Elementary_Normed_Spaces.thy --- a/src/HOL/Analysis/Elementary_Normed_Spaces.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Elementary_Normed_Spaces.thy Tue Mar 19 16:14:51 2019 +0000 @@ -915,7 +915,7 @@ theorem Baire: fixes S::"'a::{real_normed_vector,heine_borel} set" assumes "closed S" "countable \" - and ope: "\T. T \ \ \ openin (subtopology euclidean S) T \ S \ closure T" + and ope: "\T. T \ \ \ openin (top_of_set S) T \ S \ closure T" shows "S \ closure(\\)" proof (cases "\ = {}") case True @@ -930,31 +930,31 @@ proof (clarsimp simp: closure_approachable) fix x and e::real assume "x \ S" "0 < e" - obtain TF where opeF: "\n. openin (subtopology euclidean S) (TF n)" + obtain TF where opeF: "\n. openin (top_of_set S) (TF n)" and ne: "\n. TF n \ {}" and subg: "\n. S \ closure(TF n) \ ?g n" and subball: "\n. closure(TF n) \ ball x e" and decr: "\n. TF(Suc n) \ TF n" proof - - have *: "\Y. (openin (subtopology euclidean S) Y \ Y \ {} \ + have *: "\Y. (openin (top_of_set S) Y \ Y \ {} \ S \ closure Y \ ?g n \ closure Y \ ball x e) \ Y \ U" - if opeU: "openin (subtopology euclidean S) U" and "U \ {}" and cloU: "closure U \ ball x e" for U n + if opeU: "openin (top_of_set S) U" and "U \ {}" and cloU: "closure U \ ball x e" for U n proof - obtain T where T: "open T" "U = T \ S" - using \openin (subtopology euclidean S) U\ by (auto simp: openin_subtopology) + using \openin (top_of_set S) U\ by (auto simp: openin_subtopology) with \U \ {}\ have "T \ closure (?g n) \ {}" using gin ope by fastforce then have "T \ ?g n \ {}" using \open T\ open_Int_closure_eq_empty by blast then obtain y where "y \ U" "y \ ?g n" using T ope [of "?g n", OF gin] by (blast dest: openin_imp_subset) - moreover have "openin (subtopology euclidean S) (U \ ?g n)" + moreover have "openin (top_of_set S) (U \ ?g n)" using gin ope opeU by blast ultimately obtain d where U: "U \ ?g n \ S" and "d > 0" and d: "ball y d \ S \ U \ ?g n" by (force simp: openin_contains_ball) show ?thesis proof (intro exI conjI) - show "openin (subtopology euclidean S) (S \ ball y (d/2))" + show "openin (top_of_set S) (S \ ball y (d/2))" by (simp add: openin_open_Int) show "S \ ball y (d/2) \ {}" using \0 < d\ \y \ U\ opeU openin_imp_subset by fastforce @@ -979,14 +979,14 @@ using ball_divide_subset_numeral d by blast qed qed - let ?\ = "\n X. openin (subtopology euclidean S) X \ X \ {} \ + let ?\ = "\n X. openin (top_of_set S) X \ X \ {} \ S \ closure X \ ?g n \ closure X \ ball x e" have "closure (S \ ball x (e / 2)) \ closure(ball x (e/2))" by (simp add: closure_mono) also have "... \ ball x e" using \e > 0\ by auto finally have "closure (S \ ball x (e / 2)) \ ball x e" . - moreover have"openin (subtopology euclidean S) (S \ ball x (e / 2))" "S \ ball x (e / 2) \ {}" + moreover have"openin (top_of_set S) (S \ ball x (e / 2))" "S \ ball x (e / 2) \ {}" using \0 < e\ \x \ S\ by auto ultimately obtain Y where Y: "?\ 0 Y \ Y \ S \ ball x (e / 2)" using * [of "S \ ball x (e/2)" 0] by metis diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy --- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Tue Mar 19 16:14:51 2019 +0000 @@ -2032,7 +2032,7 @@ subsection\Negligibility is a local property\ lemma locally_negligible_alt: - "negligible S \ (\x \ S. \U. openin (subtopology euclidean S) U \ x \ U \ negligible U)" + "negligible S \ (\x \ S. \U. openin (top_of_set S) U \ x \ U \ negligible U)" (is "_ = ?rhs") proof assume "negligible S" @@ -2040,7 +2040,7 @@ using openin_subtopology_self by blast next assume ?rhs - then obtain U where ope: "\x. x \ S \ openin (subtopology euclidean S) (U x)" + then obtain U where ope: "\x. x \ S \ openin (top_of_set S) (U x)" and cov: "\x. x \ S \ x \ U x" and neg: "\x. x \ S \ negligible (U x)" by metis @@ -4493,7 +4493,7 @@ fix a assume a: "a \ U n" and fa: "f a \ T" define V where "V \ {x. x \ U n \ f x \ T} \ ball a (1 / n / 2)" - show "\V. openin (subtopology euclidean {x \ U n. f x \ T}) V \ a \ V \ negligible V" + show "\V. openin (top_of_set {x \ U n. f x \ T}) V \ a \ V \ negligible V" proof (intro exI conjI) have noxy: "norm(x - y) \ n * norm(f x - f y)" if "x \ V" "y \ V" for x y using that unfolding U_def V_def mem_Collect_eq Int_iff mem_ball dist_norm diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Function_Topology.thy --- a/src/HOL/Analysis/Function_Topology.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Function_Topology.thy Tue Mar 19 16:14:51 2019 +0000 @@ -7,9 +7,9 @@ begin -section \Product Topology\ (* FIX see comments by the author *) +section \Function Topology\ -text \We want to define the product topology. +text \We want to define the general product topology. The product topology on a product of topological spaces is generated by the sets which are products of open sets along finitely many coordinates, and the whole @@ -48,7 +48,7 @@ be the natural continuity definition of a map from the topology \T1\ to the topology \T2\. Then the current \continuous_on\ (with type classes) can be redefined as @{text [display] \continuous_on s f = - continuous_on_topo (subtopology euclidean s) (topology euclidean) f\} + continuous_on_topo (top_of_set s) (topology euclidean) f\} In fact, I need \continuous_on_topo\ to express the continuity of the projection on subfactors for the product topology, in Lemma~\continuous_on_restrict_product_topology\, and I show @@ -61,369 +61,6 @@ I realized afterwards that this theory has a lot in common with \<^file>\~~/src/HOL/Library/Finite_Map.thy\. \ -subsection \Topology without type classes\ - -subsubsection%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%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: - "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]) - -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 - -subsubsection%important \Continuity\ - -text \We will need to deal with continuous maps in terms of topologies and not in terms -of type classes, as defined below.\ - -definition%important continuous_on_topo::"'a topology \ 'b topology \ ('a \ 'b) \ bool" - where "continuous_on_topo T1 T2 f = ((\ U. openin T2 U \ openin T1 (f-`U \ topspace(T1))) - \ (f`(topspace T1) \ (topspace T2)))" - -lemma continuous_on_continuous_on_topo: - "continuous_on s f \ continuous_on_topo (subtopology euclidean s) euclidean f" - by (auto simp: continuous_on_topo_def Int_commute continuous_openin_preimage_eq) - -lemma continuous_on_topo_UNIV: - "continuous_on UNIV f \ continuous_on_topo euclidean euclidean f" -using continuous_on_continuous_on_topo[of UNIV f] subtopology_UNIV[of euclidean] by auto - -lemma continuous_on_topo_open [intro]: - "continuous_on_topo T1 T2 f \ openin T2 U \ openin T1 (f-`U \ topspace(T1))" - unfolding continuous_on_topo_def by auto - -lemma continuous_on_topo_topspace [intro]: - "continuous_on_topo T1 T2 f \ f`(topspace T1) \ (topspace T2)" -unfolding continuous_on_topo_def by auto - -lemma continuous_on_generated_topo_iff: - "continuous_on_topo T1 (topology_generated_by S) f \ - ((\U. U \ S \ openin T1 (f-`U \ topspace(T1))) \ (f`(topspace T1) \ (\ S)))" -unfolding continuous_on_topo_def 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_on_topo T1 (topology_generated_by S) f" - using assms continuous_on_generated_topo_iff by blast - -proposition continuous_on_topo_compose: - assumes "continuous_on_topo T1 T2 f" "continuous_on_topo T2 T3 g" - shows "continuous_on_topo T1 T3 (g o f)" - using assms unfolding continuous_on_topo_def -proof (auto) - fix U :: "'c set" - assume H: "openin T3 U" - have "openin T1 (f -` (g -` U \ topspace T2) \ topspace T1)" - using H assms by blast - moreover have "f -` (g -` U \ topspace T2) \ topspace T1 = (g \ f) -` U \ topspace T1" - using H assms continuous_on_topo_topspace by fastforce - ultimately show "openin T1 ((g \ f) -` U \ topspace T1)" - by simp -qed (blast) - -lemma continuous_on_topo_preimage_topspace [intro]: - assumes "continuous_on_topo T1 T2 f" - shows "f-`(topspace T2) \ topspace T1 = topspace T1" -using assms unfolding continuous_on_topo_def by auto - - -subsubsection%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%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_on_topo_pullback [intro]: - assumes "continuous_on_topo T1 T2 g" - shows "continuous_on_topo (pullback_topology A f T1) T2 (g o f)" -unfolding continuous_on_topo_def -proof (auto) - fix U::"'b set" assume "openin T2 U" - then have "openin T1 (g-`U \ topspace T1)" - using assms unfolding continuous_on_topo_def 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_on_topo_def by auto -qed - -proposition continuous_on_topo_pullback' [intro]: - assumes "continuous_on_topo T1 T2 (f o g)" "topspace T1 \ g-`A" - shows "continuous_on_topo T1 (pullback_topology A f T2) g" -unfolding continuous_on_topo_def -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_on_topo_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 \The product topology\ text \We can now define the product topology, as generated by @@ -605,6 +242,9 @@ then show ?thesis using \x \ U\ by auto qed +lemma product_topology_empty_discrete: + "product_topology T {} = discrete_topology {(\x. undefined)}" + by (simp add: subtopology_eq_discrete_topology_sing) lemma openin_product_topology: "openin (product_topology X I) = @@ -793,6 +433,15 @@ using PiE_singleton closedin_product_topology [of X I] by (metis (no_types, lifting) all_not_in_conv insertI1) +lemma product_topology_empty: + "product_topology X {} = topology (\S. S \ {{},{\k. undefined}})" + unfolding product_topology union_of_def intersection_of_def arbitrary_def relative_to_def + by (auto intro: arg_cong [where f=topology]) + +lemma openin_product_topology_empty: "openin (product_topology X {}) S \ S \ {{},{\k. undefined}}" + unfolding union_of_def intersection_of_def arbitrary_def relative_to_def openin_product_topology + by auto + subsubsection \The basic property of the product topology is the continuity of projections:\ @@ -949,10 +598,10 @@ by (rule continuous_on_topo_product_coordinates, simp) lemma continuous_on_coordinatewise_then_product [intro, continuous_intros]: - fixes f :: "('a \ real) \ 'b \ real" + fixes f :: "'a::topological_space \ 'b \ 'c::topological_space" assumes "\i. continuous_on S (\x. f x i)" shows "continuous_on S f" - using continuous_on_topo_coordinatewise_then_product [of UNIV, where T = "\i. euclideanreal"] + using continuous_on_topo_coordinatewise_then_product [of UNIV, where T = "\i. euclidean"] by (metis UNIV_I assms continuous_on_continuous_on_topo euclidean_product_topology) lemma continuous_on_product_then_coordinatewise_UNIV: @@ -1583,4 +1232,354 @@ instance "fun" :: (countable, polish_space) polish_space by standard + +subsection\The Alexander subbase theorem\ + +theorem Alexander_subbase: + assumes X: "topology (arbitrary union_of (finite intersection_of (\x. x \ \) relative_to \\)) = X" + and fin: "\C. \C \ \; \C = topspace X\ \ \C'. finite C' \ C' \ C \ \C' = topspace X" + shows "compact_space X" +proof - + have U\: "\\ = topspace X" + by (simp flip: X) + have False if \: "\U\\. openin X U" and sub: "topspace X \ \\" + and neg: "\\. \\ \ \; finite \\ \ \ topspace X \ \\" for \ + proof - + define \ where "\ \ {\. (\U \ \. openin X U) \ topspace X \ \\ \ (\\. finite \ \ \ \ \ \ ~(topspace X \ \\))}" + have 1: "\ \ {}" + unfolding \_def using sub \ neg by force + have 2: "\\ \ \" if "\\{}" and \: "subset.chain \ \" for \ + unfolding \_def + proof (intro CollectI conjI ballI allI impI notI) + show "openin X U" if U: "U \ \\" for U + using U \ unfolding \_def subset_chain_def by force + have "\ \ \" + using subset_chain_def \ by blast + with that \_def show UUC: "topspace X \ \(\\)" + by blast + show "False" if "finite \" and "\ \ \\" and "topspace X \ \\" for \ + proof - + obtain \ where "\ \ \" "\ \ \" + by (metis Sup_empty \ \\ \ \\\ \finite \\ UUC empty_subsetI finite.emptyI finite_subset_Union_chain neg) + then show False + using \_def \\ \ \\ \finite \\ \topspace X \ \\\ by blast + qed + qed + obtain \ where "\ \ \" and "\X. \X\\; \ \ X\ \ X = \" + using subset_Zorn_nonempty [OF 1 2] by metis + then have *: "\\. \\W. W\\ \ openin X W; topspace X \ \\; \ \ \; + \\. \finite \; \ \ \; topspace X \ \\\ \ False\ + \ \ = \" + and ope: "\U\\. openin X U" and top: "topspace X \ \\" + and non: "\\. \finite \; \ \ \; topspace X \ \\\ \ False" + unfolding \_def by simp_all metis+ + then obtain x where "x \ topspace X" "x \ \(\ \ \)" + proof - + have "\(\ \ \) \ \\" + by (metis \\\ = topspace X\ fin inf.bounded_iff non order_refl) + then have "\a. a \ \(\ \ \) \ a \ \\" + by blast + then show ?thesis + using that by (metis U\) + qed + obtain C where C: "openin X C" "C \ \" "x \ C" + using \x \ topspace X\ ope top by auto + then have "C \ topspace X" + by (metis openin_subset) + then have "(arbitrary union_of (finite intersection_of (\x. x \ \) relative_to \\)) C" + using openin_subbase C unfolding X [symmetric] by blast + moreover have "C \ topspace X" + using \\ \ \\ \C \ \\ unfolding \_def by blast + ultimately obtain \ W where W: "(finite intersection_of (\x. x \ \) relative_to topspace X) W" + and "x \ W" "W \ \" "\\ \ topspace X" "C = \\" + using C by (auto simp: union_of_def U\) + then have "\\ \ topspace X" + by (metis \C \ topspace X\) + then have "topspace X \ \" + using \\\ \ topspace X\ by blast + then obtain \' where \': "finite \'" "\' \ \" "x \ \\'" "W = topspace X \ \\'" + using W \x \ W\ unfolding relative_to_def intersection_of_def by auto + then have "\\' \ \\" + using \W \ \\ \\\ \ topspace X\ \\\ \ topspace X\ by blast + then have "\\' \ C" + using U\ \C = \\\ \W = topspace X \ \\'\ \W \ \\ by auto + have "\b \ \'. \C'. finite C' \ C' \ \ \ topspace X \ \(insert b C')" + proof + fix b + assume "b \ \'" + have "insert b \ = \" if neg: "\ (\C'. finite C' \ C' \ \ \ topspace X \ \(insert b C'))" + proof (rule *) + show "openin X W" if "W \ insert b \" for W + using that + proof + have "b \ \" + using \b \ \'\ \\' \ \\ by blast + then have "\\. finite \ \ \ \ \ \ \\ = b" + by (rule_tac x="{b}" in exI) auto + moreover have "\\ \ b = b" + using \'(2) \b \ \'\ by auto + ultimately show "openin X W" if "W = b" + using that \b \ \'\ + apply (simp add: openin_subbase flip: X) + apply (auto simp: arbitrary_def intersection_of_def relative_to_def intro!: union_of_inc) + done + show "openin X W" if "W \ \" + by (simp add: \W \ \\ ope) + qed + next + show "topspace X \ \ (insert b \)" + using top by auto + next + show False if "finite \" and "\ \ insert b \" "topspace X \ \\" for \ + proof - + have "insert b (\ \ \) = \" + using non that by blast + then show False + by (metis Int_lower2 finite_insert neg that(1) that(3)) + qed + qed auto + then show "\C'. finite C' \ C' \ \ \ topspace X \ \(insert b C')" + using \b \ \'\ \x \ \(\ \ \)\ \' + by (metis IntI InterE Union_iff subsetD insertI1) + qed + then obtain F where F: "\b \ \'. finite (F b) \ F b \ \ \ topspace X \ \(insert b (F b))" + by metis + let ?\ = "insert C (\(F ` \'))" + show False + proof (rule non) + have "topspace X \ (\b \ \'. \(insert b (F b)))" + using F by (simp add: INT_greatest) + also have "\ \ \?\" + using \\\' \ C\ by force + finally show "topspace X \ \?\" . + show "?\ \ \" + using \C \ \\ F by auto + show "finite ?\" + using \finite \'\ F by auto + qed + qed + then show ?thesis + by (force simp: compact_space_def compactin_def) +qed + + +corollary Alexander_subbase_alt: + assumes "U \ \\" + and fin: "\C. \C \ \; U \ \C\ \ \C'. finite C' \ C' \ C \ U \ \C'" + and X: "topology + (arbitrary union_of + (finite intersection_of (\x. x \ \) relative_to U)) = X" + shows "compact_space X" +proof - + have "topspace X = U" + using X topspace_subbase by fastforce + have eq: "\ (Collect ((\x. x \ \) relative_to U)) = U" + unfolding relative_to_def + using \U \ \\\ by blast + have *: "\\. finite \ \ \ \ \ \ \\ = topspace X" + if "\ \ Collect ((\x. x \ \) relative_to topspace X)" and UC: "\\ = topspace X" for \ + proof - + have "\ \ (\U. topspace X \ U) ` \" + using that by (auto simp: relative_to_def) + then obtain \' where "\' \ \" and \': "\ = (\) (topspace X) ` \'" + by (auto simp: subset_image_iff) + moreover have "U \ \\'" + using \' \topspace X = U\ UC by auto + ultimately obtain \' where "finite \'" "\' \ \'" "U \ \\'" + using fin [of \'] \topspace X = U\ \U \ \\\ by blast + then show ?thesis + unfolding \' exists_finite_subset_image \topspace X = U\ by auto + qed + show ?thesis + apply (rule Alexander_subbase [where \ = "Collect ((\x. x \ \) relative_to (topspace X))"]) + apply (simp flip: X) + apply (metis finite_intersection_of_relative_to eq) + apply (blast intro: *) + done +qed + +proposition continuous_map_componentwise: + "continuous_map X (product_topology Y I) f \ + f ` (topspace X) \ extensional I \ (\k \ I. continuous_map X (Y k) (\x. f x k))" + (is "?lhs \ _ \ ?rhs") +proof (cases "\x \ topspace X. f x \ extensional I") + case True + then have "f ` (topspace X) \ extensional I" + by force + moreover have ?rhs if L: ?lhs + proof - + have "openin X {x \ topspace X. f x k \ U}" if "k \ I" and "openin (Y k) U" for k U + proof - + have "openin (product_topology Y I) ({Y. Y k \ U} \ (\\<^sub>E i\I. topspace (Y i)))" + apply (simp add: openin_product_topology flip: arbitrary_union_of_relative_to) + apply (simp add: relative_to_def) + using that apply (blast intro: arbitrary_union_of_inc finite_intersection_of_inc) + done + with that have "openin X {x \ topspace X. f x \ ({Y. Y k \ U} \ (\\<^sub>E i\I. topspace (Y i)))}" + using L unfolding continuous_map_def by blast + moreover have "{x \ topspace X. f x \ ({Y. Y k \ U} \ (\\<^sub>E i\I. topspace (Y i)))} = {x \ topspace X. f x k \ U}" + using L by (auto simp: continuous_map_def) + ultimately show ?thesis + by metis + qed + with that + show ?thesis + by (auto simp: continuous_map_def) + qed + moreover have ?lhs if ?rhs + proof - + have 1: "\x. x \ topspace X \ f x \ (\\<^sub>E i\I. topspace (Y i))" + using that True by (auto simp: continuous_map_def PiE_iff) + have 2: "{x \ S. \T\\. f x \ T} = (\T\\. {x \ S. f x \ T})" for S \ + by blast + have 3: "{x \ S. \U\\. f x \ U} = (\ (insert S ((\U. {x \ S. f x \ U}) ` \)))" for S \ + by blast + show ?thesis + unfolding continuous_map_def openin_product_topology arbitrary_def + proof (clarsimp simp: all_union_of 1 2) + fix \ + assume \: "\ \ Collect (finite intersection_of (\F. \i U. F = {f. f i \ U} \ i \ I \ openin (Y i) U) + relative_to (\\<^sub>E i\I. topspace (Y i)))" + show "openin X (\T\\. {x \ topspace X. f x \ T})" + proof (rule openin_Union; clarify) + fix S T + assume "T \ \" + obtain \ where "T = (\\<^sub>E i\I. topspace (Y i)) \ \\" and "finite \" + "\ \ {{f. f i \ U} |i U. i \ I \ openin (Y i) U}" + using subsetD [OF \ \T \ \\] by (auto simp: intersection_of_def relative_to_def) + with that show "openin X {x \ topspace X. f x \ T}" + apply (simp add: continuous_map_def 1 cong: conj_cong) + unfolding 3 + apply (rule openin_Inter; auto) + done + qed + qed + qed + ultimately show ?thesis + by metis +next + case False + then show ?thesis + by (auto simp: continuous_map_def PiE_def) +qed + + +lemma continuous_map_componentwise_UNIV: + "continuous_map X (product_topology Y UNIV) f \ (\k. continuous_map X (Y k) (\x. f x k))" + by (simp add: continuous_map_componentwise) + +lemma continuous_map_product_projection [continuous_intros]: + "k \ I \ continuous_map (product_topology X I) (X k) (\x. x k)" + using continuous_map_componentwise [of "product_topology X I" X I id] by simp + +proposition open_map_product_projection: + assumes "i \ I" + shows "open_map (product_topology Y I) (Y i) (\f. f i)" + unfolding openin_product_topology all_union_of arbitrary_def open_map_def image_Union +proof clarify + fix \ + assume \: "\ \ Collect + (finite intersection_of + (\F. \i U. F = {f. f i \ U} \ i \ I \ openin (Y i) U) relative_to + topspace (product_topology Y I))" + show "openin (Y i) (\x\\. (\f. f i) ` x)" + proof (rule openin_Union, clarify) + fix S V + assume "V \ \" + obtain \ where "finite \" + and V: "V = (\\<^sub>E i\I. topspace (Y i)) \ \\" + and \: "\ \ {{f. f i \ U} |i U. i \ I \ openin (Y i) U}" + using subsetD [OF \ \V \ \\] + by (auto simp: intersection_of_def relative_to_def) + show "openin (Y i) ((\f. f i) ` V)" + proof (subst openin_subopen; clarify) + fix x f + assume "f \ V" + let ?T = "{a \ topspace(Y i). + (\j. if j = i then a + else if j \ I then f j else undefined) \ (\\<^sub>E i\I. topspace (Y i)) \ \\}" + show "\T. openin (Y i) T \ f i \ T \ T \ (\f. f i) ` V" + proof (intro exI conjI) + show "openin (Y i) ?T" + proof (rule openin_continuous_map_preimage) + have "continuous_map (Y i) (Y k) (\x. if k = i then x else f k)" if "k \ I" for k + proof (cases "k=i") + case True + then show ?thesis + by (metis (mono_tags) continuous_map_id eq_id_iff) + next + case False + then show ?thesis + by simp (metis IntD1 PiE_iff V \f \ V\ that) + qed + then show "continuous_map (Y i) (product_topology Y I) + (\x j. if j = i then x else if j \ I then f j else undefined)" + by (auto simp: continuous_map_componentwise assms extensional_def) + next + have "openin (product_topology Y I) (\\<^sub>E i\I. topspace (Y i))" + by (metis openin_topspace topspace_product_topology) + moreover have "openin (product_topology Y I) (\B\\. (\\<^sub>E i\I. topspace (Y i)) \ B)" + if "\ \ {}" + proof - + show ?thesis + proof (rule openin_Inter) + show "\X. X \ (\) (\\<^sub>E i\I. topspace (Y i)) ` \ \ openin (product_topology Y I) X" + unfolding openin_product_topology relative_to_def + apply (clarify intro!: arbitrary_union_of_inc) + apply (rename_tac F) + apply (rule_tac x=F in exI) + using subsetD [OF \] + apply (force intro: finite_intersection_of_inc) + done + qed (use \finite \\ \\ \ {}\ in auto) + qed + ultimately show "openin (product_topology Y I) ((\\<^sub>E i\I. topspace (Y i)) \ \\)" + by (auto simp only: Int_Inter_eq split: if_split) + qed + next + have eqf: "(\j. if j = i then f i else if j \ I then f j else undefined) = f" + using PiE_arb V \f \ V\ by force + show "f i \ ?T" + using V assms \f \ V\ by (auto simp: PiE_iff eqf) + next + show "?T \ (\f. f i) ` V" + unfolding V by (auto simp: intro!: rev_image_eqI) + qed + qed + qed +qed + +lemma retraction_map_product_projection: + assumes "i \ I" + shows "(retraction_map (product_topology X I) (X i) (\x. x i) \ + (topspace (product_topology X I) = {}) \ topspace (X i) = {})" + (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + using retraction_imp_surjective_map by force +next + assume R: ?rhs + show ?lhs + proof (cases "topspace (product_topology X I) = {}") + case True + then show ?thesis + using R by (auto simp: retraction_map_def retraction_maps_def continuous_map_on_empty) + next + case False + have *: "\g. continuous_map (X i) (product_topology X I) g \ (\x\topspace (X i). g x i = x)" + if z: "z \ (\\<^sub>E i\I. topspace (X i))" for z + proof - + have cm: "continuous_map (X i) (X j) (\x. if j = i then x else z j)" if "j \ I" for j + using \j \ I\ z by (case_tac "j = i") auto + show ?thesis + using \i \ I\ that + by (rule_tac x="\x j. if j = i then x else z j" in exI) (auto simp: continuous_map_componentwise PiE_iff extensional_def cm) + qed + show ?thesis + using \i \ I\ False + by (auto simp: retraction_map_def retraction_maps_def assms continuous_map_product_projection *) + qed +qed + end diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Further_Topology.thy --- a/src/HOL/Analysis/Further_Topology.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Further_Topology.thy Tue Mar 19 16:14:51 2019 +0000 @@ -1149,7 +1149,7 @@ 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 (subtopology euclidean U) S" and K: "disjnt K S" "K \ U" + 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 @@ -1176,11 +1176,11 @@ 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 (subtopology euclidean U) C" + have opeUC: "openin (top_of_set U) C" proof (rule openin_trans) - show "openin (subtopology euclidean (U-S)) C" + 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 (subtopology euclidean U) (U - S)" + 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" @@ -1192,10 +1192,10 @@ 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 (subtopology euclidean C) (ball a d \ U)" + 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 (subtopology euclidean (affine hull C)) C" - by (metis \a \ C\ \openin (subtopology euclidean U) C\ affine_hull_eq affine_hull_openin all_not_in_conv \affine U\) + 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)" @@ -1323,13 +1323,14 @@ 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\, of "S \ UF" "\C. S \ (C - {a C})" h]) - show "S \ UF \ (\C\F. S \ (C - {a C}))" + 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 (subtopology euclidean (S \ UF)) (S \ (C - {a C}))" + 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 (subtopology euclidean 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 @@ -1346,22 +1347,26 @@ show "S \ (C - {a C}) = (S \ C) \ (S \ UF)" using F_def UF_def components_nonoverlap that by auto qed - next - show "continuous_on (S \ (C' - {a C'})) (h C')" if "C' \ F" for C' - using ah F_def that by blast + 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 \ (S \ UF) \ (S \ (i - {a i})) \ (S \ (j - {a j}))\ + 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 blast + 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 (subtopology euclidean (S \ \G \ (S \ UF))) (S \ \G)" + have clo1: "closedin (top_of_set (S \ \G \ (S \ UF))) (S \ \G)" proof (rule closedin_closed_subset [OF _ SU']) - have *: "\C. C \ F \ openin (subtopology euclidean U) C" + 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 (subtopology euclidean U) (U - UF)" + 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))" @@ -1370,9 +1375,9 @@ apply (metis DiffD1 UnionI Union_components) by (metis (no_types, lifting) IntI components_nonoverlap empty_iff) qed - have clo2: "closedin (subtopology euclidean (S \ \G \ (S \ UF))) (S \ UF)" + have clo2: "closedin (top_of_set (S \ \G \ (S \ UF))) (S \ UF)" proof (rule closedin_closed_subset [OF _ SU']) - show "closedin (subtopology euclidean U) (\C\F. S \ C)" + 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 @@ -1467,7 +1472,7 @@ 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 (subtopology euclidean T) S" + 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 @@ -2008,8 +2013,8 @@ 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 (subtopology euclidean S) U" - shows "openin (subtopology euclidean S) (f ` U)" + 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 @@ -2033,7 +2038,7 @@ 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 (subtopology euclidean (k ` S)) T" + 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]) @@ -2053,14 +2058,14 @@ 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 (subtopology euclidean S) U" - shows "openin (subtopology euclidean S) (f ` U)" + 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 (subtopology euclidean (S \ S')) (g ` (U \ S'))" + 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) @@ -2072,7 +2077,7 @@ 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 (subtopology euclidean (S \ S')) (U \ S')" + 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) @@ -2092,11 +2097,11 @@ corollary invariance_of_domain_subspaces: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes ope: "openin (subtopology euclidean U) S" + 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 (subtopology euclidean V) (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] @@ -2119,7 +2124,7 @@ by (simp add: homeomorphism_symD homhk) have hfV': "(h \ f) ` S \ V'" using fim homeomorphism_image1 homhk by fastforce - moreover have "openin (subtopology euclidean U) ((h \ f) ` S)" + 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) @@ -2129,14 +2134,14 @@ show "(h \ f) ` S \ U" using \V' \ U\ hfV' by auto qed (auto simp: assms) - ultimately show "openin (subtopology euclidean V') ((h \ f) ` S)" + 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 (subtopology euclidean U) S" + 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 \ {}" @@ -2158,7 +2163,7 @@ 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 (subtopology euclidean U) ((h \ f) ` S)" + 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 @@ -2176,11 +2181,11 @@ corollary invariance_of_domain_affine_sets: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes ope: "openin (subtopology euclidean U) S" + 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 (subtopology euclidean V) (f ` S)" + shows "openin (top_of_set V) (f ` S)" proof (cases "S = {}") case True then show ?thesis by auto @@ -2188,9 +2193,9 @@ case False obtain a b where "a \ S" "a \ U" "b \ V" using False fim ope openin_contains_cball by fastforce - have "openin (subtopology euclidean ((+) (- b) ` V)) (((+) (- b) \ f \ (+) a) ` (+) (- a) ` S)" + have "openin (top_of_set ((+) (- b) ` V)) (((+) (- b) \ f \ (+) a) ` (+) (- a) ` S)" proof (rule invariance_of_domain_subspaces) - show "openin (subtopology euclidean ((+) (- a) ` U)) ((+) (- a) ` S)" + 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) @@ -2211,7 +2216,7 @@ corollary invariance_of_dimension_affine_sets: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes ope: "openin (subtopology euclidean U) S" + 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 \ {}" @@ -2221,7 +2226,7 @@ using \S \ {}\ fim ope openin_contains_cball by fastforce have "dim ((+) (- a) ` U) \ dim ((+) (- b) ` V)" proof (rule invariance_of_dimension_subspaces) - show "openin (subtopology euclidean ((+) (- a) ` U)) ((+) (- a) ` S)" + 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) @@ -2269,7 +2274,7 @@ case False have "aff_dim (affine hull S) \ aff_dim (affine hull T)" proof (rule invariance_of_dimension_affine_sets) - show "openin (subtopology euclidean (affine hull S)) (rel_interior S)" + 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 @@ -2335,7 +2340,7 @@ assume "open T" have eq: "f ` S \ g -` T = f ` (S \ T)" by (auto simp: gf) - show "openin (subtopology euclidean (f ` S)) (f ` S \ g -` T)" + 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) @@ -2530,9 +2535,9 @@ proof (rule rel_interior_maximal) show "f ` rel_interior S \ f ` S" by(simp add: image_mono rel_interior_subset) - show "openin (subtopology euclidean (affine hull f ` S)) (f ` rel_interior S)" + show "openin (top_of_set (affine hull f ` S)) (f ` rel_interior S)" proof (rule invariance_of_domain_affine_sets) - show "openin (subtopology euclidean (affine hull S)) (rel_interior S)" + 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) @@ -2826,8 +2831,8 @@ 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 (subtopology euclidean (rel_frontier U)) S" - shows "openin (subtopology euclidean T) (f ` S)" + 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 @@ -2857,9 +2862,9 @@ by (simp_all add: homeomorphism_def subset_eq) have [simp]: "aff_dim T \ aff_dim V" by (simp add: affTU affV) - have "openin (subtopology euclidean T) ((f \ h) ` g ` (S - {b}))" + have "openin (top_of_set T) ((f \ h) ` g ` (S - {b}))" proof (rule invariance_of_domain_affine_sets [OF _ \affine V\]) - show "openin (subtopology euclidean V) (g ` (S - {b}))" + 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)" @@ -2874,9 +2879,9 @@ by (metis fim image_comp image_mono hgsub subset_trans) qed (auto simp: assms) moreover - have "openin (subtopology euclidean T) ((f \ k) ` j ` (S - {c}))" + have "openin (top_of_set T) ((f \ k) ` j ` (S - {c}))" proof (rule invariance_of_domain_affine_sets [OF _ \affine V\]) - show "openin (subtopology euclidean V) (j ` (S - {c}))" + 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)" @@ -2890,7 +2895,7 @@ show "(f \ k) ` j ` (S - {c}) \ T" by (metis fim image_comp image_mono kjsub subset_trans) qed (auto simp: assms) - ultimately have "openin (subtopology euclidean T) ((f \ h) ` g ` (S - {b}) \ ((f \ k) ` j ` (S - {c})))" + 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 - @@ -2931,8 +2936,8 @@ 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 (subtopology euclidean (sphere a r)) S" - shows "openin (subtopology euclidean T) (f ` S)" + 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 @@ -2943,7 +2948,7 @@ 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 (subtopology euclidean (rel_frontier (cball a r))) S" + show "openin (top_of_set (rel_frontier (cball a r))) S" by (simp add: \r \ 0\ ope) qed qed @@ -3379,7 +3384,7 @@ 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 (subtopology euclidean (- {0})) T \ + 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 @@ -3399,7 +3404,7 @@ 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 (subtopology euclidean (- {0})) (exp ` ball (Ln z) 1)" + 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) @@ -3647,150 +3652,150 @@ text\Many similar proofs below.\ lemma upper_hemicontinuous: assumes "\x. x \ S \ f x \ T" - shows "((\U. openin (subtopology euclidean T) U - \ openin (subtopology euclidean S) {x \ S. f x \ U}) \ - (\U. closedin (subtopology euclidean T) U - \ closedin (subtopology euclidean S) {x \ S. f x \ U \ {}}))" + 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 (subtopology euclidean T) U" - then have "openin (subtopology euclidean T) (T - 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 (subtopology euclidean S) {x \ S. f x \ T - U}" + 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 (subtopology euclidean S) {x \ S. f x \ U \ {}}" + 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 (subtopology euclidean T) U" - then have "closedin (subtopology euclidean T) (T - 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 (subtopology euclidean S) {x \ S. f x \ (T - U) \ {}}" + 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 (subtopology euclidean S) {x \ S. f x \ U}" + 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 (subtopology euclidean T) U - \ closedin (subtopology euclidean S) {x \ S. f x \ U}) \ - (\U. openin (subtopology euclidean T) U - \ openin (subtopology euclidean S) {x \ S. f x \ U \ {}}))" + 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 (subtopology euclidean T) U" - then have "closedin (subtopology euclidean T) (T - 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 (subtopology euclidean S) {x \ S. f x \ T-U}" + 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 (subtopology euclidean S) {x \ S. f x \ U \ {}}" + 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 (subtopology euclidean T) U" - then have "openin (subtopology euclidean T) (T - 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 (subtopology euclidean S) {x \ S. f x \ (T - U) \ {}}" + 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 (subtopology euclidean S) {x \ S. f x \ U}" + 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 (subtopology euclidean S) U - \ openin (subtopology euclidean T) (f ` U)) \ - (\U. closedin (subtopology euclidean S) U - \ closedin (subtopology euclidean T) {y \ T. {x. x \ S \ f x = y} \ U}))" + 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 (subtopology euclidean S) U" - then have "openin (subtopology euclidean S) (S - 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 (subtopology euclidean T) (f ` (S - U))" + 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 (subtopology euclidean T) {y \ T. {x \ S. f x = y} \ U}" + 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 (subtopology euclidean S) U" - then have "closedin (subtopology euclidean S) (S - 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 (subtopology euclidean T) {y \ T. {x \ S. f x = y} \ S - U}" + 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 (subtopology euclidean T) (f ` U)" + 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 (subtopology euclidean S) U - \ closedin (subtopology euclidean T) (f ` U)) \ - (\U. openin (subtopology euclidean S) U - \ openin (subtopology euclidean T) {y \ T. {x. x \ S \ f x = y} \ U}))" + 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 (subtopology euclidean S) U" - then have "closedin (subtopology euclidean S) (S - 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 (subtopology euclidean T) (f ` (S - U))" + 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 (subtopology euclidean T) {y \ T. {x. x \ S \ f x = y} \ U}" + 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 (subtopology euclidean S) U" - then have "openin (subtopology euclidean S) (S - 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 (subtopology euclidean T) {y \ T. {x \ S. f x = y} \ S - U}" + 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 (subtopology euclidean T) (f ` U)" + 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 (subtopology euclidean T) U - \ openin (subtopology euclidean S) {x \ S. f x \ U}" - and clo: "\U. closedin (subtopology euclidean T) U - \ closedin (subtopology euclidean S) {x \ S. f x \ U}" + 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 (subtopology euclidean T) (T \ (\a\f x. \b\ball 0 e. {a + b}))" + 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 (subtopology euclidean S) + 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 (subtopology euclidean T) U \ - openin (subtopology euclidean S) {x \ S. f x \ U \ {}}" + 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))" @@ -3806,9 +3811,9 @@ 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 (subtopology euclidean S) {x \ S. f x \ (T \ ball a (e/2)) \ {}}" for a + 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 (subtopology euclidean S) (\a \ C. {x \ S. f x \ T \ ball a (e/2) \ {}})" + 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) \ {}" @@ -4375,8 +4380,8 @@ proposition Borsukian_open_Un: fixes S :: "'a::real_normed_vector set" - assumes opeS: "openin (subtopology euclidean (S \ T)) S" - and opeT: "openin (subtopology euclidean (S \ T)) T" + 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) @@ -4445,8 +4450,8 @@ 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 (subtopology euclidean (S \ T)) S" - and cloT: "closedin (subtopology euclidean (S \ T)) T" + 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) @@ -4586,8 +4591,8 @@ 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 (subtopology euclidean S) U - \ openin (subtopology euclidean T) (f ` U)" + 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" @@ -4636,12 +4641,12 @@ \ (\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 (subtopology euclidean S) U - \ openin (subtopology euclidean T) {x \ T. {z \ S. f z = x} \ U}" + show "\U. openin (top_of_set S) U + \ openin (top_of_set T) {x \ T. {z \ S. f z = x} \ U}" using continuous_imp_closed_map closed_map_iff_upper_hemicontinuous_preimage [OF fim [THEN equalityD1]] by (simp add: continuous_imp_closed_map \compact S\ contf fim) - show "\U. closedin (subtopology euclidean S) U \ - closedin (subtopology euclidean T) {x \ T. {z \ S. f z = x} \ U}" + 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}" @@ -4812,17 +4817,17 @@ definition%important unicoherent where "unicoherent U \ \S T. connected S \ connected T \ S \ T = U \ - closedin (subtopology euclidean U) S \ closedin (subtopology euclidean U) T + 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 (subtopology euclidean U) S; closedin (subtopology euclidean U) T\ + 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 (subtopology euclidean U) S" "closedin (subtopology euclidean U) T" + 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 @@ -4837,8 +4842,8 @@ proof fix U V assume "connected U" "connected V" and T: "T = U \ V" - and cloU: "closedin (subtopology euclidean T) U" - and cloV: "closedin (subtopology euclidean T) 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)" @@ -4858,10 +4863,10 @@ 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 (subtopology euclidean T) U" "closedin (subtopology euclidean T) V" + have "closedin (top_of_set T) U" "closedin (top_of_set T) V" by (simp_all add: cloU cloV) - then show "closedin (subtopology euclidean S) (g ` U)" - "closedin (subtopology euclidean S) (g ` V)" + 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 @@ -4894,16 +4899,16 @@ proof clarify fix S T assume "connected S" "connected T" "U = S \ T" - and cloS: "closedin (subtopology euclidean (S \ T)) S" - and cloT: "closedin (subtopology euclidean (S \ T)) 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 (subtopology euclidean (S \ T)) V" - and "closedin (subtopology euclidean (S \ T)) 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 (subtopology euclidean U) V" and cloW: "closedin (subtopology euclidean U) 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}" @@ -5020,8 +5025,8 @@ proof fix U V assume UV: "connected U" "connected V" "T = U \ V" - and cloU: "closedin (subtopology euclidean T) U" - and cloV: "closedin (subtopology euclidean T) 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" @@ -5035,14 +5040,14 @@ 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 (subtopology euclidean S) C \ closedin (subtopology euclidean T) (f ` C)" + 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 (subtopology euclidean S) (S \ f -` U)" - "closedin (subtopology euclidean S) (S \ f -` V)" + 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 @@ -5348,7 +5353,7 @@ 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 (subtopology euclidean UNIV) S" + 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 @@ -5396,8 +5401,8 @@ lemma inessential_on_clopen_Union: fixes \ :: "'a::euclidean_space set set" assumes T: "path_connected T" - and "\S. S \ \ \ closedin (subtopology euclidean (\\)) S" - and "\S. S \ \ \ openin (subtopology euclidean (\\)) S" + and "\S. S \ \ \ closedin (top_of_set (\\)) S" + and "\S. S \ \ \ openin (top_of_set (\\)) S" and hom: "\S. S \ \ \ \a. homotopic_with (\x. True) S T f (\x. a)" obtains a where "homotopic_with (\x. True) (\\) T f (\x. a)" proof (cases "\\ = {}") @@ -5408,16 +5413,16 @@ case False then obtain C where "C \ \" "C \ {}" by blast - then obtain a where clo: "closedin (subtopology euclidean (\\)) C" - and ope: "openin (subtopology euclidean (\\)) C" + then obtain a where clo: "closedin (top_of_set (\\)) C" + and ope: "openin (top_of_set (\\)) C" and "homotopic_with (\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 (\x. True) (\\) T f (\x. a)" proof (rule homotopic_on_clopen_Union) - show "\S. S \ \ \ closedin (subtopology euclidean (\\)) S" - "\S. S \ \ \ openin (subtopology euclidean (\\)) S" + show "\S. S \ \ \ closedin (top_of_set (\\)) S" + "\S. S \ \ \ openin (top_of_set (\\)) S" by (simp_all add: assms) show "homotopic_with (\x. True) S T f (\x. a)" if "S \ \" for S proof (cases "S = {}") diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Henstock_Kurzweil_Integration.thy --- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Tue Mar 19 16:14:51 2019 +0000 @@ -4415,9 +4415,9 @@ then show "\e>0. ball x e \ (S \ f -` {f x})" using \0 < e\ by blast qed - then have "openin (subtopology euclidean S) (S \ f -` {y})" + then have "openin (top_of_set S) (S \ f -` {y})" by (auto intro!: open_openin_trans[OF \open S\] simp: open_contains_ball) - moreover have "closedin (subtopology euclidean S) (S \ f -` {y})" + moreover have "closedin (top_of_set S) (S \ f -` {y})" by (force intro!: continuous_closedin_preimage [OF contf]) ultimately have "(S \ f -` {y}) = {} \ (S \ f -` {y}) = S" using \connected S\ by (simp add: connected_clopen) diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Homeomorphism.thy --- a/src/HOL/Analysis/Homeomorphism.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Homeomorphism.thy Tue Mar 19 16:14:51 2019 +0000 @@ -916,7 +916,7 @@ fixes S :: "'m::euclidean_space set" assumes "aff_dim S < DIM('n)" obtains U and T :: "'n::euclidean_space set" - where "convex U" "U \ {}" "closedin (subtopology euclidean U) T" + where "convex U" "U \ {}" "closedin (top_of_set U) T" "S homeomorphic T" proof (cases "S = {}") case True then show ?thesis @@ -1051,7 +1051,7 @@ lemma locally_compact_closedin_open: fixes S :: "'a :: metric_space set" assumes "locally compact S" - obtains T where "open T" "closedin (subtopology euclidean T) S" + obtains T where "open T" "closedin (top_of_set T) S" by (metis locally_compact_open_Int_closure [OF assms] closed_closure closedin_closed_Int) @@ -1079,7 +1079,7 @@ by (simp add: Ucomp setdist_eq_0_sing_1) then have homU: "homeomorphism U (f`U) f fst" by (auto simp: f_def homeomorphism_def image_iff continuous_intros) - have cloS: "closedin (subtopology euclidean U) S" + have cloS: "closedin (top_of_set U) S" by (metis US closed_closure closedin_closed_Int) have cont: "isCont ((\x. setdist {x} (- U)) o fst) z" for z :: "'a \ 'b" by (rule continuous_at_compose continuous_intros continuous_at_setdist)+ @@ -1275,9 +1275,9 @@ where "covering_space c p S \ continuous_on c p \ p ` c = S \ - (\x \ S. \T. x \ T \ openin (subtopology euclidean S) T \ + (\x \ S. \T. x \ T \ openin (top_of_set S) T \ (\v. \v = c \ p -` T \ - (\u \ v. openin (subtopology euclidean c) u) \ + (\u \ v. openin (top_of_set c) u) \ pairwise disjnt v \ (\u \ v. \q. homeomorphism u T p q)))" @@ -1295,8 +1295,8 @@ lemma covering_space_local_homeomorphism: assumes "covering_space c p S" "x \ c" - obtains T u q where "x \ T" "openin (subtopology euclidean c) T" - "p x \ u" "openin (subtopology euclidean S) u" + obtains T u q where "x \ T" "openin (top_of_set c) T" + "p x \ u" "openin (top_of_set S) u" "homeomorphism T u p q" using assms apply (simp add: covering_space_def, clarify) @@ -1307,8 +1307,8 @@ lemma covering_space_local_homeomorphism_alt: assumes p: "covering_space c p S" and "y \ S" obtains x T U q where "p x = y" - "x \ T" "openin (subtopology euclidean c) T" - "y \ U" "openin (subtopology euclidean S) U" + "x \ T" "openin (top_of_set c) T" + "y \ U" "openin (top_of_set S) U" "homeomorphism T U p q" proof - obtain x where "p x = y" "x \ c" @@ -1320,26 +1320,26 @@ proposition covering_space_open_map: fixes S :: "'a :: metric_space set" and T :: "'b :: metric_space set" - assumes p: "covering_space c p S" and T: "openin (subtopology euclidean c) T" - shows "openin (subtopology euclidean S) (p ` T)" + assumes p: "covering_space c p S" and T: "openin (top_of_set c) T" + shows "openin (top_of_set S) (p ` T)" proof - have pce: "p ` c = S" and covs: "\x. x \ S \ - \X VS. x \ X \ openin (subtopology euclidean S) X \ + \X VS. x \ X \ openin (top_of_set S) X \ \VS = c \ p -` X \ - (\u \ VS. openin (subtopology euclidean c) u) \ + (\u \ VS. openin (top_of_set c) u) \ pairwise disjnt VS \ (\u \ VS. \q. homeomorphism u X p q)" using p by (auto simp: covering_space_def) have "T \ c" by (metis openin_euclidean_subtopology_iff T) - have "\X. openin (subtopology euclidean S) X \ y \ X \ X \ p ` T" + have "\X. openin (top_of_set S) X \ y \ X \ X \ p ` T" if "y \ p ` T" for y proof - have "y \ S" using \T \ c\ pce that by blast - obtain U VS where "y \ U" and U: "openin (subtopology euclidean S) U" + obtain U VS where "y \ U" and U: "openin (top_of_set S) U" and VS: "\VS = c \ p -` U" - and openVS: "\V \ VS. openin (subtopology euclidean c) V" + and openVS: "\V \ VS. openin (top_of_set c) V" and homVS: "\V. V \ VS \ \q. homeomorphism V U p q" using covs [OF \y \ S\] by auto obtain x where "x \ c" "p x \ U" "x \ T" "p x = y" @@ -1349,14 +1349,14 @@ then obtain q where q: "homeomorphism V U p q" using homVS by blast then have ptV: "p ` (T \ V) = U \ q -` (T \ V)" using VS \V \ VS\ by (auto simp: homeomorphism_def) - have ocv: "openin (subtopology euclidean c) V" + have ocv: "openin (top_of_set c) V" by (simp add: \V \ VS\ openVS) - have "openin (subtopology euclidean U) (U \ q -` (T \ V))" + have "openin (top_of_set U) (U \ q -` (T \ V))" apply (rule continuous_on_open [THEN iffD1, rule_format]) using homeomorphism_def q apply blast using openin_subtopology_Int_subset [of c] q T unfolding homeomorphism_def by (metis inf.absorb_iff2 Int_commute ocv openin_euclidean_subtopology_iff) - then have os: "openin (subtopology euclidean S) (U \ q -` (T \ V))" + then have os: "openin (top_of_set S) (U \ q -` (T \ V))" using openin_trans [of U] by (simp add: Collect_conj_eq U) show ?thesis apply (rule_tac x = "p ` (T \ V)" in exI) @@ -1386,13 +1386,13 @@ have "connected U" by (rule in_components_connected [OF u_compt]) have contu: "continuous_on U g1" "continuous_on U g2" using \U \ T\ continuous_on_subset g1 g2 by blast+ - have o12: "openin (subtopology euclidean U) G12" + have o12: "openin (top_of_set U) G12" unfolding G12_def proof (subst openin_subopen, clarify) fix z assume z: "z \ U" "g1 z - g2 z = 0" - obtain v w q where "g1 z \ v" and ocv: "openin (subtopology euclidean c) v" - and "p (g1 z) \ w" and osw: "openin (subtopology euclidean S) w" + obtain v w q where "g1 z \ v" and ocv: "openin (top_of_set c) v" + and "p (g1 z) \ w" and osw: "openin (top_of_set S) w" and hom: "homeomorphism v w p q" apply (rule_tac x = "g1 z" in covering_space_local_homeomorphism [OF cov]) using \U \ T\ \z \ U\ g1(2) apply blast+ @@ -1400,15 +1400,15 @@ have "g2 z \ v" using \g1 z \ v\ z by auto have gg: "U \ g -` v = U \ g -` (v \ g ` U)" for g by auto - have "openin (subtopology euclidean (g1 ` U)) (v \ g1 ` U)" + have "openin (top_of_set (g1 ` U)) (v \ g1 ` U)" using ocv \U \ T\ g1 by (fastforce simp add: openin_open) - then have 1: "openin (subtopology euclidean U) (U \ g1 -` v)" + then have 1: "openin (top_of_set U) (U \ g1 -` v)" unfolding gg by (blast intro: contu continuous_on_open [THEN iffD1, rule_format]) - have "openin (subtopology euclidean (g2 ` U)) (v \ g2 ` U)" + have "openin (top_of_set (g2 ` U)) (v \ g2 ` U)" using ocv \U \ T\ g2 by (fastforce simp add: openin_open) - then have 2: "openin (subtopology euclidean U) (U \ g2 -` v)" + then have 2: "openin (top_of_set U) (U \ g2 -` v)" unfolding gg by (blast intro: contu continuous_on_open [THEN iffD1, rule_format]) - show "\T. openin (subtopology euclidean U) T \ z \ T \ T \ {z \ U. g1 z - g2 z = 0}" + show "\T. openin (top_of_set U) T \ z \ T \ T \ {z \ U. g1 z - g2 z = 0}" using z apply (rule_tac x = "(U \ g1 -` v) \ (U \ g2 -` v)" in exI) apply (intro conjI) @@ -1417,7 +1417,7 @@ apply (metis \U \ T\ subsetD eq_iff_diff_eq_0 fg1 fg2 hom homeomorphism_def) done qed - have c12: "closedin (subtopology euclidean U) G12" + have c12: "closedin (top_of_set U) G12" unfolding G12_def by (intro continuous_intros continuous_closedin_preimage_constant contu) have "G12 = {} \ G12 = U" @@ -1468,37 +1468,37 @@ show ?rhs proof (rule locallyI) fix V x - assume V: "openin (subtopology euclidean C) V" and "x \ V" + assume V: "openin (top_of_set C) V" and "x \ V" have "p x \ p ` C" by (metis IntE V \x \ V\ imageI openin_open) then obtain T \ where "p x \ T" - and opeT: "openin (subtopology euclidean S) T" + and opeT: "openin (top_of_set S) T" and veq: "\\ = C \ p -` T" - and ope: "\U\\. openin (subtopology euclidean C) U" + and ope: "\U\\. openin (top_of_set C) U" and hom: "\U\\. \q. homeomorphism U T p q" using cov unfolding covering_space_def by (blast intro: that) have "x \ \\" using V veq \p x \ T\ \x \ V\ openin_imp_subset by fastforce then obtain U where "x \ U" "U \ \" by blast - then obtain q where opeU: "openin (subtopology euclidean C) U" and q: "homeomorphism U T p q" + then obtain q where opeU: "openin (top_of_set C) U" and q: "homeomorphism U T p q" using ope hom by blast - with V have "openin (subtopology euclidean C) (U \ V)" + with V have "openin (top_of_set C) (U \ V)" by blast - then have UV: "openin (subtopology euclidean S) (p ` (U \ V))" + then have UV: "openin (top_of_set S) (p ` (U \ V))" using cov covering_space_open_map by blast - obtain W W' where opeW: "openin (subtopology euclidean S) W" and "\ W'" "p x \ W" "W \ W'" and W'sub: "W' \ p ` (U \ V)" + obtain W W' where opeW: "openin (top_of_set S) W" and "\ W'" "p x \ W" "W \ W'" and W'sub: "W' \ p ` (U \ V)" using locallyE [OF L UV] \x \ U\ \x \ V\ by blast then have "W \ T" by (metis Int_lower1 q homeomorphism_image1 image_Int_subset order_trans) - show "\U Z. openin (subtopology euclidean C) U \ + show "\U Z. openin (top_of_set C) U \ \ Z \ x \ U \ U \ Z \ Z \ V" proof (intro exI conjI) - have "openin (subtopology euclidean T) W" + have "openin (top_of_set T) W" by (meson opeW opeT openin_imp_subset openin_subset_trans \W \ T\) - then have "openin (subtopology euclidean U) (q ` W)" + then have "openin (top_of_set U) (q ` W)" by (meson homeomorphism_imp_open_map homeomorphism_symD q) - then show "openin (subtopology euclidean C) (q ` W)" + then show "openin (top_of_set C) (q ` W)" using opeU openin_trans by blast show "\ (q ` W')" by (metis (mono_tags, lifting) Int_subset_iff UV W'sub \\ W'\ continuous_on_subset dual_order.trans homeomorphism_def image_Int_subset openin_imp_subset q qim) @@ -1575,18 +1575,18 @@ "\y. y \ U \ k(0, y) = f y" "\z. z \ {0..1} \ U \ h z = p(k z)" proof - - have "\V k. openin (subtopology euclidean U) V \ y \ V \ + have "\V k. openin (top_of_set U) V \ y \ V \ continuous_on ({0..1} \ V) k \ k ` ({0..1} \ V) \ C \ (\z \ V. k(0, z) = f z) \ (\z \ {0..1} \ V. h z = p(k z))" if "y \ U" for y proof - - obtain UU where UU: "\s. s \ S \ s \ (UU s) \ openin (subtopology euclidean S) (UU s) \ + obtain UU where UU: "\s. s \ S \ s \ (UU s) \ openin (top_of_set S) (UU s) \ (\\. \\ = C \ p -` UU s \ - (\U \ \. openin (subtopology euclidean C) U) \ + (\U \ \. openin (top_of_set C) U) \ pairwise disjnt \ \ (\U \ \. \q. homeomorphism U (UU s) p q))" using cov unfolding covering_space_def by (metis (mono_tags)) - then have ope: "\s. s \ S \ s \ (UU s) \ openin (subtopology euclidean S) (UU s)" + then have ope: "\s. s \ S \ s \ (UU s) \ openin (top_of_set S) (UU s)" by blast have "\k n i. open k \ open n \ t \ k \ y \ n \ i \ S \ h ` (({0..1} \ k) \ (U \ n)) \ UU i" if "t \ {0..1}" for t @@ -1595,7 +1595,7 @@ using \y \ U\ him that by blast then have "(t,y) \ ({0..1} \ U) \ h -` UU(h(t, y))" using \y \ U\ \t \ {0..1}\ by (auto simp: ope) - moreover have ope_01U: "openin (subtopology euclidean ({0..1} \ U)) (({0..1} \ U) \ h -` UU(h(t, y)))" + moreover have ope_01U: "openin (top_of_set ({0..1} \ U)) (({0..1} \ U) \ h -` UU(h(t, y)))" using hinS ope continuous_on_open_gen [OF him] conth by blast ultimately obtain V W where opeV: "open V" and "t \ {0..1} \ V" "t \ {0..1} \ V" and opeW: "open W" and "y \ U" "y \ W" @@ -1641,7 +1641,7 @@ using reals_Archimedean2 by blast then have "N > 0" using \0 < \\ order.asym by force - have *: "\V k. openin (subtopology euclidean U) V \ y \ V \ + have *: "\V k. openin (top_of_set U) V \ y \ V \ continuous_on ({0..of_nat n / N} \ V) k \ k ` ({0..of_nat n / N} \ V) \ C \ (\z\V. k (0, z) = f z) \ @@ -1657,7 +1657,7 @@ done next case (Suc n) - then obtain V k where opeUV: "openin (subtopology euclidean U) V" + then obtain V k where opeUV: "openin (top_of_set U) V" and "y \ V" and contk: "continuous_on ({0..real n / real N} \ V) k" and kim: "k ` ({0..real n / real N} \ V) \ C" @@ -1676,7 +1676,7 @@ have t01: "t \ {0..1}" using \t \ tk\ \tk \ {0..1}\ by blast obtain \ where \: "\\ = C \ p -` UU (X t)" - and opeC: "\U. U \ \ \ openin (subtopology euclidean C) U" + and opeC: "\U. U \ \ \ openin (top_of_set C) U" and "pairwise disjnt \" and homuu: "\U. U \ \ \ \q. homeomorphism U (UU (X t)) p q" using inUS [OF t01] UU by meson @@ -1703,24 +1703,24 @@ by blast then obtain W where W: "k (real n / real N, y) \ W" and "W \ \" by blast - then obtain p' where opeC': "openin (subtopology euclidean C) W" + then obtain p' where opeC': "openin (top_of_set C) W" and hom': "homeomorphism W (UU (X t)) p p'" using homuu opeC by blast then have "W \ C" using openin_imp_subset by blast define W' where "W' = UU(X t)" - have opeVW: "openin (subtopology euclidean V) (V \ (k \ Pair (n / N)) -` W)" + have opeVW: "openin (top_of_set V) (V \ (k \ Pair (n / N)) -` W)" apply (rule continuous_openin_preimage [OF _ _ opeC']) apply (intro continuous_intros continuous_on_subset [OF contk]) using kim apply (auto simp: \y \ V\ W) done - obtain N' where opeUN': "openin (subtopology euclidean U) N'" + obtain N' where opeUN': "openin (top_of_set U) N'" and "y \ N'" and kimw: "k ` ({(real n / real N)} \ N') \ W" apply (rule_tac N' = "(V \ (k \ Pair (n / N)) -` W)" in that) apply (fastforce simp: \y \ V\ W intro!: openin_trans [OF opeVW opeUV])+ done - obtain Q Q' where opeUQ: "openin (subtopology euclidean U) Q" - and cloUQ': "closedin (subtopology euclidean U) Q'" + obtain Q Q' where opeUQ: "openin (top_of_set U) Q" + and cloUQ': "closedin (top_of_set U) Q'" and "y \ Q" "Q \ Q'" and Q': "Q' \ (U \ NN(t)) \ N' \ V" proof - @@ -1732,9 +1732,9 @@ by (metis Int_iff \N' = U \ VX\ \V = U \ VO\ \y \ N'\ \y \ V\ inUS open_contains_cball t01) show ?thesis proof - show "openin (subtopology euclidean U) (U \ ball y e)" + show "openin (top_of_set U) (U \ ball y e)" by blast - show "closedin (subtopology euclidean U) (U \ cball y e)" + show "closedin (top_of_set U) (U \ cball y e)" using e by (auto simp: closedin_closed) qed (use \y \ U\ \e > 0\ VO VX e in auto) qed @@ -1749,13 +1749,13 @@ (\x. if x \ {0..real n / real N} \ Q' then k x else (p' \ h) x)" unfolding neqQ' [symmetric] proof (rule continuous_on_cases_local, simp_all add: neqQ' del: comp_apply) - show "closedin (subtopology euclidean ({0..(1 + real n) / real N} \ Q')) + show "closedin (top_of_set ({0..(1 + real n) / real N} \ Q')) ({0..real n / real N} \ Q')" apply (simp add: closedin_closed) apply (rule_tac x="{0 .. real n / real N} \ UNIV" in exI) using n_div_N_in apply (auto simp: closed_Times) done - show "closedin (subtopology euclidean ({0..(1 + real n) / real N} \ Q')) + show "closedin (top_of_set ({0..(1 + real n) / real N} \ Q')) ({real n / real N..(1 + real n) / real N} \ Q')" apply (simp add: closedin_closed) apply (rule_tac x="{real n / real N .. (1 + real n) / real N} \ UNIV" in exI) @@ -1845,7 +1845,7 @@ show ?thesis using*[OF order_refl] N \0 < \\ by (simp add: split: if_split_asm) qed - then obtain V fs where opeV: "\y. y \ U \ openin (subtopology euclidean U) (V y)" + then obtain V fs where opeV: "\y. y \ U \ openin (top_of_set U) (V y)" and V: "\y. y \ U \ y \ V y" and contfs: "\y. y \ U \ continuous_on ({0..1} \ V y) (fs y)" and *: "\y. y \ U \ (fs y) ` ({0..1} \ V y) \ C \ @@ -1857,14 +1857,17 @@ obtain k where contk: "continuous_on ({0..1} \ U) k" and k: "\x i. \i \ U; x \ {0..1} \ U \ {0..1} \ V i\ \ k x = fs i x" proof (rule pasting_lemma_exists) - show "{0..1} \ U \ (\i\U. {0..1} \ V i)" - apply auto - using V by blast - show "\i. i \ U \ openin (subtopology euclidean ({0..1} \ U)) ({0..1} \ V i)" + let ?X = "top_of_set ({0..1::real} \ U)" + show "topspace ?X \ (\i\U. {0..1} \ V i)" + using V by force + show "\i. i \ U \ openin (top_of_set ({0..1} \ U)) ({0..1} \ V i)" by (simp add: opeV openin_Times) - show "\i. i \ U \ continuous_on ({0..1} \ V i) (fs i)" - using contfs by blast - show "fs i x = fs j x" if "i \ U" "j \ U" and x: "x \ {0..1} \ U \ {0..1} \ V i \ {0..1} \ V j" + show "\i. i \ U \ continuous_map + (subtopology (top_of_set ({0..1} \ U)) ({0..1} \ V i)) euclidean (fs i)" + using contfs + apply simp + by (metis continuous_map_iff_continuous continuous_on_subset openin_imp_subset openin_subtopology_self subtopology_subtopology) + show "fs i x = fs j x" if "i \ U" "j \ U" and x: "x \ topspace ?X \ {0..1} \ V i \ {0..1} \ V j" for i j x proof - obtain u y where "x = (u, y)" "y \ V i" "y \ V j" "0 \ u" "u \ 1" @@ -1900,7 +1903,7 @@ using \x = (u, y)\ x by blast qed qed - qed blast + qed force show ?thesis proof show "k ` ({0..1} \ U) \ C" @@ -2282,17 +2285,17 @@ using l [of "linepath z z" z "linepath a a"] by (auto simp: assms) show LC: "l ` U \ C" by (clarify dest!: *) (metis (full_types) l pathfinish_in_path_image subsetCE) - have "\T. openin (subtopology euclidean U) T \ y \ T \ T \ U \ l -` X" - if X: "openin (subtopology euclidean C) X" and "y \ U" "l y \ X" for X y + have "\T. openin (top_of_set U) T \ y \ T \ T \ U \ l -` X" + if X: "openin (top_of_set C) X" and "y \ U" "l y \ X" for X y proof - have "X \ C" using X openin_euclidean_subtopology_iff by blast have "f y \ S" using fim \y \ U\ by blast then obtain W \ - where WV: "f y \ W \ openin (subtopology euclidean S) W \ + where WV: "f y \ W \ openin (top_of_set S) W \ (\\ = C \ p -` W \ - (\U \ \. openin (subtopology euclidean C) U) \ + (\U \ \. openin (top_of_set C) U) \ pairwise disjnt \ \ (\U \ \. \q. homeomorphism U W p q))" using cov by (force simp: covering_space_def) @@ -2300,15 +2303,15 @@ using \X \ C\ pleq that by auto then obtain W' where "l y \ W'" and "W' \ \" by blast - with WV obtain p' where opeCW': "openin (subtopology euclidean C) W'" + with WV obtain p' where opeCW': "openin (top_of_set C) W'" and homUW': "homeomorphism W' W p p'" by blast then have contp': "continuous_on W p'" and p'im: "p' ` W \ W'" using homUW' homeomorphism_image2 homeomorphism_cont2 by fastforce+ obtain V where "y \ V" "y \ U" and fimW: "f ` V \ W" "V \ U" - and "path_connected V" and opeUV: "openin (subtopology euclidean U) V" + and "path_connected V" and opeUV: "openin (top_of_set U) V" proof - - have "openin (subtopology euclidean U) (U \ f -` W)" + have "openin (top_of_set U) (U \ f -` W)" using WV contf continuous_on_open_gen fim by auto then show ?thesis using U WV @@ -2324,16 +2327,16 @@ using homUW' homeomorphism_image2 by fastforce show ?thesis proof (intro exI conjI) - have "openin (subtopology euclidean S) (W \ p' -` (W' \ X))" + have "openin (top_of_set S) (W \ p' -` (W' \ X))" proof (rule openin_trans) - show "openin (subtopology euclidean W) (W \ p' -` (W' \ X))" + show "openin (top_of_set W) (W \ p' -` (W' \ X))" apply (rule continuous_openin_preimage [OF contp' p'im]) using X \W' \ C\ apply (auto simp: openin_open) done - show "openin (subtopology euclidean S) W" + show "openin (top_of_set S) W" using WV by blast qed - then show "openin (subtopology euclidean U) (V \ (U \ (f -` (W \ (p' -` (W' \ X))))))" + then show "openin (top_of_set U) (V \ (U \ (f -` (W \ (p' -` (W' \ X))))))" by (blast intro: opeUV openin_subtopology_self continuous_openin_preimage [OF contf fim]) have "p' (f y) \ X" using \l y \ W'\ homeomorphism_apply1 [OF homUW'] pleq \y \ U\ \l y \ X\ by fastforce diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Homotopy.thy --- a/src/HOL/Analysis/Homotopy.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Homotopy.thy Tue Mar 19 16:14:51 2019 +0000 @@ -258,12 +258,12 @@ assumes "homotopic_with P X Y f g" and "homotopic_with P X Y g h" shows "homotopic_with P X Y f h" proof - - have clo1: "closedin (subtopology euclidean ({0..1/2} \ X \ {1/2..1} \ X)) ({0..1/2::real} \ X)" + have clo1: "closedin (top_of_set ({0..1/2} \ X \ {1/2..1} \ X)) ({0..1/2::real} \ X)" apply (simp add: closedin_closed split_01_prod [symmetric]) apply (rule_tac x="{0..1/2} \ UNIV" in exI) apply (force simp: closed_Times) done - have clo2: "closedin (subtopology euclidean ({0..1/2} \ X \ {1/2..1} \ X)) ({1/2..1::real} \ X)" + have clo2: "closedin (top_of_set ({0..1/2} \ X \ {1/2..1} \ X)) ({1/2..1::real} \ X)" apply (simp add: closedin_closed split_01_prod [symmetric]) apply (rule_tac x="{1/2..1} \ UNIV" in exI) apply (force simp: closed_Times) @@ -1472,20 +1472,20 @@ definition%important locally :: "('a::topological_space set \ bool) \ 'a set \ bool" where "locally P S \ - \w x. openin (subtopology euclidean S) w \ x \ w - \ (\u v. openin (subtopology euclidean S) u \ P v \ + \w x. openin (top_of_set S) w \ x \ w + \ (\u v. openin (top_of_set S) u \ P v \ x \ u \ u \ v \ v \ w)" lemma locallyI: - assumes "\w x. \openin (subtopology euclidean S) w; x \ w\ - \ \u v. openin (subtopology euclidean S) u \ P v \ + assumes "\w x. \openin (top_of_set S) w; x \ w\ + \ \u v. openin (top_of_set S) u \ P v \ x \ u \ u \ v \ v \ w" shows "locally P S" using assms by (force simp: locally_def) lemma locallyE: - assumes "locally P S" "openin (subtopology euclidean S) w" "x \ w" - obtains u v where "openin (subtopology euclidean S) u" + assumes "locally P S" "openin (top_of_set S) w" "x \ w" + obtains u v where "openin (top_of_set S) u" "P v" "x \ u" "u \ v" "v \ w" using assms unfolding locally_def by meson @@ -1495,7 +1495,7 @@ by (metis assms locally_def) lemma locally_open_subset: - assumes "locally P S" "openin (subtopology euclidean S) t" + assumes "locally P S" "openin (top_of_set S) t" shows "locally P t" using assms apply (simp add: locally_def) @@ -1507,7 +1507,7 @@ by (metis (no_types, hide_lams) Int_absorb1 Int_lower1 Int_subset_iff openin_open openin_subtopology_Int_subset) lemma locally_diff_closed: - "\locally P S; closedin (subtopology euclidean S) t\ \ locally P (S - t)" + "\locally P S; closedin (top_of_set S) t\ \ locally P (S - t)" using locally_open_subset closedin_def by fastforce lemma locally_empty [iff]: "locally P {}" @@ -1550,15 +1550,15 @@ unfolding locally_def proof (clarify) fix W x y - assume W: "openin (subtopology euclidean (S \ T)) W" and xy: "(x, y) \ W" - then obtain U V where "openin (subtopology euclidean S) U" "x \ U" - "openin (subtopology euclidean T) V" "y \ V" "U \ V \ W" + assume W: "openin (top_of_set (S \ T)) W" and xy: "(x, y) \ W" + then obtain U V where "openin (top_of_set S) U" "x \ U" + "openin (top_of_set T) V" "y \ V" "U \ V \ W" using Times_in_interior_subtopology by metis then obtain U1 U2 V1 V2 - where opeS: "openin (subtopology euclidean S) U1 \ P U2 \ x \ U1 \ U1 \ U2 \ U2 \ U" - and opeT: "openin (subtopology euclidean T) V1 \ Q V2 \ y \ V1 \ V1 \ V2 \ V2 \ V" + where opeS: "openin (top_of_set S) U1 \ P U2 \ x \ U1 \ U1 \ U2 \ U2 \ U" + and opeT: "openin (top_of_set T) V1 \ Q V2 \ y \ V1 \ V1 \ V2 \ V2 \ V" by (meson PS QT locallyE) - with \U \ V \ W\ show "\u v. openin (subtopology euclidean (S \ T)) u \ R v \ (x,y) \ u \ u \ v \ v \ W" + with \U \ V \ W\ show "\u v. openin (top_of_set (S \ T)) u \ R v \ (x,y) \ u \ u \ v \ v \ W" apply (rule_tac x="U1 \ V1" in exI) apply (rule_tac x="U2 \ V2" in exI) apply (auto simp: openin_Times R) @@ -1573,7 +1573,7 @@ shows "locally Q t" proof (clarsimp simp: locally_def) fix W y - assume "y \ W" and "openin (subtopology euclidean t) W" + assume "y \ W" and "openin (top_of_set t) W" then obtain T where T: "open T" "W = t \ T" by (force simp: openin_open) then have "W \ t" by auto @@ -1586,15 +1586,15 @@ using \g ` t = S\ \W \ t\ apply blast using g \W \ t\ apply auto[1] by (simp add: f rev_image_eqI) - have \: "openin (subtopology euclidean S) (g ` W)" + have \: "openin (top_of_set S) (g ` W)" proof - have "continuous_on S f" using f(3) by blast - then show "openin (subtopology euclidean S) (g ` W)" - by (simp add: gw Collect_conj_eq \openin (subtopology euclidean t) W\ continuous_on_open f(2)) + then show "openin (top_of_set S) (g ` W)" + by (simp add: gw Collect_conj_eq \openin (top_of_set t) W\ continuous_on_open f(2)) qed then obtain u v - where osu: "openin (subtopology euclidean S) u" and uv: "P v" "g y \ u" "u \ v" "v \ g ` W" + where osu: "openin (top_of_set S) u" and uv: "P v" "g y \ u" "u \ v" "v \ g ` W" using S [unfolded locally_def, rule_format, of "g ` W" "g y"] \y \ W\ by force have "v \ S" using uv by (simp add: gw) have fv: "f ` v = t \ {x. g x \ v}" @@ -1609,7 +1609,7 @@ using \v \ S\ \W \ t\ f apply (simp add: homeomorphism_def contvf contvg, auto) by (metis f(1) rev_image_eqI rev_subsetD) - have 1: "openin (subtopology euclidean t) (t \ g -` u)" + have 1: "openin (top_of_set t) (t \ g -` u)" apply (rule continuous_on_open [THEN iffD1, rule_format]) apply (rule \continuous_on t g\) using \g ` t = S\ apply (simp add: osu) @@ -1619,7 +1619,7 @@ apply (intro conjI Q [OF \P v\ homv]) using \W \ t\ \y \ W\ \f ` v \ W\ uv apply (auto simp: fv) done - show "\U. openin (subtopology euclidean t) U \ (\v. Q v \ y \ U \ U \ v \ v \ W)" + show "\U. openin (top_of_set t) U \ (\v. Q v \ y \ U \ U \ v \ v \ W)" by (meson 1 2) qed @@ -1674,23 +1674,23 @@ fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes P: "locally P S" and f: "continuous_on S f" - and oo: "\t. openin (subtopology euclidean S) t - \ openin (subtopology euclidean (f ` S)) (f ` t)" + and oo: "\t. openin (top_of_set S) t + \ openin (top_of_set (f ` S)) (f ` t)" and Q: "\t. \t \ S; P t\ \ Q(f ` t)" shows "locally Q (f ` S)" proof (clarsimp simp add: locally_def) fix W y - assume oiw: "openin (subtopology euclidean (f ` S)) W" and "y \ W" + assume oiw: "openin (top_of_set (f ` S)) W" and "y \ W" then have "W \ f ` S" by (simp add: openin_euclidean_subtopology_iff) - have oivf: "openin (subtopology euclidean S) (S \ f -` W)" + have oivf: "openin (top_of_set S) (S \ f -` W)" by (rule continuous_on_open [THEN iffD1, rule_format, OF f oiw]) then obtain x where "x \ S" "f x = y" using \W \ f ` S\ \y \ W\ by blast then obtain U V - where "openin (subtopology euclidean S) U" "P V" "x \ U" "U \ V" "V \ S \ f -` W" + where "openin (top_of_set S) U" "P V" "x \ U" "U \ V" "V \ S \ f -` W" using P [unfolded locally_def, rule_format, of "(S \ f -` W)" x] oivf \y \ W\ by auto - then show "\X. openin (subtopology euclidean (f ` S)) X \ (\Y. Q Y \ y \ X \ X \ Y \ Y \ W)" + then show "\X. openin (top_of_set (f ` S)) X \ (\Y. Q Y \ y \ X \ X \ Y \ Y \ W)" apply (rule_tac x="f ` U" in exI) apply (rule conjI, blast intro!: oo) apply (rule_tac x="f ` V" in exI) @@ -1703,21 +1703,21 @@ proposition connected_induction: assumes "connected S" - and opD: "\T a. \openin (subtopology euclidean S) T; a \ T\ \ \z. z \ T \ P z" + and opD: "\T a. \openin (top_of_set S) T; a \ T\ \ \z. z \ T \ P z" and opI: "\a. a \ S - \ \T. openin (subtopology euclidean S) T \ a \ T \ + \ \T. openin (top_of_set S) T \ a \ T \ (\x \ T. \y \ T. P x \ P y \ Q x \ Q y)" and etc: "a \ S" "b \ S" "P a" "P b" "Q a" shows "Q b" proof - - have 1: "openin (subtopology euclidean S) - {b. \T. openin (subtopology euclidean S) T \ + have 1: "openin (top_of_set S) + {b. \T. openin (top_of_set S) T \ b \ T \ (\x\T. P x \ Q x)}" apply (subst openin_subopen, clarify) apply (rule_tac x=T in exI, auto) done - have 2: "openin (subtopology euclidean S) - {b. \T. openin (subtopology euclidean S) T \ + have 2: "openin (top_of_set S) + {b. \T. openin (top_of_set S) T \ b \ T \ (\x\T. P x \ \ Q x)}" apply (subst openin_subopen, clarify) apply (rule_tac x=T in exI, auto) @@ -1738,9 +1738,9 @@ assumes "connected S" and etc: "a \ S" "b \ S" "P a" "P b" and trans: "\x y z. \R x y; R y z\ \ R x z" - and opD: "\T a. \openin (subtopology euclidean S) T; a \ T\ \ \z. z \ T \ P z" + and opD: "\T a. \openin (top_of_set S) T; a \ T\ \ \z. z \ T \ P z" and opI: "\a. a \ S - \ \T. openin (subtopology euclidean S) T \ a \ T \ + \ \T. openin (top_of_set S) T \ a \ T \ (\x \ T. \y \ T. P x \ P y \ R x y)" shows "R a b" proof - @@ -1754,7 +1754,7 @@ assumes "connected S" and etc: "a \ S" "b \ S" "P a" and opI: "\a. a \ S - \ \T. openin (subtopology euclidean S) T \ a \ T \ + \ \T. openin (top_of_set S) T \ a \ T \ (\x \ T. \y \ T. P x \ P y)" shows "P b" apply (rule connected_induction [OF \connected S\ _, where P = "\x. True"], blast) @@ -1767,7 +1767,7 @@ and etc: "a \ S" "b \ S" and sym: "\x y. \R x y; x \ S; y \ S\ \ R y x" and trans: "\x y z. \R x y; R y z; x \ S; y \ S; z \ S\ \ R x z" - and opI: "\a. a \ S \ \T. openin (subtopology euclidean S) T \ a \ T \ (\x \ T. R a x)" + and opI: "\a. a \ S \ \T. openin (top_of_set S) T \ a \ T \ (\x \ T. R a x)" shows "R a b" proof - have "\a b c. \a \ S; b \ S; c \ S; R a b\ \ R a c" @@ -1779,7 +1779,7 @@ lemma locally_constant_imp_constant: assumes "connected S" and opI: "\a. a \ S - \ \T. openin (subtopology euclidean S) T \ a \ T \ (\x \ T. f x = f a)" + \ \T. openin (top_of_set S) T \ a \ T \ (\x \ T. f x = f a)" shows "f constant_on S" proof - have "\x y. x \ S \ y \ S \ f x = f y" @@ -1805,7 +1805,7 @@ shows "locally compact s \ (\x \ s. \u v. x \ u \ u \ v \ v \ s \ - openin (subtopology euclidean s) u \ compact v)" + openin (top_of_set s) u \ compact v)" (is "?lhs = ?rhs") proof assume ?lhs @@ -1816,11 +1816,11 @@ next assume r [rule_format]: ?rhs have *: "\u v. - openin (subtopology euclidean s) u \ + openin (top_of_set s) u \ compact v \ x \ u \ u \ v \ v \ s \ T" if "open T" "x \ s" "x \ T" for x T proof - - obtain u v where uv: "x \ u" "u \ v" "v \ s" "compact v" "openin (subtopology euclidean s) u" + obtain u v where uv: "x \ u" "u \ v" "v \ s" "compact v" "openin (top_of_set s) u" using r [OF \x \ s\] by auto obtain e where "e>0" and e: "cball x e \ T" using open_contains_cball \open T\ \x \ T\ by blast @@ -1842,7 +1842,7 @@ fixes s :: "'a :: metric_space set" assumes "locally compact s" obtains u v where "\x. x \ s \ x \ u x \ u x \ v x \ v x \ s \ - openin (subtopology euclidean s) (u x) \ compact (v x)" + openin (top_of_set s) (u x) \ compact (v x)" using assms unfolding locally_compact by metis @@ -1850,7 +1850,7 @@ fixes s :: "'a :: heine_borel set" shows "locally compact s \ (\x \ s. \u. x \ u \ - openin (subtopology euclidean s) u \ compact(closure u) \ closure u \ s)" + openin (top_of_set s) u \ compact(closure u) \ closure u \ s)" apply (simp add: locally_compact) apply (intro ball_cong ex_cong refl iffI) apply (metis bounded_subset closure_eq closure_mono compact_eq_bounded_closed dual_order.trans) @@ -1884,21 +1884,21 @@ shows "locally compact s \ (\k. k \ s \ compact k \ (\u v. k \ u \ u \ v \ v \ s \ - openin (subtopology euclidean s) u \ compact v))" + openin (top_of_set s) u \ compact v))" (is "?lhs = ?rhs") proof assume ?lhs then obtain u v where uv: "\x. x \ s \ x \ u x \ u x \ v x \ v x \ s \ - openin (subtopology euclidean s) (u x) \ compact (v x)" + openin (top_of_set s) (u x) \ compact (v x)" by (metis locally_compactE) - have *: "\u v. k \ u \ u \ v \ v \ s \ openin (subtopology euclidean s) u \ compact v" + have *: "\u v. k \ u \ u \ v \ v \ s \ openin (top_of_set s) u \ compact v" if "k \ s" "compact k" for k proof - - have "\C. (\c\C. openin (subtopology euclidean k) c) \ k \ \C \ + have "\C. (\c\C. openin (top_of_set k) c) \ k \ \C \ \D\C. finite D \ k \ \D" using that by (simp add: compact_eq_openin_cover) - moreover have "\c \ (\x. k \ u x) ` k. openin (subtopology euclidean k) c" + moreover have "\c \ (\x. k \ u x) ` k. openin (top_of_set k) c" using that by clarify (metis subsetD inf.absorb_iff2 openin_subset openin_subtopology_Int_subset topspace_euclidean_subtopology uv) moreover have "k \ \((\x. k \ u x) ` k)" using that by clarsimp (meson subsetCE uv) @@ -1931,12 +1931,12 @@ assumes "open s" shows "locally compact s" proof - - have *: "\u v. x \ u \ u \ v \ v \ s \ openin (subtopology euclidean s) u \ compact v" + have *: "\u v. x \ u \ u \ v \ v \ s \ openin (top_of_set s) u \ compact v" if "x \ s" for x proof - obtain e where "e>0" and e: "cball x e \ s" using open_contains_cball assms \x \ s\ by blast - have ope: "openin (subtopology euclidean s) (ball x e)" + have ope: "openin (top_of_set s) (ball x e)" by (meson e open_ball ball_subset_cball dual_order.trans open_subset) show ?thesis apply (rule_tac x="ball x e" in exI) @@ -1955,7 +1955,7 @@ shows "locally compact s" proof - have *: "\u v. x \ u \ u \ v \ v \ s \ - openin (subtopology euclidean s) u \ compact v" + openin (top_of_set s) u \ compact v" if "x \ s" for x proof - show ?thesis @@ -1979,7 +1979,7 @@ lemma locally_compact_closedin: fixes s :: "'a :: heine_borel set" - shows "\closedin (subtopology euclidean s) t; locally compact s\ + shows "\closedin (top_of_set s) t; locally compact s\ \ locally compact t" unfolding closedin_closed using closed_imp_locally_compact locally_compact_Int by blast @@ -2010,8 +2010,8 @@ lemma locally_compact_openin_Un: fixes S :: "'a::euclidean_space set" assumes LCS: "locally compact S" and LCT:"locally compact T" - and opS: "openin (subtopology euclidean (S \ T)) S" - and opT: "openin (subtopology euclidean (S \ T)) T" + and opS: "openin (top_of_set (S \ T)) S" + and opT: "openin (top_of_set (S \ T)) T" shows "locally compact (S \ T)" proof - have "\e>0. closed (cball x e \ (S \ T))" if "x \ S" for x @@ -2047,8 +2047,8 @@ lemma locally_compact_closedin_Un: fixes S :: "'a::euclidean_space set" assumes LCS: "locally compact S" and LCT:"locally compact T" - and clS: "closedin (subtopology euclidean (S \ T)) S" - and clT: "closedin (subtopology euclidean (S \ T)) T" + and clS: "closedin (top_of_set (S \ T)) S" + and clT: "closedin (top_of_set (S \ T)) T" shows "locally compact (S \ T)" proof - have "\e>0. closed (cball x e \ (S \ T))" if "x \ S" "x \ T" for x @@ -2113,7 +2113,7 @@ "locally compact S \ (\K T. K \ S \ compact K \ open T \ K \ T \ (\U V. K \ U \ U \ V \ U \ T \ V \ S \ - openin (subtopology euclidean S) U \ compact V))" + openin (top_of_set S) U \ compact V))" (is "?lhs = ?rhs") proof assume L: ?lhs @@ -2122,10 +2122,10 @@ fix K :: "'a set" and T :: "'a set" assume "K \ S" and "compact K" and "open T" and "K \ T" obtain U V where "K \ U" "U \ V" "V \ S" "compact V" - and ope: "openin (subtopology euclidean S) U" + and ope: "openin (top_of_set S) U" using L unfolding locally_compact_compact by (meson \K \ S\ \compact K\) show "\U V. K \ U \ U \ V \ U \ T \ V \ S \ - openin (subtopology euclidean S) U \ compact V" + openin (top_of_set S) U \ compact V" proof (intro exI conjI) show "K \ U \ T" by (simp add: \K \ T\ \K \ U\) @@ -2133,7 +2133,7 @@ by (rule closure_subset) show "closure (U \ T) \ S" by (metis \U \ V\ \V \ S\ \compact V\ closure_closed closure_mono compact_imp_closed inf.cobounded1 subset_trans) - show "openin (subtopology euclidean S) (U \ T)" + show "openin (top_of_set S) (U \ T)" by (simp add: \open T\ ope openin_Int_open) show "compact (closure (U \ T))" by (meson Int_lower1 \U \ V\ \compact V\ bounded_subset compact_closure compact_eq_bounded_closed) @@ -2151,8 +2151,8 @@ proposition Sura_Bura_compact: fixes S :: "'a::euclidean_space set" assumes "compact S" and C: "C \ components S" - shows "C = \{T. C \ T \ openin (subtopology euclidean S) T \ - closedin (subtopology euclidean S) T}" + shows "C = \{T. C \ T \ openin (top_of_set S) T \ + closedin (top_of_set S) T}" (is "C = \?\") proof obtain x where x: "C = connected_component_set S x" and "x \ S" @@ -2168,8 +2168,8 @@ have clo: "closed (\?\)" by (simp add: \compact S\ closed_Inter closedin_compact_eq compact_imp_closed) have False - if K1: "closedin (subtopology euclidean (\?\)) K1" and - K2: "closedin (subtopology euclidean (\?\)) K2" and + if K1: "closedin (top_of_set (\?\)) K1" and + K2: "closedin (top_of_set (\?\)) K2" and K12_Int: "K1 \ K2 = {}" and K12_Un: "K1 \ K2 = \?\" and "K1 \ {}" "K2 \ {}" for K1 K2 proof - @@ -2187,8 +2187,8 @@ show "(S - (V1 \ V2)) \ \\ \ {}" if "finite \" and \: "\ \ ?\" for \ proof assume djo: "(S - (V1 \ V2)) \ \\ = {}" - obtain D where opeD: "openin (subtopology euclidean S) D" - and cloD: "closedin (subtopology euclidean S) D" + obtain D where opeD: "openin (top_of_set S) D" + and cloD: "closedin (top_of_set S) D" and "C \ D" and DV12: "D \ V1 \ V2" proof (cases "\ = {}") case True @@ -2197,9 +2197,9 @@ next case False show ?thesis proof - show ope: "openin (subtopology euclidean S) (\\)" + show ope: "openin (top_of_set S) (\\)" using openin_Inter \finite \\ False \ by blast - then show "closedin (subtopology euclidean S) (\\)" + then show "closedin (top_of_set S) (\\)" by (meson clo\ \ closed_Inter closed_subset openin_imp_subset subset_eq) show "C \ \\" using \ by auto @@ -2211,16 +2211,16 @@ by (simp add: x) have "closed D" using \compact S\ cloD closedin_closed_trans compact_imp_closed by blast - have cloV1: "closedin (subtopology euclidean D) (D \ closure V1)" - and cloV2: "closedin (subtopology euclidean D) (D \ closure V2)" + have cloV1: "closedin (top_of_set D) (D \ closure V1)" + and cloV2: "closedin (top_of_set D) (D \ closure V2)" by (simp_all add: closedin_closed_Int) moreover have "D \ closure V1 = D \ V1" "D \ closure V2 = D \ V2" apply safe using \D \ V1 \ V2\ \open V1\ \open V2\ V12 apply (simp_all add: closure_subset [THEN subsetD] closure_iff_nhds_not_empty, blast+) done - ultimately have cloDV1: "closedin (subtopology euclidean D) (D \ V1)" - and cloDV2: "closedin (subtopology euclidean D) (D \ V2)" + ultimately have cloDV1: "closedin (top_of_set D) (D \ V1)" + and cloDV2: "closedin (top_of_set D) (D \ V2)" by metis+ then obtain U1 U2 where "closed U1" "closed U2" and D1: "D \ V1 = D \ U1" and D2: "D \ V2 = D \ U2" @@ -2274,21 +2274,21 @@ fixes S :: "'a::euclidean_space set" assumes S: "locally compact S" and C: "C \ components S" and "compact C" and U: "open U" "C \ U" - obtains K where "openin (subtopology euclidean S) K" "compact K" "C \ K" "K \ U" + obtains K where "openin (top_of_set S) K" "compact K" "C \ K" "K \ U" proof (rule ccontr) assume "\ thesis" - with that have neg: "\K. openin (subtopology euclidean S) K \ compact K \ C \ K \ K \ U" + with that have neg: "\K. openin (top_of_set S) K \ compact K \ C \ K \ K \ U" by metis obtain V K where "C \ V" "V \ U" "V \ K" "K \ S" "compact K" - and opeSV: "openin (subtopology euclidean S) V" + and opeSV: "openin (top_of_set S) V" using S U \compact C\ apply (simp add: locally_compact_compact_subopen) by (meson C in_components_subset) - let ?\ = "{T. C \ T \ openin (subtopology euclidean K) T \ compact T \ T \ K}" + let ?\ = "{T. C \ T \ openin (top_of_set K) T \ compact T \ T \ K}" have CK: "C \ components K" by (meson C \C \ V\ \K \ S\ \V \ K\ components_intermediate_subset subset_trans) with \compact K\ - have "C = \{T. C \ T \ openin (subtopology euclidean K) T \ closedin (subtopology euclidean K) T}" + have "C = \{T. C \ T \ openin (top_of_set K) T \ closedin (top_of_set K) T}" by (simp add: Sura_Bura_compact) then have Ceq: "C = \?\" by (simp add: closedin_compact_eq \compact K\) @@ -2322,11 +2322,11 @@ by (metis (no_types, lifting) compact_Inter False mem_Collect_eq subsetCE \) moreover have "\\ \ K" using False that(2) by fastforce - moreover have opeKF: "openin (subtopology euclidean K) (\\)" + moreover have opeKF: "openin (top_of_set K) (\\)" using False \ \finite \\ by blast - then have opeVF: "openin (subtopology euclidean V) (\\)" + then have opeVF: "openin (top_of_set V) (\\)" using W \K \ S\ \V \ K\ opeKF \\\ \ K\ FUW openin_subset_trans by fastforce - then have "openin (subtopology euclidean S) (\\)" + then have "openin (top_of_set S) (\\)" by (metis opeSV openin_trans) moreover have "\\ \ U" by (meson \V \ U\ opeVF dual_order.trans openin_imp_subset) @@ -2343,8 +2343,8 @@ corollary Sura_Bura_clopen_subset_alt: fixes S :: "'a::euclidean_space set" assumes S: "locally compact S" and C: "C \ components S" and "compact C" - and opeSU: "openin (subtopology euclidean S) U" and "C \ U" - obtains K where "openin (subtopology euclidean S) K" "compact K" "C \ K" "K \ U" + and opeSU: "openin (top_of_set S) U" and "C \ U" + obtains K where "openin (top_of_set S) K" "compact K" "C \ K" "K \ U" proof - obtain V where "open V" "U = S \ V" using opeSU by (auto simp: openin_open) @@ -2358,13 +2358,13 @@ corollary Sura_Bura: fixes S :: "'a::euclidean_space set" assumes "locally compact S" "C \ components S" "compact C" - shows "C = \ {K. C \ K \ compact K \ openin (subtopology euclidean S) K}" + shows "C = \ {K. C \ K \ compact K \ openin (top_of_set S) K}" (is "C = ?rhs") proof show "?rhs \ C" proof (clarsimp, rule ccontr) fix x - assume *: "\X. C \ X \ compact X \ openin (subtopology euclidean S) X \ x \ X" + assume *: "\X. C \ X \ compact X \ openin (top_of_set S) X \ x \ X" and "x \ C" obtain U V where "open U" "open V" "{x} \ U" "C \ V" "U \ V = {}" using separation_normal [of "{x}" C] @@ -2381,8 +2381,8 @@ lemma locally_connected_1: assumes - "\v x. \openin (subtopology euclidean S) v; x \ v\ - \ \u. openin (subtopology euclidean S) u \ + "\v x. \openin (top_of_set S) v; x \ v\ + \ \u. openin (top_of_set S) u \ connected u \ x \ u \ u \ v" shows "locally connected S" apply (clarsimp simp add: locally_def) @@ -2391,12 +2391,12 @@ lemma locally_connected_2: assumes "locally connected S" - "openin (subtopology euclidean S) t" + "openin (top_of_set S) t" "x \ t" - shows "openin (subtopology euclidean S) (connected_component_set t x)" + shows "openin (top_of_set S) (connected_component_set t x)" proof - { fix y :: 'a - let ?SS = "subtopology euclidean S" + let ?SS = "top_of_set S" assume 1: "openin ?SS t" "\w x. openin ?SS w \ x \ w \ (\u. openin ?SS u \ (\v. connected v \ x \ u \ u \ v \ v \ w))" and "connected_component t x y" @@ -2420,29 +2420,29 @@ qed lemma locally_connected_3: - assumes "\t x. \openin (subtopology euclidean S) t; x \ t\ - \ openin (subtopology euclidean S) + assumes "\t x. \openin (top_of_set S) t; x \ t\ + \ openin (top_of_set S) (connected_component_set t x)" - "openin (subtopology euclidean S) v" "x \ v" - shows "\u. openin (subtopology euclidean S) u \ connected u \ x \ u \ u \ v" + "openin (top_of_set S) v" "x \ v" + shows "\u. openin (top_of_set S) u \ connected u \ x \ u \ u \ v" using assms connected_component_subset by fastforce lemma locally_connected: "locally connected S \ - (\v x. openin (subtopology euclidean S) v \ x \ v - \ (\u. openin (subtopology euclidean S) u \ connected u \ x \ u \ u \ v))" + (\v x. openin (top_of_set S) v \ x \ v + \ (\u. openin (top_of_set S) u \ connected u \ x \ u \ u \ v))" by (metis locally_connected_1 locally_connected_2 locally_connected_3) lemma locally_connected_open_connected_component: "locally connected S \ - (\t x. openin (subtopology euclidean S) t \ x \ t - \ openin (subtopology euclidean S) (connected_component_set t x))" + (\t x. openin (top_of_set S) t \ x \ t + \ openin (top_of_set S) (connected_component_set t x))" by (metis locally_connected_1 locally_connected_2 locally_connected_3) lemma locally_path_connected_1: assumes - "\v x. \openin (subtopology euclidean S) v; x \ v\ - \ \u. openin (subtopology euclidean S) u \ path_connected u \ x \ u \ u \ v" + "\v x. \openin (top_of_set S) v; x \ v\ + \ \u. openin (top_of_set S) u \ path_connected u \ x \ u \ u \ v" shows "locally path_connected S" apply (clarsimp simp add: locally_def) apply (drule assms; blast) @@ -2450,12 +2450,12 @@ lemma locally_path_connected_2: assumes "locally path_connected S" - "openin (subtopology euclidean S) t" + "openin (top_of_set S) t" "x \ t" - shows "openin (subtopology euclidean S) (path_component_set t x)" + shows "openin (top_of_set S) (path_component_set t x)" proof - { fix y :: 'a - let ?SS = "subtopology euclidean S" + let ?SS = "top_of_set S" assume 1: "openin ?SS t" "\w x. openin ?SS w \ x \ w \ (\u. openin ?SS u \ (\v. path_connected v \ x \ u \ u \ v \ v \ w))" and "path_component t x y" @@ -2479,10 +2479,10 @@ qed lemma locally_path_connected_3: - assumes "\t x. \openin (subtopology euclidean S) t; x \ t\ - \ openin (subtopology euclidean S) (path_component_set t x)" - "openin (subtopology euclidean S) v" "x \ v" - shows "\u. openin (subtopology euclidean S) u \ path_connected u \ x \ u \ u \ v" + assumes "\t x. \openin (top_of_set S) t; x \ t\ + \ openin (top_of_set S) (path_component_set t x)" + "openin (top_of_set S) v" "x \ v" + shows "\u. openin (top_of_set S) u \ path_connected u \ x \ u \ u \ v" proof - have "path_component v x x" by (meson assms(3) path_component_refl) @@ -2492,26 +2492,26 @@ proposition locally_path_connected: "locally path_connected S \ - (\v x. openin (subtopology euclidean S) v \ x \ v - \ (\u. openin (subtopology euclidean S) u \ path_connected u \ x \ u \ u \ v))" + (\v x. openin (top_of_set S) v \ x \ v + \ (\u. openin (top_of_set S) u \ path_connected u \ x \ u \ u \ v))" by (metis locally_path_connected_1 locally_path_connected_2 locally_path_connected_3) proposition locally_path_connected_open_path_component: "locally path_connected S \ - (\t x. openin (subtopology euclidean S) t \ x \ t - \ openin (subtopology euclidean S) (path_component_set t x))" + (\t x. openin (top_of_set S) t \ x \ t + \ openin (top_of_set S) (path_component_set t x))" by (metis locally_path_connected_1 locally_path_connected_2 locally_path_connected_3) lemma locally_connected_open_component: "locally connected S \ - (\t c. openin (subtopology euclidean S) t \ c \ components t - \ openin (subtopology euclidean S) c)" + (\t c. openin (top_of_set S) t \ c \ components t + \ openin (top_of_set S) c)" by (metis components_iff locally_connected_open_connected_component) proposition locally_connected_im_kleinen: "locally connected S \ - (\v x. openin (subtopology euclidean S) v \ x \ v - \ (\u. openin (subtopology euclidean S) u \ + (\v x. openin (top_of_set S) v \ x \ v + \ (\u. openin (top_of_set S) u \ x \ u \ u \ v \ (\y. y \ u \ (\c. connected c \ c \ v \ x \ c \ y \ c))))" (is "?lhs = ?rhs") @@ -2521,12 +2521,12 @@ by (fastforce simp add: locally_connected) next assume ?rhs - have *: "\T. openin (subtopology euclidean S) T \ x \ T \ T \ c" - if "openin (subtopology euclidean S) t" and c: "c \ components t" and "x \ c" for t c x + have *: "\T. openin (top_of_set S) T \ x \ T \ T \ c" + if "openin (top_of_set S) t" and c: "c \ components t" and "x \ c" for t c x proof - from that \?rhs\ [rule_format, of t x] obtain u where u: - "openin (subtopology euclidean S) u \ x \ u \ u \ t \ + "openin (top_of_set S) u \ x \ u \ u \ t \ (\y. y \ u \ (\c. connected c \ c \ t \ x \ c \ y \ c))" using in_components_subset by auto obtain F :: "'a set \ 'a set \ 'a" where @@ -2551,8 +2551,8 @@ proposition locally_path_connected_im_kleinen: "locally path_connected S \ - (\v x. openin (subtopology euclidean S) v \ x \ v - \ (\u. openin (subtopology euclidean S) u \ + (\v x. openin (top_of_set S) v \ x \ v + \ (\u. openin (top_of_set S) u \ x \ u \ u \ v \ (\y. y \ u \ (\p. path p \ path_image p \ v \ pathstart p = x \ pathfinish p = y))))" @@ -2565,15 +2565,15 @@ by (meson dual_order.trans) next assume ?rhs - have *: "\T. openin (subtopology euclidean S) T \ + have *: "\T. openin (top_of_set S) T \ x \ T \ T \ path_component_set u z" - if "openin (subtopology euclidean S) u" and "z \ u" and c: "path_component u z x" for u z x + if "openin (top_of_set S) u" and "z \ u" and c: "path_component u z x" for u z x proof - have "x \ u" by (meson c path_component_mem(2)) with that \?rhs\ [rule_format, of u x] obtain U where U: - "openin (subtopology euclidean S) U \ x \ U \ U \ u \ + "openin (top_of_set S) U \ x \ U \ U \ u \ (\y. y \ U \ (\p. path p \ path_image p \ u \ pathstart p = x \ pathfinish p = y))" by blast show ?thesis @@ -2626,22 +2626,22 @@ lemma openin_connected_component_locally_connected: "locally connected S - \ openin (subtopology euclidean S) (connected_component_set S x)" + \ openin (top_of_set S) (connected_component_set S x)" apply (simp add: locally_connected_open_connected_component) by (metis connected_component_eq_empty connected_component_subset open_empty open_subset openin_subtopology_self) lemma openin_components_locally_connected: - "\locally connected S; c \ components S\ \ openin (subtopology euclidean S) c" + "\locally connected S; c \ components S\ \ openin (top_of_set S) c" using locally_connected_open_component openin_subtopology_self by blast lemma openin_path_component_locally_path_connected: "locally path_connected S - \ openin (subtopology euclidean S) (path_component_set S x)" + \ openin (top_of_set S) (path_component_set S x)" by (metis (no_types) empty_iff locally_path_connected_2 openin_subopen openin_subtopology_self path_component_eq_empty) lemma closedin_path_component_locally_path_connected: "locally path_connected S - \ closedin (subtopology euclidean S) (path_component_set S x)" + \ closedin (top_of_set S) (path_component_set S x)" apply (simp add: closedin_def path_component_subset complement_path_component_Union) apply (rule openin_Union) using openin_path_component_locally_path_connected by auto @@ -2670,12 +2670,12 @@ shows "(path_component S x = connected_component S x)" proof (cases "x \ S") case True - have "openin (subtopology euclidean (connected_component_set S x)) (path_component_set S x)" + have "openin (top_of_set (connected_component_set S x)) (path_component_set S x)" apply (rule openin_subset_trans [of S]) apply (intro conjI openin_path_component_locally_path_connected [OF assms]) using path_component_subset_connected_component apply (auto simp: connected_component_subset) done - moreover have "closedin (subtopology euclidean (connected_component_set S x)) (path_component_set S x)" + moreover have "closedin (top_of_set (connected_component_set S x)) (path_component_set S x)" apply (rule closedin_subset_trans [of S]) apply (intro conjI closedin_path_component_locally_path_connected [OF assms]) using path_component_subset_connected_component apply (auto simp: connected_component_subset) @@ -2710,26 +2710,26 @@ proposition locally_connected_quotient_image: assumes lcS: "locally connected S" and oo: "\T. T \ f ` S - \ openin (subtopology euclidean S) (S \ f -` T) \ - openin (subtopology euclidean (f ` S)) T" + \ openin (top_of_set S) (S \ f -` T) \ + openin (top_of_set (f ` S)) T" shows "locally connected (f ` S)" proof (clarsimp simp: locally_connected_open_component) fix U C - assume opefSU: "openin (subtopology euclidean (f ` S)) U" and "C \ components U" + assume opefSU: "openin (top_of_set (f ` S)) U" and "C \ components U" then have "C \ U" "U \ f ` S" by (meson in_components_subset openin_imp_subset)+ - then have "openin (subtopology euclidean (f ` S)) C \ - openin (subtopology euclidean S) (S \ f -` C)" + then have "openin (top_of_set (f ` S)) C \ + openin (top_of_set S) (S \ f -` C)" by (auto simp: oo) - moreover have "openin (subtopology euclidean S) (S \ f -` C)" + moreover have "openin (top_of_set S) (S \ f -` C)" proof (subst openin_subopen, clarify) fix x assume "x \ S" "f x \ C" - show "\T. openin (subtopology euclidean S) T \ x \ T \ T \ (S \ f -` C)" + show "\T. openin (top_of_set S) T \ x \ T \ T \ (S \ f -` C)" proof (intro conjI exI) - show "openin (subtopology euclidean S) (connected_component_set (S \ f -` U) x)" + show "openin (top_of_set S) (connected_component_set (S \ f -` U) x)" proof (rule ccontr) - assume **: "\ openin (subtopology euclidean S) (connected_component_set (S \ f -` U) x)" + assume **: "\ openin (top_of_set S) (connected_component_set (S \ f -` U) x)" then have "x \ (S \ f -` U)" using \U \ f ` S\ opefSU lcS locally_connected_2 oo by blast with ** show False @@ -2768,7 +2768,7 @@ finally show "connected_component_set (S \ f -` U) x \ (S \ f -` C)" . qed qed - ultimately show "openin (subtopology euclidean (f ` S)) C" + ultimately show "openin (top_of_set (f ` S)) C" by metis qed @@ -2776,32 +2776,32 @@ proposition locally_path_connected_quotient_image: assumes lcS: "locally path_connected S" and oo: "\T. T \ f ` S - \ openin (subtopology euclidean S) (S \ f -` T) \ openin (subtopology euclidean (f ` S)) T" + \ openin (top_of_set S) (S \ f -` T) \ openin (top_of_set (f ` S)) T" shows "locally path_connected (f ` S)" proof (clarsimp simp: locally_path_connected_open_path_component) fix U y - assume opefSU: "openin (subtopology euclidean (f ` S)) U" and "y \ U" + assume opefSU: "openin (top_of_set (f ` S)) U" and "y \ U" then have "path_component_set U y \ U" "U \ f ` S" by (meson path_component_subset openin_imp_subset)+ - then have "openin (subtopology euclidean (f ` S)) (path_component_set U y) \ - openin (subtopology euclidean S) (S \ f -` path_component_set U y)" + then have "openin (top_of_set (f ` S)) (path_component_set U y) \ + openin (top_of_set S) (S \ f -` path_component_set U y)" proof - have "path_component_set U y \ f ` S" using \U \ f ` S\ \path_component_set U y \ U\ by blast then show ?thesis using oo by blast qed - moreover have "openin (subtopology euclidean S) (S \ f -` path_component_set U y)" + moreover have "openin (top_of_set S) (S \ f -` path_component_set U y)" proof (subst openin_subopen, clarify) fix x assume "x \ S" and Uyfx: "path_component U y (f x)" then have "f x \ U" using path_component_mem by blast - show "\T. openin (subtopology euclidean S) T \ x \ T \ T \ (S \ f -` path_component_set U y)" + show "\T. openin (top_of_set S) T \ x \ T \ T \ (S \ f -` path_component_set U y)" proof (intro conjI exI) - show "openin (subtopology euclidean S) (path_component_set (S \ f -` U) x)" + show "openin (top_of_set S) (path_component_set (S \ f -` U) x)" proof (rule ccontr) - assume **: "\ openin (subtopology euclidean S) (path_component_set (S \ f -` U) x)" + assume **: "\ openin (top_of_set S) (path_component_set (S \ f -` U) x)" then have "x \ (S \ f -` U)" by (metis (no_types, lifting) \U \ f ` S\ opefSU lcS oo locally_path_connected_open_path_component) then show False @@ -2842,7 +2842,7 @@ finally show "path_component_set (S \ f -` U) x \ (S \ f -` path_component_set U y)" . qed qed - ultimately show "openin (subtopology euclidean (f ` S)) (path_component_set U y)" + ultimately show "openin (top_of_set (f ` S)) (path_component_set U y)" by metis qed @@ -2851,14 +2851,14 @@ lemma continuous_on_components_gen: fixes f :: "'a::topological_space \ 'b::topological_space" assumes "\c. c \ components S \ - openin (subtopology euclidean S) c \ continuous_on c f" + openin (top_of_set S) c \ continuous_on c f" shows "continuous_on S f" proof (clarsimp simp: continuous_openin_preimage_eq) fix t :: "'b set" assume "open t" have *: "S \ f -` t = (\c \ components S. c \ f -` t)" by auto - show "openin (subtopology euclidean S) (S \ f -` t)" + show "openin (top_of_set S) (S \ f -` t)" unfolding * using \open t\ assms continuous_openin_preimage_gen openin_trans openin_Union by blast qed @@ -2891,9 +2891,9 @@ lemma closedin_union_complement_components: assumes u: "locally connected u" - and S: "closedin (subtopology euclidean u) S" + and S: "closedin (top_of_set u) S" and cuS: "c \ components(u - S)" - shows "closedin (subtopology euclidean u) (S \ \c)" + shows "closedin (top_of_set u) (S \ \c)" proof - have di: "(\S t. S \ c \ t \ c' \ disjnt S t) \ disjnt (\ c) (\ c')" for c' by (simp add: disjnt_def) blast @@ -2906,13 +2906,13 @@ by (metis DiffD1 DiffD2 assms(3) components_nonoverlap disjnt_def subsetCE) ultimately have eq: "S \ \c = u - (\(components(u - S) - c))" by (auto simp: disjnt_def) - have *: "openin (subtopology euclidean u) (\(components (u - S) - c))" + have *: "openin (top_of_set u) (\(components (u - S) - c))" apply (rule openin_Union) apply (rule openin_trans [of "u - S"]) apply (simp add: u S locally_diff_closed openin_components_locally_connected) apply (simp add: openin_diff S) done - have "openin (subtopology euclidean u) (u - (u - \(components (u - S) - c)))" + have "openin (top_of_set u) (u - (u - \(components (u - S) - c)))" apply (rule openin_diff, simp) apply (metis closedin_diff closedin_topspace topspace_euclidean_subtopology *) done @@ -2925,7 +2925,7 @@ assumes S: "closed S" and c: "c \ components(- S)" shows "closed(S \ \ c)" proof - - have "closedin (subtopology euclidean UNIV) (S \ \c)" + have "closedin (top_of_set UNIV) (S \ \c)" apply (rule closedin_union_complement_components [OF locally_connected_UNIV]) using S c apply (simp_all add: Compl_eq_Diff_UNIV) done @@ -2935,11 +2935,11 @@ lemma closedin_Un_complement_component: fixes S :: "'a::real_normed_vector set" assumes u: "locally connected u" - and S: "closedin (subtopology euclidean u) S" + and S: "closedin (top_of_set u) S" and c: " c \ components(u - S)" - shows "closedin (subtopology euclidean u) (S \ c)" + shows "closedin (top_of_set u) (S \ c)" proof - - have "closedin (subtopology euclidean u) (S \ \{c})" + have "closedin (top_of_set u) (S \ \{c})" using c by (blast intro: closedin_union_complement_components [OF u S]) then show ?thesis by simp @@ -3987,7 +3987,7 @@ proposition path_connected_openin_diff_countable: fixes S :: "'a::euclidean_space set" - assumes "connected S" and ope: "openin (subtopology euclidean (affine hull S)) S" + assumes "connected S" and ope: "openin (top_of_set (affine hull S)) S" and "\ collinear S" "countable T" shows "path_connected(S - T)" proof (clarsimp simp add: path_connected_component) @@ -3995,9 +3995,9 @@ assume xy: "x \ S" "x \ T" "y \ S" "y \ T" show "path_component (S - T) x y" proof (rule connected_equivalence_relation_gen [OF \connected S\, where P = "\x. x \ T"]) - show "\z. z \ U \ z \ T" if opeU: "openin (subtopology euclidean S) U" and "x \ U" for U x + show "\z. z \ U \ z \ T" if opeU: "openin (top_of_set S) U" and "x \ U" for U x proof - - have "openin (subtopology euclidean (affine hull S)) U" + have "openin (top_of_set (affine hull S)) U" using opeU ope openin_trans by blast with \x \ U\ obtain r where Usub: "U \ affine hull S" and "r > 0" and subU: "ball x r \ affine hull S \ U" @@ -4024,7 +4024,7 @@ using \countable T\ countable_subset by blast then show ?thesis by blast qed - show "\U. openin (subtopology euclidean S) U \ x \ U \ + show "\U. openin (top_of_set S) U \ x \ U \ (\x\U. \y\U. x \ T \ y \ T \ path_component (S - T) x y)" if "x \ S" for x proof - @@ -4046,9 +4046,9 @@ proof (intro exI conjI) show "x \ ball x r \ affine hull S" using \x \ S\ \r > 0\ by (simp add: hull_inc) - have "openin (subtopology euclidean (affine hull S)) (ball x r \ affine hull S)" + have "openin (top_of_set (affine hull S)) (ball x r \ affine hull S)" by (subst inf.commute) (simp add: openin_Int_open) - then show "openin (subtopology euclidean S) (ball x r \ affine hull S)" + then show "openin (top_of_set S) (ball x r \ affine hull S)" by (rule openin_subset_trans [OF _ subS Ssub]) qed (use * path_component_trans in \auto simp: path_connected_component path_component_of_subset [OF ST]\) qed @@ -4057,7 +4057,7 @@ corollary connected_openin_diff_countable: fixes S :: "'a::euclidean_space set" - assumes "connected S" and ope: "openin (subtopology euclidean (affine hull S)) S" + assumes "connected S" and ope: "openin (top_of_set (affine hull S)) S" and "\ collinear S" "countable T" shows "connected(S - T)" by (metis path_connected_imp_connected path_connected_openin_diff_countable [OF assms]) @@ -4074,7 +4074,7 @@ case False show ?thesis proof (rule path_connected_openin_diff_countable) - show "openin (subtopology euclidean (affine hull S)) S" + show "openin (top_of_set (affine hull S)) S" by (simp add: assms hull_subset open_subset) show "\ collinear S" using assms False by (simp add: collinear_aff_dim aff_dim_open) @@ -4329,7 +4329,7 @@ proposition%unimportant homeomorphism_moving_point: fixes a :: "'a::euclidean_space" - assumes ope: "openin (subtopology euclidean (affine hull S)) S" + assumes ope: "openin (top_of_set (affine hull S)) S" and "S \ T" and TS: "T \ affine hull S" and S: "connected S" "a \ S" "b \ S" @@ -4371,7 +4371,7 @@ then show "bounded {x. \ ((f2 \ f1) x = x \ (g1 \ g2) x = x)}" by (rule bounded_subset) auto qed - have 3: "\U. openin (subtopology euclidean S) U \ + have 3: "\U. openin (top_of_set S) U \ d \ U \ (\x\U. \f g. homeomorphism T T f g \ f d = x \ @@ -4410,7 +4410,7 @@ assumes K: "finite K" "\i. i \ K \ x i \ S \ y i \ S" "pairwise (\i j. (x i \ x j) \ (y i \ y j)) K" and "2 \ aff_dim S" - and ope: "openin (subtopology euclidean (affine hull S)) S" + and ope: "openin (top_of_set (affine hull S)) S" and "S \ T" "T \ affine hull S" "connected S" shows "\f g. homeomorphism T T f g \ (\i \ K. f(x i) = y i) \ {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" @@ -4449,7 +4449,7 @@ and hk_sub: "{x. \ (h x = x \ k x = x)} \ S - y ` K" and bo_hk: "bounded {x. \ (h x = x \ k x = x)}" proof (rule homeomorphism_moving_point [of "S - y`K" T "f(x i)" "y i"]) - show "openin (subtopology euclidean (affine hull (S - y ` K))) (S - y ` K)" + show "openin (top_of_set (affine hull (S - y ` K))) (S - y ` K)" by (simp add: aff_eq openin_diff finite_imp_closedin image_subset_iff hull_inc insert xyS) show "S - y ` K \ T" using \S \ T\ by auto @@ -4499,7 +4499,7 @@ case False then have affS: "affine hull S = UNIV" by (simp add: affine_hull_open \open S\) - then have ope: "openin (subtopology euclidean (affine hull S)) S" + then have ope: "openin (top_of_set (affine hull S)) S" using \open S\ open_openin by auto have "2 \ DIM('a)" by (rule 2) also have "\ = aff_dim (UNIV :: 'a set)" @@ -4714,7 +4714,7 @@ using \cbox w z \ S\ \S \ T\ by auto show "homeomorphism T T f' g'" proof - have clo: "closedin (subtopology euclidean (cbox w z \ (T - box w z))) (T - box w z)" + have clo: "closedin (top_of_set (cbox w z \ (T - box w z))) (T - box w z)" by (metis Diff_Diff_Int Diff_subset T closedin_def open_box openin_open_Int topspace_euclidean_subtopology) have "continuous_on (cbox w z \ (T - box w z)) f'" "continuous_on (cbox w z \ (T - box w z)) g'" unfolding f'_def g'_def @@ -4830,14 +4830,14 @@ proposition%unimportant homeomorphism_grouping_points_exists_gen: fixes S :: "'a::euclidean_space set" - assumes opeU: "openin (subtopology euclidean S) U" - and opeS: "openin (subtopology euclidean (affine hull S)) S" + assumes opeU: "openin (top_of_set S) U" + and opeS: "openin (top_of_set (affine hull S)) S" and "U \ {}" "finite K" "K \ S" and S: "S \ T" "T \ affine hull S" "connected S" obtains f g where "homeomorphism T T f g" "{x. (\ (f x = x \ g x = x))} \ S" "bounded {x. (\ (f x = x \ g x = x))}" "\x. x \ K \ f x \ U" proof (cases "2 \ aff_dim S") case True - have opeU': "openin (subtopology euclidean (affine hull S)) U" + have opeU': "openin (top_of_set (affine hull S)) U" using opeS opeU openin_trans by blast obtain u where "u \ U" "u \ S" using \U \ {}\ opeU openin_imp_subset by fastforce+ @@ -4892,7 +4892,7 @@ by (meson Topological_Spaces.connected_continuous_image \connected S\ homeomorphism_cont1 homeomorphism_of_subsets homhj hull_subset top_greatest) have hUS: "h ` U \ h ` S" by (meson homeomorphism_imp_open_map homeomorphism_of_subsets homhj hull_subset opeS opeU open_UNIV openin_open_eq) - have opn: "openin (subtopology euclidean (affine hull S)) U \ open (h ` U)" for U + have opn: "openin (top_of_set (affine hull S)) U \ open (h ` U)" for U using homeomorphism_imp_open_map [OF homhj] by simp have "open (h ` U)" "open (h ` S)" by (auto intro: opeS opeU openin_trans opn) diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Jordan_Curve.thy --- a/src/HOL/Analysis/Jordan_Curve.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Jordan_Curve.thy Tue Mar 19 16:14:51 2019 +0000 @@ -18,7 +18,7 @@ proof - have [simp]: "a \ S" "a \ T" "b \ S" "b \ T" by (meson ComplD ccS ccT connected_component_in)+ - have clo: "closedin (subtopology euclidean (S \ T)) S" "closedin (subtopology euclidean (S \ T)) T" + have clo: "closedin (top_of_set (S \ T)) S" "closedin (top_of_set (S \ T)) T" by (simp_all add: assms closed_subset compact_imp_closed) obtain g where contg: "continuous_on S g" and g: "\x. x \ S \ exp (\* of_real (g x)) = (x - a) /\<^sub>R cmod (x - a) / ((x - b) /\<^sub>R cmod (x - b))" @@ -276,8 +276,8 @@ and "arc g" "pathstart g = a" "pathfinish g = b" and pag_sub: "path_image g - {a,b} \ middle" proof (rule dense_accessible_frontier_point_pairs [OF \open middle\ \connected middle\, of "path_image c \ ball a0 (dist a0 b0)" "path_image c \ ball b0 (dist a0 b0)"]) - show "openin (subtopology euclidean (frontier middle)) (path_image c \ ball a0 (dist a0 b0))" - "openin (subtopology euclidean (frontier middle)) (path_image c \ ball b0 (dist a0 b0))" + show "openin (top_of_set (frontier middle)) (path_image c \ ball a0 (dist a0 b0))" + "openin (top_of_set (frontier middle)) (path_image c \ ball b0 (dist a0 b0))" by (simp_all add: \frontier middle = path_image c\ openin_open_Int) show "path_image c \ ball a0 (dist a0 b0) \ path_image c \ ball b0 (dist a0 b0)" using \a0 \ b0\ \b0 \ path_image c\ by auto diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Lebesgue_Measure.thy --- a/src/HOL/Analysis/Lebesgue_Measure.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Lebesgue_Measure.thy Tue Mar 19 16:14:51 2019 +0000 @@ -1341,11 +1341,11 @@ qed lemma lebesgue_openin: - "\openin (subtopology euclidean S) T; S \ sets lebesgue\ \ T \ sets lebesgue" + "\openin (top_of_set S) T; S \ sets lebesgue\ \ T \ sets lebesgue" by (metis borel_open openin_open sets.Int sets_completionI_sets sets_lborel) lemma lebesgue_closedin: - "\closedin (subtopology euclidean S) T; S \ sets lebesgue\ \ T \ sets lebesgue" + "\closedin (top_of_set S) T; S \ sets lebesgue\ \ T \ sets lebesgue" by (metis borel_closed closedin_closed sets.Int sets_completionI_sets sets_lborel) end diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Polytope.thy --- a/src/HOL/Analysis/Polytope.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Polytope.thy Tue Mar 19 16:14:51 2019 +0000 @@ -3921,9 +3921,9 @@ 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 (subtopology euclidean U) (U-C)" + 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 (subtopology euclidean U) T \ U \ closure T" + 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 diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Product_Topology.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Product_Topology.thy Tue Mar 19 16:14:51 2019 +0000 @@ -0,0 +1,459 @@ +theory Product_Topology +imports Function_Topology +begin + +lemma subset_UnE: (*FIXME MOVE*) + assumes "C \ A \ B" + obtains A' B' where "A' \ A" "B' \ B" "C = A' \ B'" + by (metis assms Int_Un_distrib inf.order_iff inf_le2) + +section \Product Topology\ + +subsection\A binary product topology where the two types can be different.\ + +definition prod_topology :: "'a topology \ 'b topology \ ('a \ 'b) topology" where + "prod_topology X Y \ topology (arbitrary union_of (\U. U \ {S \ T |S T. openin X S \ openin Y T}))" + +lemma open_product_open: + assumes "open A" + shows "\\. \ \ {S \ T |S T. open S \ open T} \ \ \ = A" +proof - + obtain f g where *: "\u. u \ A \ open (f u) \ open (g u) \ u \ (f u) \ (g u) \ (f u) \ (g u) \ A" + using open_prod_def [of A] assms by metis + let ?\ = "(\u. f u \ g u) ` A" + show ?thesis + by (rule_tac x="?\" in exI) (auto simp: dest: *) +qed + +lemma open_product_open_eq: "(arbitrary union_of (\U. \S T. U = S \ T \ open S \ open T)) = open" + by (force simp: union_of_def arbitrary_def intro: open_product_open open_Times) + +lemma openin_prod_topology: + "openin (prod_topology X Y) = arbitrary union_of (\U. U \ {S \ T |S T. openin X S \ openin Y T})" + unfolding prod_topology_def +proof (rule topology_inverse') + show "istopology (arbitrary union_of (\U. U \ {S \ T |S T. openin X S \ openin Y T}))" + apply (rule istopology_base, simp) + by (metis openin_Int times_Int_times) +qed + +lemma topspace_prod_topology [simp]: + "topspace (prod_topology X Y) = topspace X \ topspace Y" +proof - + have "topspace(prod_topology X Y) = \ (Collect (openin (prod_topology X Y)))" (is "_ = ?Z") + unfolding topspace_def .. + also have "\ = topspace X \ topspace Y" + proof + show "?Z \ topspace X \ topspace Y" + apply (auto simp: openin_prod_topology union_of_def arbitrary_def) + using openin_subset by force+ + next + have *: "\A B. topspace X \ topspace Y = A \ B \ openin X A \ openin Y B" + by blast + show "topspace X \ topspace Y \ ?Z" + apply (rule Union_upper) + using * by (simp add: openin_prod_topology arbitrary_union_of_inc) + qed + finally show ?thesis . +qed + +lemma subtopology_Times: + shows "subtopology (prod_topology X Y) (S \ T) = prod_topology (subtopology X S) (subtopology Y T)" +proof - + have "((\U. \S T. U = S \ T \ openin X S \ openin Y T) relative_to S \ T) = + (\U. \S' T'. U = S' \ T' \ (openin X relative_to S) S' \ (openin Y relative_to T) T')" + by (auto simp: relative_to_def times_Int_times fun_eq_iff) metis + then show ?thesis + by (simp add: topology_eq openin_prod_topology arbitrary_union_of_relative_to flip: openin_relative_to) +qed + +lemma prod_topology_subtopology: + "prod_topology (subtopology X S) Y = subtopology (prod_topology X Y) (S \ topspace Y)" + "prod_topology X (subtopology Y T) = subtopology (prod_topology X Y) (topspace X \ T)" +by (auto simp: subtopology_Times) + +lemma prod_topology_discrete_topology: + "discrete_topology (S \ T) = prod_topology (discrete_topology S) (discrete_topology T)" + by (auto simp: discrete_topology_unique openin_prod_topology intro: arbitrary_union_of_inc) + +lemma prod_topology_euclidean [simp]: "prod_topology euclidean euclidean = euclidean" + by (simp add: prod_topology_def open_product_open_eq) + +lemma prod_topology_subtopology_eu [simp]: + "prod_topology (subtopology euclidean S) (subtopology euclidean T) = subtopology euclidean (S \ T)" + by (simp add: prod_topology_subtopology subtopology_subtopology times_Int_times) + +lemma continuous_map_subtopology_eu [simp]: + "continuous_map (subtopology euclidean 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 openin_prod_topology_alt: + "openin (prod_topology X Y) S \ + (\x y. (x,y) \ S \ (\U V. openin X U \ openin Y V \ x \ U \ y \ V \ U \ V \ S))" + apply (auto simp: openin_prod_topology arbitrary_union_of_alt, fastforce) + by (metis mem_Sigma_iff) + +lemma open_map_fst: "open_map (prod_topology X Y) X fst" + unfolding open_map_def openin_prod_topology_alt + by (force simp: openin_subopen [of X "fst ` _"] intro: subset_fst_imageI) + +lemma open_map_snd: "open_map (prod_topology X Y) Y snd" + unfolding open_map_def openin_prod_topology_alt + by (force simp: openin_subopen [of Y "snd ` _"] intro: subset_snd_imageI) + +lemma openin_Times: + "openin (prod_topology X Y) (S \ T) \ S = {} \ T = {} \ openin X S \ openin Y T" +proof (cases "S = {} \ T = {}") + case False + then show ?thesis + apply (simp add: openin_prod_topology_alt openin_subopen [of X S] openin_subopen [of Y T] times_subset_iff, safe) + apply (meson|force)+ + done +qed force + + +lemma closure_of_Times: + "(prod_topology X Y) closure_of (S \ T) = (X closure_of S) \ (Y closure_of T)" (is "?lhs = ?rhs") +proof + show "?lhs \ ?rhs" + by (clarsimp simp: closure_of_def openin_prod_topology_alt) blast + show "?rhs \ ?lhs" + by (clarsimp simp: closure_of_def openin_prod_topology_alt) (meson SigmaI subsetD) +qed + +lemma closedin_Times: + "closedin (prod_topology X Y) (S \ T) \ S = {} \ T = {} \ closedin X S \ closedin Y T" + by (auto simp: closure_of_Times times_eq_iff simp flip: closure_of_eq) + +lemma continuous_map_pairwise: + "continuous_map Z (prod_topology X Y) f \ continuous_map Z X (fst \ f) \ continuous_map Z Y (snd \ f)" + (is "?lhs = ?rhs") +proof - + let ?g = "fst \ f" and ?h = "snd \ f" + have f: "f x = (?g x, ?h x)" for x + by auto + show ?thesis + proof (cases "(\x \ topspace Z. ?g x \ topspace X) \ (\x \ topspace Z. ?h x \ topspace Y)") + case True + show ?thesis + proof safe + assume "continuous_map Z (prod_topology X Y) f" + then have "openin Z {x \ topspace Z. fst (f x) \ U}" if "openin X U" for U + unfolding continuous_map_def using True that + apply clarify + apply (drule_tac x="U \ topspace Y" in spec) + by (simp add: openin_Times mem_Times_iff cong: conj_cong) + with True show "continuous_map Z X (fst \ f)" + by (auto simp: continuous_map_def) + next + assume "continuous_map Z (prod_topology X Y) f" + then have "openin Z {x \ topspace Z. snd (f x) \ V}" if "openin Y V" for V + unfolding continuous_map_def using True that + apply clarify + apply (drule_tac x="topspace X \ V" in spec) + by (simp add: openin_Times mem_Times_iff cong: conj_cong) + with True show "continuous_map Z Y (snd \ f)" + by (auto simp: continuous_map_def) + next + assume Z: "continuous_map Z X (fst \ f)" "continuous_map Z Y (snd \ f)" + have *: "openin Z {x \ topspace Z. f x \ W}" + if "\w. w \ W \ \U V. openin X U \ openin Y V \ w \ U \ V \ U \ V \ W" for W + proof (subst openin_subopen, clarify) + fix x :: "'a" + assume "x \ topspace Z" and "f x \ W" + with that [OF \f x \ W\] + obtain U V where UV: "openin X U" "openin Y V" "f x \ U \ V" "U \ V \ W" + by auto + with Z UV show "\T. openin Z T \ x \ T \ T \ {x \ topspace Z. f x \ W}" + apply (rule_tac x="{x \ topspace Z. ?g x \ U} \ {x \ topspace Z. ?h x \ V}" in exI) + apply (auto simp: \x \ topspace Z\ continuous_map_def) + done + qed + show "continuous_map Z (prod_topology X Y) f" + using True by (simp add: continuous_map_def openin_prod_topology_alt mem_Times_iff *) + qed + qed (auto simp: continuous_map_def) +qed + +lemma continuous_map_paired: + "continuous_map Z (prod_topology X Y) (\x. (f x,g x)) \ continuous_map Z X f \ continuous_map Z Y g" + by (simp add: continuous_map_pairwise o_def) + +lemma continuous_map_pairedI [continuous_intros]: + "\continuous_map Z X f; continuous_map Z Y g\ \ continuous_map Z (prod_topology X Y) (\x. (f x,g x))" + by (simp add: continuous_map_pairwise o_def) + +lemma continuous_map_fst [continuous_intros]: "continuous_map (prod_topology X Y) X fst" + using continuous_map_pairwise [of "prod_topology X Y" X Y id] + by (simp add: continuous_map_pairwise) + +lemma continuous_map_snd [continuous_intros]: "continuous_map (prod_topology X Y) Y snd" + using continuous_map_pairwise [of "prod_topology X Y" X Y id] + by (simp add: continuous_map_pairwise) + +lemma continuous_map_fst_of [continuous_intros]: + "continuous_map Z (prod_topology X Y) f \ continuous_map Z X (fst \ f)" + by (simp add: continuous_map_pairwise) + +lemma continuous_map_snd_of [continuous_intros]: + "continuous_map Z (prod_topology X Y) f \ continuous_map Z Y (snd \ f)" + by (simp add: continuous_map_pairwise) + +lemma continuous_map_if_iff [simp]: "continuous_map X Y (\x. if P then f x else g x) \ continuous_map X Y (if P then f else g)" + by simp + +lemma continuous_map_if [continuous_intros]: "\P \ continuous_map X Y f; ~P \ continuous_map X Y g\ + \ continuous_map X Y (\x. if P then f x else g x)" + by simp + +lemma continuous_map_subtopology_fst [continuous_intros]: "continuous_map (subtopology (prod_topology X Y) Z) X fst" + using continuous_map_from_subtopology continuous_map_fst by force + +lemma continuous_map_subtopology_snd [continuous_intros]: "continuous_map (subtopology (prod_topology X Y) Z) Y snd" + using continuous_map_from_subtopology continuous_map_snd by force + +lemma quotient_map_fst [simp]: + "quotient_map(prod_topology X Y) X fst \ (topspace Y = {} \ topspace X = {})" + by (auto simp: continuous_open_quotient_map open_map_fst continuous_map_fst) + +lemma quotient_map_snd [simp]: + "quotient_map(prod_topology X Y) Y snd \ (topspace X = {} \ topspace Y = {})" + by (auto simp: continuous_open_quotient_map open_map_snd continuous_map_snd) + +lemma retraction_map_fst: + "retraction_map (prod_topology X Y) X fst \ (topspace Y = {} \ topspace X = {})" +proof (cases "topspace Y = {}") + case True + then show ?thesis + using continuous_map_image_subset_topspace + by (fastforce simp: retraction_map_def retraction_maps_def continuous_map_fst continuous_map_on_empty) +next + case False + have "\g. continuous_map X (prod_topology X Y) g \ (\x\topspace X. fst (g x) = x)" + if y: "y \ topspace Y" for y + by (rule_tac x="\x. (x,y)" in exI) (auto simp: y continuous_map_paired) + with False have "retraction_map (prod_topology X Y) X fst" + by (fastforce simp: retraction_map_def retraction_maps_def continuous_map_fst) + with False show ?thesis + by simp +qed + +lemma retraction_map_snd: + "retraction_map (prod_topology X Y) Y snd \ (topspace X = {} \ topspace Y = {})" +proof (cases "topspace X = {}") + case True + then show ?thesis + using continuous_map_image_subset_topspace + by (fastforce simp: retraction_map_def retraction_maps_def continuous_map_fst continuous_map_on_empty) +next + case False + have "\g. continuous_map Y (prod_topology X Y) g \ (\y\topspace Y. snd (g y) = y)" + if x: "x \ topspace X" for x + by (rule_tac x="\y. (x,y)" in exI) (auto simp: x continuous_map_paired) + with False have "retraction_map (prod_topology X Y) Y snd" + by (fastforce simp: retraction_map_def retraction_maps_def continuous_map_snd) + with False show ?thesis + by simp +qed + + +lemma continuous_map_of_fst: + "continuous_map (prod_topology X Y) Z (f \ fst) \ topspace Y = {} \ continuous_map X Z f" +proof (cases "topspace Y = {}") + case True + then show ?thesis + by (simp add: continuous_map_on_empty) +next + case False + then show ?thesis + by (simp add: continuous_compose_quotient_map_eq quotient_map_fst) +qed + +lemma continuous_map_of_snd: + "continuous_map (prod_topology X Y) Z (f \ snd) \ topspace X = {} \ continuous_map Y Z f" +proof (cases "topspace X = {}") + case True + then show ?thesis + by (simp add: continuous_map_on_empty) +next + case False + then show ?thesis + by (simp add: continuous_compose_quotient_map_eq quotient_map_snd) +qed + +lemma continuous_map_prod_top: + "continuous_map (prod_topology X Y) (prod_topology X' Y') (\(x,y). (f x, g y)) \ + topspace (prod_topology X Y) = {} \ continuous_map X X' f \ continuous_map Y Y' g" +proof (cases "topspace (prod_topology X Y) = {}") + case True + then show ?thesis + by (simp add: continuous_map_on_empty) +next + case False + then show ?thesis + by (simp add: continuous_map_paired case_prod_unfold continuous_map_of_fst [unfolded o_def] continuous_map_of_snd [unfolded o_def]) +qed + +lemma homeomorphic_maps_prod: + "homeomorphic_maps (prod_topology X Y) (prod_topology X' Y') (\(x,y). (f x, g y)) (\(x,y). (f' x, g' y)) \ + topspace(prod_topology X Y) = {} \ + topspace(prod_topology X' Y') = {} \ + homeomorphic_maps X X' f f' \ + homeomorphic_maps Y Y' g g'" + unfolding homeomorphic_maps_def continuous_map_prod_top + by (auto simp: continuous_map_def homeomorphic_maps_def continuous_map_prod_top) + +lemma embedding_map_graph: + "embedding_map X (prod_topology X Y) (\x. (x, f x)) \ continuous_map X Y f" + (is "?lhs = ?rhs") +proof + assume L: ?lhs + have "snd \ (\x. (x, f x)) = f" + by force + moreover have "continuous_map X Y (snd \ (\x. (x, f x)))" + using L + unfolding embedding_map_def + by (meson continuous_map_in_subtopology continuous_map_snd_of homeomorphic_imp_continuous_map) + ultimately show ?rhs + by simp +next + assume R: ?rhs + then show ?lhs + unfolding homeomorphic_map_maps embedding_map_def homeomorphic_maps_def + by (rule_tac x=fst in exI) + (auto simp: continuous_map_in_subtopology continuous_map_paired continuous_map_from_subtopology + continuous_map_fst topspace_subtopology) +qed + +lemma in_prod_topology_closure_of: + assumes "z \ (prod_topology X Y) closure_of S" + shows "fst z \ X closure_of (fst ` S)" "snd z \ Y closure_of (snd ` S)" + using assms continuous_map_eq_image_closure_subset continuous_map_fst apply fastforce + using assms continuous_map_eq_image_closure_subset continuous_map_snd apply fastforce + done + + +proposition compact_space_prod_topology: + "compact_space(prod_topology X Y) \ topspace(prod_topology X Y) = {} \ compact_space X \ compact_space Y" +proof (cases "topspace(prod_topology X Y) = {}") + case True + then show ?thesis + using compact_space_topspace_empty by blast +next + case False + then have non_mt: "topspace X \ {}" "topspace Y \ {}" + by auto + have "compact_space X" "compact_space Y" if "compact_space(prod_topology X Y)" + proof - + have "compactin X (fst ` (topspace X \ topspace Y))" + by (metis compact_space_def continuous_map_fst image_compactin that topspace_prod_topology) + moreover + have "compactin Y (snd ` (topspace X \ topspace Y))" + by (metis compact_space_def continuous_map_snd image_compactin that topspace_prod_topology) + ultimately show "compact_space X" "compact_space Y" + by (simp_all add: non_mt compact_space_def) + qed + moreover + define \ where "\ \ (\V. topspace X \ V) ` Collect (openin Y)" + define \ where "\ \ (\U. U \ topspace Y) ` Collect (openin X)" + have "compact_space(prod_topology X Y)" if "compact_space X" "compact_space Y" + proof (rule Alexander_subbase_alt) + show toptop: "topspace X \ topspace Y \ \(\ \ \)" + unfolding \_def \_def by auto + fix \ :: "('a \ 'b) set set" + assume \: "\ \ \ \ \" "topspace X \ topspace Y \ \\" + then obtain \' \' where XY: "\' \ \" "\' \ \" and \eq: "\ = \' \ \'" + using subset_UnE by metis + then have sub: "topspace X \ topspace Y \ \(\' \ \')" + using \ by simp + obtain \ \ where \: "\U. U \ \ \ openin X U" "\' = (\U. U \ topspace Y) ` \" + and \: "\V. V \ \ \ openin Y V" "\' = (\V. topspace X \ V) ` \" + using XY by (clarsimp simp add: \_def \_def subset_image_iff) (force simp add: subset_iff) + have "\\. finite \ \ \ \ \' \ \' \ topspace X \ topspace Y \ \ \" + proof - + have "topspace X \ \\ \ topspace Y \ \\" + using \ \ \ \eq by auto + then have *: "\\. finite \ \ + (\x \ \. x \ (\V. topspace X \ V) ` \ \ x \ (\U. U \ topspace Y) ` \) \ + (topspace X \ topspace Y \ \\)" + proof + assume "topspace X \ \\" + with \compact_space X\ \ obtain \ where "finite \" "\ \ \" "topspace X \ \\" + by (meson compact_space_alt) + with that show ?thesis + by (rule_tac x="(\D. D \ topspace Y) ` \" in exI) auto + next + assume "topspace Y \ \\" + with \compact_space Y\ \ obtain \ where "finite \" "\ \ \" "topspace Y \ \\" + by (meson compact_space_alt) + with that show ?thesis + by (rule_tac x="(\C. topspace X \ C) ` \" in exI) auto + qed + then show ?thesis + using that \ \ by blast + qed + then show "\\. finite \ \ \ \ \ \ topspace X \ topspace Y \ \ \" + using \ \eq by blast + next + have "(finite intersection_of (\x. x \ \ \ x \ \) relative_to topspace X \ topspace Y) + = (\U. \S T. U = S \ T \ openin X S \ openin Y T)" + (is "?lhs = ?rhs") + proof - + have "?rhs U" if "?lhs U" for U + proof - + have "topspace X \ topspace Y \ \ T \ {A \ B |A B. A \ Collect (openin X) \ B \ Collect (openin Y)}" + if "finite T" "T \ \ \ \" for T + using that + proof induction + case (insert B \) + then show ?case + unfolding \_def \_def + apply (simp add: Int_ac subset_eq image_def) + apply (metis (no_types) openin_Int openin_topspace times_Int_times) + done + qed auto + then show ?thesis + using that + by (auto simp: subset_eq elim!: relative_toE intersection_ofE) + qed + moreover + have "?lhs Z" if Z: "?rhs Z" for Z + proof - + obtain U V where "Z = U \ V" "openin X U" "openin Y V" + using Z by blast + then have UV: "U \ V = (topspace X \ topspace Y) \ (U \ V)" + by (simp add: Sigma_mono inf_absorb2 openin_subset) + moreover + have "?lhs ((topspace X \ topspace Y) \ (U \ V))" + proof (rule relative_to_inc) + show "(finite intersection_of (\x. x \ \ \ x \ \)) (U \ V)" + apply (simp add: intersection_of_def \_def \_def) + apply (rule_tac x="{(U \ topspace Y),(topspace X \ V)}" in exI) + using \openin X U\ \openin Y V\ openin_subset UV apply (fastforce simp add:) + done + qed + ultimately show ?thesis + using \Z = U \ V\ by auto + qed + ultimately show ?thesis + by meson + qed + then show "topology (arbitrary union_of (finite intersection_of (\x. x \ \ \ \) + relative_to (topspace X \ topspace Y))) = + prod_topology X Y" + by (simp add: prod_topology_def) + qed + ultimately show ?thesis + using False by blast +qed + + +lemma compactin_Times: + "compactin (prod_topology X Y) (S \ T) \ S = {} \ T = {} \ compactin X S \ compactin Y T" + by (auto simp: compactin_subspace subtopology_Times compact_space_prod_topology) + +end + diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Riemann_Mapping.thy --- a/src/HOL/Analysis/Riemann_Mapping.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Riemann_Mapping.thy Tue Mar 19 16:14:51 2019 +0000 @@ -1029,14 +1029,14 @@ by (metis frontierS C) obtain K where "C \ K" "compact K" and Ksub: "K \ \(range X)" and clo: "closed(\(range X) - K)" - proof (cases "{k. C \ k \ compact k \ openin (subtopology euclidean (\(range X))) k} = {}") + proof (cases "{k. C \ k \ compact k \ openin (top_of_set (\(range X))) k} = {}") case True then show ?thesis using Sura_Bura [OF lcX Cco \compact C\] boC by (simp add: True) next case False - then obtain L where "compact L" "C \ L" and K: "openin (subtopology euclidean (\x. X x)) L" + then obtain L where "compact L" "C \ L" and K: "openin (top_of_set (\x. X x)) L" by blast show ?thesis proof @@ -1089,9 +1089,9 @@ apply safe using DiffI J empty apply auto[1] using closure_subset by blast - then have "openin (subtopology euclidean (X j)) (X j \ closure U)" + then have "openin (top_of_set (X j)) (X j \ closure U)" by (simp add: openin_open_Int \open U\) - moreover have "closedin (subtopology euclidean (X j)) (X j \ closure U)" + moreover have "closedin (top_of_set (X j)) (X j \ closure U)" by (simp add: closedin_closed_Int) moreover have "X j \ closure U \ X j" by (metis unboundedX \compact (closure U)\ bounded_subset compact_eq_bounded_closed inf.order_iff) diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Starlike.thy --- a/src/HOL/Analysis/Starlike.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Starlike.thy Tue Mar 19 16:14:51 2019 +0000 @@ -8,7 +8,7 @@ chapter \Unsorted\ theory Starlike -imports Convex_Euclidean_Space +imports Convex_Euclidean_Space Abstract_Limits begin section \Line Segments\ @@ -1838,7 +1838,7 @@ assumes "convex S" shows "rel_interior (rel_interior S) = rel_interior S" proof - - have "openin (subtopology euclidean (affine hull (rel_interior S))) (rel_interior S)" + have "openin (top_of_set (affine hull (rel_interior S))) (rel_interior S)" using openin_rel_interior[of S] rel_interior_same_affine_hull[of S] assms by auto then show ?thesis using rel_interior_def by auto @@ -2083,7 +2083,7 @@ fixes S :: "'n::euclidean_space set" shows "closed (rel_frontier S)" proof - - have *: "closedin (subtopology euclidean (affine hull S)) (closure S - rel_interior S)" + have *: "closedin (top_of_set (affine hull S)) (closure S - rel_interior S)" by (simp add: closed_subset closedin_diff closure_affine_hull openin_rel_interior) show ?thesis apply (rule closedin_closed_trans[of "affine hull S" "rel_frontier S"]) @@ -4792,11 +4792,11 @@ lemma separation_normal_local: fixes S :: "'a::euclidean_space set" - assumes US: "closedin (subtopology euclidean U) S" - and UT: "closedin (subtopology euclidean U) T" + assumes US: "closedin (top_of_set U) S" + and UT: "closedin (top_of_set U) T" and "S \ T = {}" - obtains S' T' where "openin (subtopology euclidean U) S'" - "openin (subtopology euclidean U) T'" + obtains S' T' where "openin (top_of_set U) S'" + "openin (top_of_set U) T'" "S \ S'" "T \ T'" "S' \ T' = {}" proof (cases "S = {} \ T = {}") case True with that show ?thesis @@ -4810,10 +4810,10 @@ proof (rule_tac S' = "(U \ f -` {0<..})" and T' = "(U \ f -` {..<0})" in that) show "(U \ f -` {0<..}) \ (U \ f -` {..<0}) = {}" by auto - show "openin (subtopology euclidean U) (U \ f -` {0<..})" + show "openin (top_of_set U) (U \ f -` {0<..})" by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf) next - show "openin (subtopology euclidean U) (U \ f -` {..<0})" + show "openin (top_of_set U) (U \ f -` {..<0})" by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf) next have "S \ U" "T \ U" @@ -4993,10 +4993,10 @@ proposition proper_map: fixes f :: "'a::heine_borel \ 'b::heine_borel" - assumes "closedin (subtopology euclidean S) K" + assumes "closedin (top_of_set S) K" and com: "\U. \U \ T; compact U\ \ compact (S \ f -` U)" and "f ` S \ T" - shows "closedin (subtopology euclidean T) (f ` K)" + shows "closedin (top_of_set T) (f ` K)" proof - have "K \ S" using assms closedin_imp_subset by metis @@ -5221,7 +5221,7 @@ lemma proper_map_from_compact: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and imf: "f ` S \ T" and "compact S" - "closedin (subtopology euclidean T) K" + "closedin (top_of_set T) K" shows "compact (S \ f -` K)" by (rule closedin_compact [OF \compact S\] continuous_closedin_preimage_gen assms)+ @@ -5236,8 +5236,8 @@ lemma closed_map_fst: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" - assumes "compact T" "closedin (subtopology euclidean (S \ T)) c" - shows "closedin (subtopology euclidean S) (fst ` c)" + assumes "compact T" "closedin (top_of_set (S \ T)) c" + shows "closedin (top_of_set S) (fst ` c)" proof - have *: "fst ` (S \ T) \ S" by auto @@ -5256,8 +5256,8 @@ lemma closed_map_snd: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" - assumes "compact S" "closedin (subtopology euclidean (S \ T)) c" - shows "closedin (subtopology euclidean T) (snd ` c)" + assumes "compact S" "closedin (top_of_set (S \ T)) c" + shows "closedin (top_of_set T) (snd ` c)" proof - have *: "snd ` (S \ T) \ T" by auto @@ -5267,14 +5267,14 @@ lemma closedin_compact_projection: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" - assumes "compact S" and clo: "closedin (subtopology euclidean (S \ T)) U" - shows "closedin (subtopology euclidean T) {y. \x. x \ S \ (x, y) \ U}" + assumes "compact S" and clo: "closedin (top_of_set (S \ T)) U" + shows "closedin (top_of_set T) {y. \x. x \ S \ (x, y) \ U}" proof - have "U \ S \ T" by (metis clo closedin_imp_subset) then have "{y. \x. x \ S \ (x, y) \ U} = snd ` U" by force - moreover have "closedin (subtopology euclidean T) (snd ` U)" + moreover have "closedin (top_of_set T) (snd ` U)" by (rule closed_map_snd [OF assms]) ultimately show ?thesis by simp @@ -5366,14 +5366,14 @@ corollary affine_hull_convex_Int_openin: fixes S :: "'a::real_normed_vector set" - assumes "convex S" "openin (subtopology euclidean (affine hull S)) T" "S \ T \ {}" + assumes "convex S" "openin (top_of_set (affine hull S)) T" "S \ T \ {}" shows "affine hull (S \ T) = affine hull S" using assms unfolding openin_open by (metis affine_hull_convex_Int_open hull_subset inf.orderE inf_assoc) corollary affine_hull_openin: fixes S :: "'a::real_normed_vector set" - assumes "openin (subtopology euclidean (affine hull T)) S" "S \ {}" + assumes "openin (top_of_set (affine hull T)) S" "S \ {}" shows "affine hull S = affine hull T" using assms unfolding openin_open by (metis affine_affine_hull affine_hull_affine_Int_open hull_hull) @@ -5396,10 +5396,10 @@ lemma affine_hull_Diff: fixes S:: "'a::real_normed_vector set" - assumes ope: "openin (subtopology euclidean (affine hull S)) S" and "finite F" "F \ S" + assumes ope: "openin (top_of_set (affine hull S)) S" and "finite F" "F \ S" shows "affine hull (S - F) = affine hull S" proof - - have clo: "closedin (subtopology euclidean S) F" + have clo: "closedin (top_of_set S) F" using assms finite_imp_closedin by auto moreover have "S - F \ {}" using assms by auto @@ -6342,7 +6342,7 @@ lemma aff_dim_openin: fixes S :: "'a::euclidean_space set" - assumes ope: "openin (subtopology euclidean T) S" and "affine T" "S \ {}" + assumes ope: "openin (top_of_set T) S" and "affine T" "S \ {}" shows "aff_dim S = aff_dim T" proof - show ?thesis @@ -6394,7 +6394,7 @@ lemma dim_openin: fixes S :: "'a::euclidean_space set" - assumes ope: "openin (subtopology euclidean T) S" and "subspace T" "S \ {}" + assumes ope: "openin (top_of_set T) S" and "subspace T" "S \ {}" shows "dim S = dim T" proof (rule order_antisym) show "dim S \ dim T" @@ -6490,7 +6490,7 @@ corollary%unimportant dense_complement_openin_affine_hull: fixes S :: "'a :: euclidean_space set" assumes less: "aff_dim T < aff_dim S" - and ope: "openin (subtopology euclidean (affine hull S)) S" + and ope: "openin (top_of_set (affine hull S)) S" shows "closure(S - T) = closure S" proof - have "affine hull S - T \ affine hull S" @@ -6802,13 +6802,13 @@ corollary paracompact_closedin: fixes S :: "'a :: {metric_space,second_countable_topology} set" - assumes cin: "closedin (subtopology euclidean U) S" - and oin: "\T. T \ \ \ openin (subtopology euclidean U) T" + assumes cin: "closedin (top_of_set U) S" + and oin: "\T. T \ \ \ openin (top_of_set U) T" and "S \ \\" obtains \' where "S \ \ \'" - and "\V. V \ \' \ openin (subtopology euclidean U) V \ (\T. T \ \ \ V \ T)" + and "\V. V \ \' \ openin (top_of_set U) V \ (\T. T \ \ \ V \ T)" and "\x. x \ U - \ \V. openin (subtopology euclidean U) V \ x \ V \ + \ \V. openin (top_of_set U) V \ x \ V \ finite {X. X \ \' \ (X \ V \ {})}" proof - have "\Z. open Z \ (T = U \ Z)" if "T \ \" for T @@ -6830,12 +6830,12 @@ proof (rule_tac \' = "{U \ V |V. V \ \ \ (V \ K \ {})}" in that) show "S \ \?C" using \U \ K = S\ \U \ \\\ K by (blast dest!: subsetD) - show "\V. V \ ?C \ openin (subtopology euclidean U) V \ (\T. T \ \ \ V \ T)" + show "\V. V \ ?C \ openin (top_of_set U) V \ (\T. T \ \ \ V \ T)" using D1 intF by fastforce have *: "{X. (\V. X = U \ V \ V \ \ \ V \ K \ {}) \ X \ (U \ V) \ {}} \ (\x. U \ x) ` {U \ \. U \ V \ {}}" for V by blast - show "\V. openin (subtopology euclidean U) V \ x \ V \ finite {X \ ?C. X \ V \ {}}" + show "\V. openin (top_of_set U) V \ x \ V \ finite {X \ ?C. X \ V \ {}}" if "x \ U" for x using D2 [OF that] apply clarify @@ -6862,7 +6862,7 @@ lemma continuous_closed_graph_gen: fixes T :: "'b::real_normed_vector set" assumes contf: "continuous_on S f" and fim: "f ` S \ T" - shows "closedin (subtopology euclidean (S \ T)) ((\x. Pair x (f x)) ` S)" + shows "closedin (top_of_set (S \ T)) ((\x. Pair x (f x)) ` S)" proof - have eq: "((\x. Pair x (f x)) ` S) =(S \ T \ (\z. (f \ fst)z - snd z) -` {0})" using fim by auto @@ -6876,16 +6876,16 @@ fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact T" and fim: "f ` S \ T" shows "continuous_on S f \ - closedin (subtopology euclidean (S \ T)) ((\x. Pair x (f x)) ` S)" + closedin (top_of_set (S \ T)) ((\x. Pair x (f x)) ` S)" (is "?lhs = ?rhs") proof - have "?lhs" if ?rhs proof (clarsimp simp add: continuous_on_closed_gen [OF fim]) fix U - assume U: "closedin (subtopology euclidean T) U" + assume U: "closedin (top_of_set T) U" have eq: "(S \ f -` U) = fst ` (((\x. Pair x (f x)) ` S) \ (S \ U))" by (force simp: image_iff) - show "closedin (subtopology euclidean S) (S \ f -` U)" + show "closedin (top_of_set S) (S \ f -` U)" by (simp add: U closedin_Int closedin_Times closed_map_fst [OF \compact T\] that eq) qed with continuous_closed_graph_gen assms show ?thesis by blast @@ -6907,15 +6907,16 @@ by (auto intro: closed_subset simp: continuous_closed_graph_eq [OF \compact T\ fim]) lemma continuous_on_Un_local_open: - assumes opS: "openin (subtopology euclidean (S \ T)) S" - and opT: "openin (subtopology euclidean (S \ T)) T" + assumes opS: "openin (top_of_set (S \ T)) S" + and opT: "openin (top_of_set (S \ T)) T" and contf: "continuous_on S f" and contg: "continuous_on T f" shows "continuous_on (S \ T) f" -using pasting_lemma [of "{S,T}" "S \ T" "\i. i" "\i. f" f] contf contg opS opT by blast + using pasting_lemma [of "{S,T}" "top_of_set (S \ T)" id euclidean "\i. f" f] contf contg opS opT + by (simp add: subtopology_subtopology) (metis inf.absorb2 openin_imp_subset) lemma continuous_on_cases_local_open: - assumes opS: "openin (subtopology euclidean (S \ T)) S" - and opT: "openin (subtopology euclidean (S \ T)) T" + assumes opS: "openin (top_of_set (S \ T)) S" + and opT: "openin (top_of_set (S \ T)) T" and contf: "continuous_on S f" and contg: "continuous_on T g" and fg: "\x. x \ S \ \P x \ x \ T \ P x \ f x = g x" shows "continuous_on (S \ T) (\x. if P x then f x else g x)" @@ -6927,7 +6928,49 @@ then show ?thesis by (rule continuous_on_Un_local_open [OF opS opT]) qed - + +lemma continuous_map_cases_le: + assumes contp: "continuous_map X euclideanreal p" + and contq: "continuous_map X euclideanreal q" + and contf: "continuous_map (subtopology X {x. x \ topspace X \ p x \ q x}) Y f" + and contg: "continuous_map (subtopology X {x. x \ topspace X \ q x \ p x}) Y g" + and fg: "\x. \x \ topspace X; p x = q x\ \ f x = g x" + shows "continuous_map X Y (\x. if p x \ q x then f x else g x)" +proof - + have "continuous_map X Y (\x. if q x - p x \ {0..} then f x else g x)" + proof (rule continuous_map_cases_function) + show "continuous_map X euclideanreal (\x. q x - p x)" + by (intro contp contq continuous_intros) + show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of {0..}}) Y f" + by (simp add: contf) + show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of (topspace euclideanreal - {0..})}) Y g" + by (simp add: contg flip: Compl_eq_Diff_UNIV) + qed (auto simp: fg) + then show ?thesis + by simp +qed + +lemma continuous_map_cases_lt: + assumes contp: "continuous_map X euclideanreal p" + and contq: "continuous_map X euclideanreal q" + and contf: "continuous_map (subtopology X {x. x \ topspace X \ p x \ q x}) Y f" + and contg: "continuous_map (subtopology X {x. x \ topspace X \ q x \ p x}) Y g" + and fg: "\x. \x \ topspace X; p x = q x\ \ f x = g x" + shows "continuous_map X Y (\x. if p x < q x then f x else g x)" +proof - + have "continuous_map X Y (\x. if q x - p x \ {0<..} then f x else g x)" + proof (rule continuous_map_cases_function) + show "continuous_map X euclideanreal (\x. q x - p x)" + by (intro contp contq continuous_intros) + show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of {0<..}}) Y f" + by (simp add: contf) + show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of (topspace euclideanreal - {0<..})}) Y g" + by (simp add: contg flip: Compl_eq_Diff_UNIV) + qed (auto simp: fg) + then show ?thesis + by simp +qed + subsection%unimportant\The union of two collinear segments is another segment\ proposition%unimportant in_convex_hull_exchange: diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Analysis/Topology_Euclidean_Space.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy Tue Mar 19 16:14:51 2019 +0000 @@ -1874,8 +1874,8 @@ obtain \ :: "'a set set" where "countable \" and "{} \ \" - and ope: "\C. C \ \ \ openin(subtopology euclidean S) C" - and if_ope: "\T. openin(subtopology euclidean S) T \ \\. \ \ \ \ T = \\" + and ope: "\C. C \ \ \ openin(top_of_set S) C" + and if_ope: "\T. openin(top_of_set S) T \ \\. \ \ \ \ T = \\" by (meson subset_second_countable) then obtain f where f: "\C. C \ \ \ f C \ C" by (metis equals0I) @@ -1889,7 +1889,7 @@ proof (clarsimp simp: closure_approachable) fix x and e::real assume "x \ S" "0 < e" - have "openin (subtopology euclidean S) (S \ ball x e)" + have "openin (top_of_set S) (S \ ball x e)" by (simp add: openin_Int_open) with if_ope obtain \ where \: "\ \ \" "S \ ball x e = \\" by meson @@ -2324,12 +2324,12 @@ by (simp add: setdist_eq_0_sing_1) lemma setdist_eq_0_closedin: - shows "\closedin (subtopology euclidean U) S; x \ U\ + shows "\closedin (top_of_set U) S; x \ U\ \ (setdist {x} S = 0 \ S = {} \ x \ S)" by (auto simp: closedin_limpt setdist_eq_0_sing_1 closure_def) lemma setdist_gt_0_closedin: - shows "\closedin (subtopology euclidean U) S; x \ U; S \ {}; x \ S\ + shows "\closedin (top_of_set U) S; x \ U; S \ {}; x \ S\ \ setdist {x} S > 0" using less_eq_real_def setdist_eq_0_closedin by fastforce diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Analysis/Vitali_Covering_Theorem.thy --- a/src/HOL/Analysis/Vitali_Covering_Theorem.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Analysis/Vitali_Covering_Theorem.thy Tue Mar 19 16:14:51 2019 +0000 @@ -542,10 +542,10 @@ next assume R [rule_format]: "\x \ S. \e > 0. ?Q x e" let ?\ = "measure lebesgue" - have "\U. openin (subtopology euclidean S) U \ z \ U \ negligible U" + have "\U. openin (top_of_set S) U \ z \ U \ negligible U" if "z \ S" for z proof (intro exI conjI) - show "openin (subtopology euclidean S) (S \ ball z 1)" + show "openin (top_of_set S) (S \ ball z 1)" by (simp add: openin_open_Int) show "z \ S \ ball z 1" using \z \ S\ by auto diff -r eddcc7c726f3 -r 4a9167f377b0 src/HOL/Product_Type.thy --- a/src/HOL/Product_Type.thy Mon Mar 18 15:35:34 2019 +0000 +++ b/src/HOL/Product_Type.thy Tue Mar 19 16:14:51 2019 +0000 @@ -1124,7 +1124,7 @@ lemma vimage_snd: "snd -` A = UNIV \ A" by auto -lemma insert_times_insert [simp]: +lemma insert_Times_insert [simp]: "insert a A \ insert b B = insert (a,b) (A \ insert b B \ insert a A \ B)" by blast