--- a/src/HOL/Analysis/Abstract_Topology.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Abstract_Topology.thy Tue Mar 19 16:14:59 2019 +0000
@@ -8,7 +8,6 @@
Complex_Main
"HOL-Library.Set_Idioms"
"HOL-Library.FuncSet"
- (* Path_Connected *)
begin
subsection \<open>General notion of a topology as a value\<close>
@@ -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 \<longleftrightarrow> 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\<open>The most basic facts about the usual topology and metric on R\<close>
@@ -445,51 +444,51 @@
subsection \<open>Basic "localization" results are handy for connectedness.\<close>
-lemma openin_open: "openin (subtopology euclidean U) S \<longleftrightarrow> (\<exists>T. open T \<and> (S = U \<inter> T))"
+lemma openin_open: "openin (top_of_set U) S \<longleftrightarrow> (\<exists>T. open T \<and> (S = U \<inter> T))"
by (auto simp: openin_subtopology)
lemma openin_Int_open:
- "\<lbrakk>openin (subtopology euclidean U) S; open T\<rbrakk>
- \<Longrightarrow> openin (subtopology euclidean U) (S \<inter> T)"
+ "\<lbrakk>openin (top_of_set U) S; open T\<rbrakk>
+ \<Longrightarrow> openin (top_of_set U) (S \<inter> T)"
by (metis open_Int Int_assoc openin_open)
-lemma openin_open_Int[intro]: "open S \<Longrightarrow> openin (subtopology euclidean U) (U \<inter> S)"
+lemma openin_open_Int[intro]: "open S \<Longrightarrow> openin (top_of_set U) (U \<inter> S)"
by (auto simp: openin_open)
lemma open_openin_trans[trans]:
- "open S \<Longrightarrow> open T \<Longrightarrow> T \<subseteq> S \<Longrightarrow> openin (subtopology euclidean S) T"
+ "open S \<Longrightarrow> open T \<Longrightarrow> T \<subseteq> S \<Longrightarrow> openin (top_of_set S) T"
by (metis Int_absorb1 openin_open_Int)
-lemma open_subset: "S \<subseteq> T \<Longrightarrow> open S \<Longrightarrow> openin (subtopology euclidean T) S"
+lemma open_subset: "S \<subseteq> T \<Longrightarrow> open S \<Longrightarrow> openin (top_of_set T) S"
by (auto simp: openin_open)
-lemma closedin_closed: "closedin (subtopology euclidean U) S \<longleftrightarrow> (\<exists>T. closed T \<and> S = U \<inter> T)"
+lemma closedin_closed: "closedin (top_of_set U) S \<longleftrightarrow> (\<exists>T. closed T \<and> S = U \<inter> T)"
by (simp add: closedin_subtopology Int_ac)
-lemma closedin_closed_Int: "closed S \<Longrightarrow> closedin (subtopology euclidean U) (U \<inter> S)"
+lemma closedin_closed_Int: "closed S \<Longrightarrow> closedin (top_of_set U) (U \<inter> S)"
by (metis closedin_closed)
-lemma closed_subset: "S \<subseteq> T \<Longrightarrow> closed S \<Longrightarrow> closedin (subtopology euclidean T) S"
+lemma closed_subset: "S \<subseteq> T \<Longrightarrow> closed S \<Longrightarrow> closedin (top_of_set T) S"
by (auto simp: closedin_closed)
lemma closedin_closed_subset:
- "\<lbrakk>closedin (subtopology euclidean U) V; T \<subseteq> U; S = V \<inter> T\<rbrakk>
- \<Longrightarrow> closedin (subtopology euclidean T) S"
+ "\<lbrakk>closedin (top_of_set U) V; T \<subseteq> U; S = V \<inter> T\<rbrakk>
+ \<Longrightarrow> 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 "\<lbrakk>finite S; S \<subseteq> T\<rbrakk> \<Longrightarrow> closedin (subtopology euclidean T) S"
+ shows "\<lbrakk>finite S; S \<subseteq> T\<rbrakk> \<Longrightarrow> 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} \<longleftrightarrow> a \<in> U"
+ shows "closedin (top_of_set U) {a} \<longleftrightarrow> a \<in> 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 \<longleftrightarrow>
+ shows "openin (top_of_set U) S \<longleftrightarrow>
S \<subseteq> U \<and> (\<forall>x\<in>S. \<exists>e>0. \<forall>x'\<in>U. dist x' x < e \<longrightarrow> x'\<in> S)"
(is "?lhs \<longleftrightarrow> ?rhs")
proof
@@ -513,16 +512,16 @@
lemma connected_openin:
"connected S \<longleftrightarrow>
- \<not>(\<exists>E1 E2. openin (subtopology euclidean S) E1 \<and>
- openin (subtopology euclidean S) E2 \<and>
+ \<not>(\<exists>E1 E2. openin (top_of_set S) E1 \<and>
+ openin (top_of_set S) E2 \<and>
S \<subseteq> E1 \<union> E2 \<and> E1 \<inter> E2 = {} \<and> E1 \<noteq> {} \<and> E2 \<noteq> {})"
apply (simp add: connected_def openin_open disjoint_iff_not_equal, safe)
by (simp_all, blast+) (* SLOW *)
lemma connected_openin_eq:
"connected S \<longleftrightarrow>
- \<not>(\<exists>E1 E2. openin (subtopology euclidean S) E1 \<and>
- openin (subtopology euclidean S) E2 \<and>
+ \<not>(\<exists>E1 E2. openin (top_of_set S) E1 \<and>
+ openin (top_of_set S) E2 \<and>
E1 \<union> E2 = S \<and> E1 \<inter> E2 = {} \<and>
E1 \<noteq> {} \<and> E2 \<noteq> {})"
apply (simp add: connected_openin, safe, blast)
@@ -531,8 +530,8 @@
lemma connected_closedin:
"connected S \<longleftrightarrow>
(\<nexists>E1 E2.
- closedin (subtopology euclidean S) E1 \<and>
- closedin (subtopology euclidean S) E2 \<and>
+ closedin (top_of_set S) E1 \<and>
+ closedin (top_of_set S) E2 \<and>
S \<subseteq> E1 \<union> E2 \<and> E1 \<inter> E2 = {} \<and> E1 \<noteq> {} \<and> E2 \<noteq> {})"
(is "?lhs = ?rhs")
proof
@@ -561,8 +560,8 @@
lemma connected_closedin_eq:
"connected S \<longleftrightarrow>
\<not>(\<exists>E1 E2.
- closedin (subtopology euclidean S) E1 \<and>
- closedin (subtopology euclidean S) E2 \<and>
+ closedin (top_of_set S) E1 \<and>
+ closedin (top_of_set S) E2 \<and>
E1 \<union> E2 = S \<and> E1 \<inter> E2 = {} \<and>
E1 \<noteq> {} \<and> E2 \<noteq> {})"
apply (simp add: connected_closedin, safe, blast)
@@ -571,26 +570,26 @@
text \<open>These "transitivity" results are handy too\<close>
lemma openin_trans[trans]:
- "openin (subtopology euclidean T) S \<Longrightarrow> openin (subtopology euclidean U) T \<Longrightarrow>
- openin (subtopology euclidean U) S"
+ "openin (top_of_set T) S \<Longrightarrow> openin (top_of_set U) T \<Longrightarrow>
+ openin (top_of_set U) S"
by (metis openin_Int_open openin_open)
-lemma openin_open_trans: "openin (subtopology euclidean T) S \<Longrightarrow> open T \<Longrightarrow> open S"
+lemma openin_open_trans: "openin (top_of_set T) S \<Longrightarrow> open T \<Longrightarrow> open S"
by (auto simp: openin_open intro: openin_trans)
lemma closedin_trans[trans]:
- "closedin (subtopology euclidean T) S \<Longrightarrow> closedin (subtopology euclidean U) T \<Longrightarrow>
- closedin (subtopology euclidean U) S"
+ "closedin (top_of_set T) S \<Longrightarrow> closedin (top_of_set U) T \<Longrightarrow>
+ closedin (top_of_set U) S"
by (auto simp: closedin_closed closed_Inter Int_assoc)
-lemma closedin_closed_trans: "closedin (subtopology euclidean T) S \<Longrightarrow> closed T \<Longrightarrow> closed S"
+lemma closedin_closed_trans: "closedin (top_of_set T) S \<Longrightarrow> closed T \<Longrightarrow> closed S"
by (auto simp: closedin_closed intro: closedin_trans)
lemma openin_subtopology_Int_subset:
- "\<lbrakk>openin (subtopology euclidean u) (u \<inter> S); v \<subseteq> u\<rbrakk> \<Longrightarrow> openin (subtopology euclidean v) (v \<inter> S)"
+ "\<lbrakk>openin (top_of_set u) (u \<inter> S); v \<subseteq> u\<rbrakk> \<Longrightarrow> openin (top_of_set v) (v \<inter> S)"
by (auto simp: openin_subtopology)
-lemma openin_open_eq: "open s \<Longrightarrow> (openin (subtopology euclidean s) t \<longleftrightarrow> open t \<and> t \<subseteq> s)"
+lemma openin_open_eq: "open s \<Longrightarrow> (openin (top_of_set s) t \<longleftrightarrow> open t \<and> t \<subseteq> s)"
using open_subset openin_open_trans openin_subset by fastforce
@@ -1684,11 +1683,11 @@
"continuous_map (discrete_topology U) X f \<longleftrightarrow> f ` U \<subseteq> 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 \<longleftrightarrow>
@@ -2844,9 +2843,8 @@
done
qed
-lemma connectedin_iff_connected_real [simp]:
- "connectedin euclideanreal S \<longleftrightarrow> connected S"
- by (simp add: connected_def connectedin)
+lemma connectedin_iff_connected [simp]: "connectedin euclidean S \<longleftrightarrow> connected S"
+ by (simp add: connected_def connectedin)
lemma connectedin_closedin:
"connectedin X S \<longleftrightarrow>
@@ -3190,7 +3188,7 @@
unfolding compact_space_alt
using openin_subset by fastforce
-lemma compactin_euclideanreal_iff [simp]: "compactin euclideanreal S \<longleftrightarrow> compact S"
+lemma compactin_euclidean_iff [simp]: "compactin euclidean S \<longleftrightarrow> 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\<open>Retraction and section maps\<close>
+
+definition retraction_maps :: "'a topology \<Rightarrow> 'b topology \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> bool"
+ where "retraction_maps X Y f g \<equiv>
+ continuous_map X Y f \<and> continuous_map Y X g \<and> (\<forall>x \<in> topspace Y. f(g x) = x)"
+
+definition section_map :: "'a topology \<Rightarrow> 'b topology \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
+ where "section_map X Y f \<equiv> \<exists>g. retraction_maps Y X g f"
+
+definition retraction_map :: "'a topology \<Rightarrow> 'b topology \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
+ where "retraction_map X Y f \<equiv> \<exists>g. retraction_maps X Y f g"
+
+lemma retraction_maps_eq:
+ "\<lbrakk>retraction_maps X Y f g; \<And>x. x \<in> topspace X \<Longrightarrow> f x = f' x; \<And>x. x \<in> topspace Y \<Longrightarrow> g x = g' x\<rbrakk>
+ \<Longrightarrow> 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:
+ "\<lbrakk>section_map X Y f; \<And>x. x \<in> topspace X \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> section_map X Y g"
+ unfolding section_map_def using retraction_maps_eq by blast
+
+lemma retraction_map_eq:
+ "\<lbrakk>retraction_map X Y f; \<And>x. x \<in> topspace X \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> 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 \<Longrightarrow> 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 \<and> retraction_map X Y f \<longleftrightarrow> 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 \<Longrightarrow> 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 \<subseteq> topspace Y"
+ have "openin Y U"
+ if "\<forall>x\<in>topspace Y. g x \<in> topspace X" "\<forall>x\<in>topspace Y. f (g x) = x"
+ "openin Y {x \<in> topspace Y. g x \<in> {x \<in> topspace X. f x \<in> U}}" for g
+ using openin_subopen U that by fastforce
+ then show "openin X {x \<in> topspace X. f x \<in> U} = openin Y U"
+ using assms by (auto simp: retraction_map_def retraction_maps_def continuous_map_def)
+qed
+
+lemma retraction_maps_compose:
+ "\<lbrakk>retraction_maps X Y f f'; retraction_maps Y Z g g'\<rbrakk> \<Longrightarrow> retraction_maps X Z (g \<circ> f) (f' \<circ> g')"
+ by (clarsimp simp: retraction_maps_def continuous_map_compose) (simp add: continuous_map_def)
+
+lemma retraction_map_compose:
+ "\<lbrakk>retraction_map X Y f; retraction_map Y Z g\<rbrakk> \<Longrightarrow> retraction_map X Z (g \<circ> f)"
+ by (meson retraction_map_def retraction_maps_compose)
+
+lemma section_map_compose:
+ "\<lbrakk>section_map X Y f; section_map Y Z g\<rbrakk> \<Longrightarrow> section_map X Z (g \<circ> f)"
+ by (meson retraction_maps_compose section_map_def)
+
+lemma surjective_section_eq_homeomorphic_map:
+ "section_map X Y f \<and> f ` (topspace X) = topspace Y \<longleftrightarrow> 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 \<Longrightarrow> retraction_map X Y f \<or> section_map X Y f \<longleftrightarrow> 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 \<Longrightarrow> f ` (topspace X) = topspace Y"
+ by (simp add: retraction_imp_quotient_map quotient_imp_surjective_map)
+
+lemma section_imp_injective_map:
+ "\<lbrakk>section_map X Y f; x \<in> topspace X; y \<in> topspace X\<rbrakk> \<Longrightarrow> f x = f y \<longleftrightarrow> 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
+ \<Longrightarrow> retraction_maps X (subtopology X (s ` (topspace Y))) (s \<circ> r) id"
+ unfolding retraction_maps_def
+ by (auto simp: continuous_map_compose continuous_map_into_subtopology continuous_map_from_subtopology)
subsection \<open>Continuity\<close>
lemma continuous_on_open:
"continuous_on S f \<longleftrightarrow>
- (\<forall>T. openin (subtopology euclidean (f ` S)) T \<longrightarrow>
- openin (subtopology euclidean S) (S \<inter> f -` T))"
+ (\<forall>T. openin (top_of_set (f ` S)) T \<longrightarrow>
+ openin (top_of_set S) (S \<inter> 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 \<longleftrightarrow>
- (\<forall>T. closedin (subtopology euclidean (f ` S)) T \<longrightarrow>
- closedin (subtopology euclidean S) (S \<inter> f -` T))"
+ (\<forall>T. closedin (top_of_set (f ` S)) T \<longrightarrow>
+ closedin (top_of_set S) (S \<inter> 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 \<inter> f -` T)"
+ assumes "continuous_on S f" "closedin (top_of_set (f ` S)) T"
+ shows "closedin (top_of_set S) (S \<inter> f -` T)"
using assms continuous_on_closed by blast
subsection%unimportant \<open>Half-global and completely global cases\<close>
lemma continuous_openin_preimage_gen:
assumes "continuous_on S f" "open T"
- shows "openin (subtopology euclidean S) (S \<inter> f -` T)"
+ shows "openin (top_of_set S) (S \<inter> f -` T)"
proof -
have *: "(S \<inter> f -` T) = (S \<inter> f -` (T \<inter> f ` S))"
by auto
- have "openin (subtopology euclidean (f ` S)) (T \<inter> f ` S)"
+ have "openin (top_of_set (f ` S)) (T \<inter> 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 \<inter> f ` S"]]
@@ -3677,11 +3766,11 @@
lemma continuous_closedin_preimage:
assumes "continuous_on S f" and "closed T"
- shows "closedin (subtopology euclidean S) (S \<inter> f -` T)"
+ shows "closedin (top_of_set S) (S \<inter> f -` T)"
proof -
have *: "(S \<inter> f -` T) = (S \<inter> f -` (T \<inter> f ` S))"
by auto
- have "closedin (subtopology euclidean (f ` S)) (T \<inter> f ` S)"
+ have "closedin (top_of_set (f ` S)) (T \<inter> 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 \<longleftrightarrow>
- (\<forall>T. open T \<longrightarrow> openin (subtopology euclidean S) (S \<inter> f -` T))"
+ (\<forall>T. open T \<longrightarrow> openin (top_of_set S) (S \<inter> 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 \<longleftrightarrow>
- (\<forall>T. closed T \<longrightarrow> closedin (subtopology euclidean S) (S \<inter> f -` T))"
+ (\<forall>T. closed T \<longrightarrow> closedin (top_of_set S) (S \<inter> 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) \<in> U" "openin (subtopology euclidean (S \<times> T)) U"
- obtains V W where "openin (subtopology euclidean S) V" "x \<in> V"
- "openin (subtopology euclidean T) W" "y \<in> W" "(V \<times> W) \<subseteq> U"
+ assumes "(x, y) \<in> U" "openin (top_of_set (S \<times> T)) U"
+ obtains V W where "openin (top_of_set S) V" "x \<in> V"
+ "openin (top_of_set T) W" "y \<in> W" "(V \<times> W) \<subseteq> U"
proof -
from assms obtain E where "open E" "U = S \<times> T \<inter> E" "(x, y) \<in> E" "x \<in> S" "y \<in> T"
by (auto simp: openin_open)
@@ -3744,8 +3833,8 @@
by blast
show ?thesis
proof
- show "openin (subtopology euclidean S) (E1 \<inter> S)"
- "openin (subtopology euclidean T) (E2 \<inter> T)"
+ show "openin (top_of_set S) (E1 \<inter> S)"
+ "openin (top_of_set T) (E2 \<inter> T)"
using \<open>open E1\<close> \<open>open E2\<close>
by (auto simp: openin_open)
show "x \<in> E1 \<inter> S" "y \<in> E2 \<inter> T"
@@ -3757,20 +3846,20 @@
qed
lemma closedin_Times:
- "closedin (subtopology euclidean S) S' \<Longrightarrow> closedin (subtopology euclidean T) T' \<Longrightarrow>
- closedin (subtopology euclidean (S \<times> T)) (S' \<times> T')"
+ "closedin (top_of_set S) S' \<Longrightarrow> closedin (top_of_set T) T' \<Longrightarrow>
+ closedin (top_of_set (S \<times> T)) (S' \<times> T')"
unfolding closedin_closed using closed_Times by blast
lemma openin_Times:
- "openin (subtopology euclidean S) S' \<Longrightarrow> openin (subtopology euclidean T) T' \<Longrightarrow>
- openin (subtopology euclidean (S \<times> T)) (S' \<times> T')"
+ "openin (top_of_set S) S' \<Longrightarrow> openin (top_of_set T) T' \<Longrightarrow>
+ openin (top_of_set (S \<times> T)) (S' \<times> 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 \<times> T)) (S' \<times> T') \<longleftrightarrow>
- S' = {} \<or> T' = {} \<or> openin (subtopology euclidean S) S' \<and> openin (subtopology euclidean T) T'"
+ "openin (top_of_set (S \<times> T)) (S' \<times> T') \<longleftrightarrow>
+ S' = {} \<or> T' = {} \<or> openin (top_of_set S) S' \<and> openin (top_of_set T) T'"
(is "?lhs = ?rhs")
proof (cases "S' = {} \<or> 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 _ \<open>?lhs\<close>])
using \<open>y \<in> T'\<close>
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 _ \<open>?lhs\<close>])
using \<open>x \<in> S'\<close>
@@ -3805,7 +3894,7 @@
lemma Lim_transform_within_openin:
assumes f: "(f \<longlongrightarrow> l) (at a within T)"
- and "openin (subtopology euclidean T) S" "a \<in> S"
+ and "openin (top_of_set T) S" "a \<in> S"
and eq: "\<And>x. \<lbrakk>x \<in> S; x \<noteq> a\<rbrakk> \<Longrightarrow> f x = g x"
shows "(g \<longlongrightarrow> l) (at a within T)"
proof -
@@ -3827,8 +3916,8 @@
lemma continuous_on_open_gen:
assumes "f ` S \<subseteq> T"
shows "continuous_on S f \<longleftrightarrow>
- (\<forall>U. openin (subtopology euclidean T) U
- \<longrightarrow> openin (subtopology euclidean S) (S \<inter> f -` U))"
+ (\<forall>U. openin (top_of_set T) U
+ \<longrightarrow> openin (top_of_set S) (S \<inter> 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 \<inter> f -` (U \<inter> T))"
+ then have "openin (top_of_set S) (S \<inter> f -` (U \<inter> T))"
by (metis R inf_commute openin_open)
- then show "openin (subtopology euclidean S) (S \<inter> f -` U)"
+ then show "openin (top_of_set S) (S \<inter> 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:
- "\<lbrakk>continuous_on S f; f ` S \<subseteq> T; openin (subtopology euclidean T) U\<rbrakk>
- \<Longrightarrow> openin (subtopology euclidean S) (S \<inter> f -` U)"
+ "\<lbrakk>continuous_on S f; f ` S \<subseteq> T; openin (top_of_set T) U\<rbrakk>
+ \<Longrightarrow> openin (top_of_set S) (S \<inter> f -` U)"
by (simp add: continuous_on_open_gen)
lemma continuous_on_closed_gen:
assumes "f ` S \<subseteq> T"
shows "continuous_on S f \<longleftrightarrow>
- (\<forall>U. closedin (subtopology euclidean T) U
- \<longrightarrow> closedin (subtopology euclidean S) (S \<inter> f -` U))"
+ (\<forall>U. closedin (top_of_set T) U
+ \<longrightarrow> closedin (top_of_set S) (S \<inter> f -` U))"
(is "?lhs = ?rhs")
proof -
have *: "U \<subseteq> T \<Longrightarrow> S \<inter> f -` (T - U) = S - (S \<inter> 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 \<inter> f -` U)"
+ assume "closedin (top_of_set T) U"
+ then show "closedin (top_of_set S) (S \<inter> 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 \<subseteq> T" "closedin (subtopology euclidean T) U"
- shows "closedin (subtopology euclidean S) (S \<inter> f -` U)"
+ assumes "continuous_on S f" "f ` S \<subseteq> T" "closedin (top_of_set T) U"
+ shows "closedin (top_of_set S) (S \<inter> 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 \<in> S"
+ and "openin (top_of_set T) S" "a \<in> S"
and eq: "\<And>x. x \<in> S \<Longrightarrow> f x = g x"
shows "continuous (at a within T) g"
using assms by (simp add: Lim_transform_within_openin continuous_within)
+subsection%important \<open>The topology generated by some (open) subsets\<close>
+
+text \<open>In the definition below of a generated topology, the \<open>Empty\<close> case is not necessary,
+as it follows from \<open>UN\<close> taking for \<open>K\<close> 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 \<open>UNIV\<close> to be an open set, as this will not be the case in applications. (We are
+thinking of a topology on a subset of \<open>UNIV\<close>, the remaining part of \<open>UNIV\<close> being irrelevant.)\<close>
+
+inductive generate_topology_on for S where
+ Empty: "generate_topology_on S {}"
+| Int: "generate_topology_on S a \<Longrightarrow> generate_topology_on S b \<Longrightarrow> generate_topology_on S (a \<inter> b)"
+| UN: "(\<And>k. k \<in> K \<Longrightarrow> generate_topology_on S k) \<Longrightarrow> generate_topology_on S (\<Union>K)"
+| Basis: "s \<in> S \<Longrightarrow> 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 \<open>The basic property of the topology generated by a set \<open>S\<close> is that it is the
+smallest topology containing all the elements of \<open>S\<close>:\<close>
+
+lemma generate_topology_on_coarsest:
+ assumes "istopology T"
+ "\<And>s. s \<in> S \<Longrightarrow> 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) \<Rightarrow> ('a topology)"
+ where "topology_generated_by S \<equiv> topology (generate_topology_on S)"
+
+lemma openin_topology_generated_by_iff:
+ "openin (topology_generated_by S) s \<longleftrightarrow> 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 \<Longrightarrow> generate_topology_on S s"
+using openin_topology_generated_by_iff by auto
+
+lemma topology_generated_by_topspace:
+ "topspace (topology_generated_by S) = (\<Union>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 \<subseteq> (\<Union>S)" by (induct, auto)
+ }
+ then show "topspace (topology_generated_by S) \<subseteq> (\<Union>S)"
+ unfolding topspace_def by auto
+next
+ have "generate_topology_on S (\<Union>S)"
+ using generate_topology_on.UN[OF generate_topology_on.Basis, of S S] by simp
+ then show "(\<Union>S) \<subseteq> topspace (topology_generated_by S)"
+ unfolding topspace_def using openin_topology_generated_by_iff by auto
+qed
+
+lemma topology_generated_by_Basis:
+ "s \<in> S \<Longrightarrow> 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:
+ "\<lbrakk>finite \<F>; \<And>K. K \<in> \<F> \<Longrightarrow> generate_topology_on \<S> K; \<F> \<noteq> {}\<rbrakk> \<Longrightarrow> generate_topology_on \<S> (\<Inter>\<F>)"
+ by (induction \<F> rule: finite_induct; force intro: generate_topology_on.intros)
+
+subsection\<open>Topology bases and sub-bases\<close>
+
+lemma istopology_base_alt:
+ "istopology (arbitrary union_of P) \<longleftrightarrow>
+ (\<forall>S T. (arbitrary union_of P) S \<and> (arbitrary union_of P) T
+ \<longrightarrow> (arbitrary union_of P) (S \<inter> T))"
+ by (simp add: istopology_def) (blast intro: arbitrary_union_of_Union)
+
+lemma istopology_base_eq:
+ "istopology (arbitrary union_of P) \<longleftrightarrow>
+ (\<forall>S T. P S \<and> P T \<longrightarrow> (arbitrary union_of P) (S \<inter> T))"
+ by (simp add: istopology_base_alt arbitrary_union_of_Int_eq)
+
+lemma istopology_base:
+ "(\<And>S T. \<lbrakk>P S; P T\<rbrakk> \<Longrightarrow> P(S \<inter> T)) \<Longrightarrow> 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 \<longleftrightarrow>
+ (\<forall>V. P V \<longrightarrow> openin X V) \<and> (\<forall>U x. openin X U \<and> x \<in> U \<longrightarrow> (\<exists>V. P V \<and> x \<in> V \<and> V \<subseteq> U))"
+ (is "?lhs = ?rhs")
+proof
+ assume ?lhs
+ then show ?rhs
+ by (auto simp: union_of_def arbitrary_def)
+next
+ assume R: ?rhs
+ then have *: "\<exists>\<U>\<subseteq>Collect P. \<Union>\<U> = S" if "openin X S" for S
+ using that by (rule_tac x="{V. P V \<and> V \<subseteq> S}" in exI) fastforce
+ from R show ?lhs
+ by (fastforce simp add: union_of_def arbitrary_def intro: *)
+qed
+
+lemma topology_base_unique:
+ "\<lbrakk>\<And>S. P S \<Longrightarrow> openin X S;
+ \<And>U x. \<lbrakk>openin X U; x \<in> U\<rbrakk> \<Longrightarrow> \<exists>B. P B \<and> x \<in> B \<and> B \<subseteq> U\<rbrakk>
+ \<Longrightarrow> topology(arbitrary union_of P) = X"
+ by (metis openin_topology_base_unique openin_inverse [of X])
+
+lemma topology_bases_eq_aux:
+ "\<lbrakk>(arbitrary union_of P) S;
+ \<And>U x. \<lbrakk>P U; x \<in> U\<rbrakk> \<Longrightarrow> \<exists>V. Q V \<and> x \<in> V \<and> V \<subseteq> U\<rbrakk>
+ \<Longrightarrow> (arbitrary union_of Q) S"
+ by (metis arbitrary_union_of_alt arbitrary_union_of_idempot)
+
+lemma topology_bases_eq:
+ "\<lbrakk>\<And>U x. \<lbrakk>P U; x \<in> U\<rbrakk> \<Longrightarrow> \<exists>V. Q V \<and> x \<in> V \<and> V \<subseteq> U;
+ \<And>V x. \<lbrakk>Q V; x \<in> V\<rbrakk> \<Longrightarrow> \<exists>U. P U \<and> x \<in> U \<and> U \<subseteq> V\<rbrakk>
+ \<Longrightarrow> 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
+ \<longleftrightarrow> (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 \<subseteq> U"
+ by (metis arbitrary_union_of_relative_to openin_subbase openin_topspace relative_to_imp_subset)
+ show "U \<subseteq> ?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:
+ "\<lbrakk>\<And>S. P S \<Longrightarrow> openin X S; openin X U;
+ openin(topology(arbitrary union_of (finite intersection_of P relative_to U))) S\<rbrakk>
+ \<Longrightarrow> 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 (\<lambda>x. x \<in> 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 \<U> where \<U>: "\<And>T. T \<in> \<U> \<Longrightarrow> \<exists>\<F>. finite' \<F> \<and> \<F> \<subseteq> S \<and> \<Inter>\<F> = T" and eq: "A = \<Union>\<U>"
+ unfolding union_of_def intersection_of_def by auto
+ show "?lhs A"
+ unfolding eq
+ proof (rule generate_topology_on.UN)
+ fix T
+ assume "T \<in> \<U>"
+ with \<U> obtain \<F> where "finite' \<F>" "\<F> \<subseteq> S" "\<Inter>\<F> = T"
+ by blast
+ have "generate_topology_on S (\<Inter>\<F>)"
+ proof (rule generate_topology_on_Inter)
+ show "finite \<F>" "\<F> \<noteq> {}"
+ by (auto simp: \<open>finite' \<F>\<close>)
+ show "\<And>K. K \<in> \<F> \<Longrightarrow> generate_topology_on S K"
+ by (metis \<open>\<F> \<subseteq> S\<close> generate_topology_on.simps subset_iff)
+ qed
+ then show "generate_topology_on S T"
+ using \<open>\<Inter>\<F> = T\<close> by blast
+ qed
+qed
+
+subsubsection%important \<open>Continuity\<close>
+
+text \<open>We will need to deal with continuous maps in terms of topologies and not in terms
+of type classes, as defined below.\<close>
+
+definition%important continuous_on_topo::"'a topology \<Rightarrow> 'b topology \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
+ where "continuous_on_topo T1 T2 f = ((\<forall> U. openin T2 U \<longrightarrow> openin T1 (f-`U \<inter> topspace(T1)))
+ \<and> (f`(topspace T1) \<subseteq> (topspace T2)))"
+
+lemma continuous_on_continuous_on_topo:
+ "continuous_on s f \<longleftrightarrow> 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 \<longleftrightarrow> 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 \<Longrightarrow> openin T2 U \<Longrightarrow> openin T1 (f-`U \<inter> topspace(T1))"
+ unfolding continuous_on_topo_def by auto
+
+lemma continuous_on_topo_topspace [intro]:
+ "continuous_on_topo T1 T2 f \<Longrightarrow> f`(topspace T1) \<subseteq> (topspace T2)"
+unfolding continuous_on_topo_def by auto
+
+lemma continuous_on_generated_topo_iff:
+ "continuous_on_topo T1 (topology_generated_by S) f \<longleftrightarrow>
+ ((\<forall>U. U \<in> S \<longrightarrow> openin T1 (f-`U \<inter> topspace(T1))) \<and> (f`(topspace T1) \<subseteq> (\<Union> S)))"
+unfolding continuous_on_topo_def topology_generated_by_topspace
+proof (auto simp add: topology_generated_by_Basis)
+ assume H: "\<forall>U. U \<in> S \<longrightarrow> openin T1 (f -` U \<inter> 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 \<inter> topspace T1)"
+ proof (induct)
+ fix a b
+ assume H: "openin T1 (f -` a \<inter> topspace T1)" "openin T1 (f -` b \<inter> topspace T1)"
+ have "f -` (a \<inter> b) \<inter> topspace T1 = (f-`a \<inter> topspace T1) \<inter> (f-`b \<inter> topspace T1)"
+ by auto
+ then show "openin T1 (f -` (a \<inter> b) \<inter> topspace T1)" using H by auto
+ next
+ fix K
+ assume H: "openin T1 (f -` k \<inter> topspace T1)" if "k\<in> K" for k
+ define L where "L = {f -` k \<inter> topspace T1|k. k \<in> K}"
+ have *: "openin T1 l" if "l \<in>L" for l using that H unfolding L_def by auto
+ have "openin T1 (\<Union>L)" using openin_Union[OF *] by simp
+ moreover have "(\<Union>L) = (f -` \<Union>K \<inter> topspace T1)" unfolding L_def by auto
+ ultimately show "openin T1 (f -` \<Union>K \<inter> topspace T1)" by simp
+ qed (auto simp add: H)
+qed
+
+lemma continuous_on_generated_topo:
+ assumes "\<And>U. U \<in>S \<Longrightarrow> openin T1 (f-`U \<inter> topspace(T1))"
+ "f`(topspace T1) \<subseteq> (\<Union> 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 \<inter> topspace T2) \<inter> topspace T1)"
+ using H assms by blast
+ moreover have "f -` (g -` U \<inter> topspace T2) \<inter> topspace T1 = (g \<circ> f) -` U \<inter> topspace T1"
+ using H assms continuous_on_topo_topspace by fastforce
+ ultimately show "openin T1 ((g \<circ> f) -` U \<inter> topspace T1)"
+ by simp
+qed (blast)
+
+lemma continuous_on_topo_preimage_topspace [intro]:
+ assumes "continuous_on_topo T1 T2 f"
+ shows "f-`(topspace T2) \<inter> topspace T1 = topspace T1"
+using assms unfolding continuous_on_topo_def by auto
+
+
+subsubsection%important \<open>Pullback topology\<close>
+
+text \<open>Pulling back a topology by map gives again a topology. \<open>subtopology\<close> 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.\<close>
+
+text \<open>\<open>pullback_topology A f T\<close> is the pullback of the topology \<open>T\<close> by the map \<open>f\<close> on
+the set \<open>A\<close>.\<close>
+
+definition%important pullback_topology::"('a set) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('b topology) \<Rightarrow> ('a topology)"
+ where "pullback_topology A f T = topology (\<lambda>S. \<exists>U. openin T U \<and> S = f-`U \<inter> A)"
+
+lemma istopology_pullback_topology:
+ "istopology (\<lambda>S. \<exists>U. openin T U \<and> S = f-`U \<inter> A)"
+ unfolding istopology_def proof (auto)
+ fix K assume "\<forall>S\<in>K. \<exists>U. openin T U \<and> S = f -` U \<inter> A"
+ then have "\<exists>U. \<forall>S\<in>K. openin T (U S) \<and> S = f-`(U S) \<inter> A"
+ by (rule bchoice)
+ then obtain U where U: "\<forall>S\<in>K. openin T (U S) \<and> S = f-`(U S) \<inter> A"
+ by blast
+ define V where "V = (\<Union>S\<in>K. U S)"
+ have "openin T V" "\<Union>K = f -` V \<inter> A" unfolding V_def using U by auto
+ then show "\<exists>V. openin T V \<and> \<Union>K = f -` V \<inter> A" by auto
+qed
+
+lemma openin_pullback_topology:
+ "openin (pullback_topology A f T) S \<longleftrightarrow> (\<exists>U. openin T U \<and> S = f-`U \<inter> 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) \<inter> 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 \<inter> topspace T1)"
+ using assms unfolding continuous_on_topo_def by auto
+ have "(g o f)-`U \<inter> topspace (pullback_topology A f T1) = (g o f)-`U \<inter> A \<inter> f-`(topspace T1)"
+ unfolding topspace_pullback_topology by auto
+ also have "... = f-`(g-`U \<inter> topspace T1) \<inter> A "
+ by auto
+ also have "openin (pullback_topology A f T1) (...)"
+ unfolding openin_pullback_topology using \<open>openin T1 (g-`U \<inter> topspace T1)\<close> by auto
+ finally show "openin (pullback_topology A f T1) ((g \<circ> f) -` U \<inter> topspace (pullback_topology A f T1))"
+ by auto
+next
+ fix x assume "x \<in> topspace (pullback_topology A f T1)"
+ then have "f x \<in> topspace T1"
+ unfolding topspace_pullback_topology by auto
+ then show "g (f x) \<in> 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 \<subseteq> 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 "\<exists>V. openin T2 V \<and> U = f-`V \<inter> A"
+ unfolding openin_pullback_topology by auto
+ then obtain V where "openin T2 V" "U = f-`V \<inter> A"
+ by blast
+ then have "g -` U \<inter> topspace T1 = g-`(f-`V \<inter> A) \<inter> topspace T1"
+ by blast
+ also have "... = (f o g)-`V \<inter> (g-`A \<inter> topspace T1)"
+ by auto
+ also have "... = (f o g)-`V \<inter> topspace T1"
+ using assms(2) by auto
+ also have "openin T1 (...)"
+ using assms(1) \<open>openin T2 V\<close> by auto
+ finally show "openin T1 (g -` U \<inter> topspace T1)" by simp
+next
+ fix x assume "x \<in> topspace T1"
+ have "(f o g) x \<in> topspace T2"
+ using assms(1) \<open>x \<in> topspace T1\<close> unfolding continuous_on_topo_def by auto
+ then have "g x \<in> f-`(topspace T2)"
+ unfolding comp_def by blast
+ moreover have "g x \<in> A" using assms(2) \<open>x \<in> topspace T1\<close> by blast
+ ultimately show "g x \<in> topspace (pullback_topology A f T2)"
+ unfolding topspace_pullback_topology by blast
+qed
+
end
--- a/src/HOL/Analysis/Abstract_Topology_2.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Abstract_Topology_2.thy Tue Mar 19 16:14:59 2019 +0000
@@ -99,14 +99,14 @@
qed
lemma closedin_limpt:
- "closedin (subtopology euclidean T) S \<longleftrightarrow> S \<subseteq> T \<and> (\<forall>x. x islimpt S \<and> x \<in> T \<longrightarrow> x \<in> S)"
+ "closedin (top_of_set T) S \<longleftrightarrow> S \<subseteq> T \<and> (\<forall>x. x islimpt S \<and> x \<in> T \<longrightarrow> x \<in> 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 \<Longrightarrow> closedin (subtopology euclidean S) T \<longleftrightarrow> closed T \<and> T \<subseteq> S"
+lemma closedin_closed_eq: "closed S \<Longrightarrow> closedin (top_of_set S) T \<longleftrightarrow> closed T \<and> T \<subseteq> 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 \<Longrightarrow> S \<subseteq> T \<Longrightarrow> T \<subseteq> U \<Longrightarrow>
- closedin (subtopology euclidean T) S"
+ "closedin (top_of_set U) S \<Longrightarrow> S \<subseteq> T \<Longrightarrow> T \<subseteq> U \<Longrightarrow>
+ closedin (top_of_set T) S"
by (meson closedin_limpt subset_iff)
lemma openin_subset_trans:
- "openin (subtopology euclidean U) S \<Longrightarrow> S \<subseteq> T \<Longrightarrow> T \<subseteq> U \<Longrightarrow>
- openin (subtopology euclidean T) S"
+ "openin (top_of_set U) S \<Longrightarrow> S \<subseteq> T \<Longrightarrow> T \<subseteq> U \<Longrightarrow>
+ openin (top_of_set T) S"
by (auto simp: openin_open)
lemma closedin_compact:
- "\<lbrakk>compact S; closedin (subtopology euclidean S) T\<rbrakk> \<Longrightarrow> compact T"
+ "\<lbrakk>compact S; closedin (top_of_set S) T\<rbrakk> \<Longrightarrow> compact T"
by (metis closedin_closed compact_Int_closed)
lemma closedin_compact_eq:
fixes S :: "'a::t2_space set"
shows
"compact S
- \<Longrightarrow> (closedin (subtopology euclidean S) T \<longleftrightarrow>
+ \<Longrightarrow> (closedin (top_of_set S) T \<longleftrightarrow>
compact T \<and> T \<subseteq> S)"
by (metis closedin_imp_subset closedin_compact closed_subset compact_imp_closed)
subsection \<open>Closure\<close>
+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 \<subseteq> U"
+ assumes ope: "openin (top_of_set U) S" and "T \<subseteq> U"
shows "closure(S \<inter> closure T) = closure(S \<inter> T)"
proof
obtain V where "open V" and S: "S = U \<inter> V"
@@ -171,16 +174,16 @@
corollary infinite_openin:
fixes S :: "'a :: t1_space set"
- shows "\<lbrakk>openin (subtopology euclidean U) S; x \<in> S; x islimpt U\<rbrakk> \<Longrightarrow> infinite S"
+ shows "\<lbrakk>openin (top_of_set U) S; x \<in> S; x islimpt U\<rbrakk> \<Longrightarrow> infinite S"
by (clarsimp simp add: openin_open islimpt_eq_acc_point inf_commute)
lemma closure_Int_ballI:
- assumes "\<And>U. \<lbrakk>openin (subtopology euclidean S) U; U \<noteq> {}\<rbrakk> \<Longrightarrow> T \<inter> U \<noteq> {}"
+ assumes "\<And>U. \<lbrakk>openin (top_of_set S) U; U \<noteq> {}\<rbrakk> \<Longrightarrow> T \<inter> U \<noteq> {}"
shows "S \<subseteq> closure T"
proof (clarsimp simp: closure_iff_nhds_not_empty)
fix x and A and V
assume "x \<in> S" "V \<subseteq> A" "open V" "x \<in> V" "T \<inter> A = {}"
- then have "openin (subtopology euclidean S) (A \<inter> V \<inter> S)"
+ then have "openin (top_of_set S) (A \<inter> V \<inter> S)"
by (auto simp: openin_open intro!: exI[where x="V"])
moreover have "A \<inter> V \<inter> S \<noteq> {}" using \<open>x \<in> V\<close> \<open>V \<subseteq> A\<close> \<open>x \<in> S\<close>
by auto
@@ -192,6 +195,12 @@
subsection \<open>Frontier\<close>
+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:
"\<lbrakk>connected s; s \<inter> t \<noteq> {}; s - t \<noteq> {}\<rbrakk> \<Longrightarrow> (s \<inter> frontier t \<noteq> {})"
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
- \<Longrightarrow> openin (subtopology euclidean u) (s - {a})"
+ shows "openin (top_of_set u) s
+ \<Longrightarrow> openin (top_of_set u) (s - {a})"
by (metis Int_Diff open_delete openin_open)
lemma compact_eq_openin_cover:
"compact S \<longleftrightarrow>
- (\<forall>C. (\<forall>c\<in>C. openin (subtopology euclidean S) c) \<and> S \<subseteq> \<Union>C \<longrightarrow>
+ (\<forall>C. (\<forall>c\<in>C. openin (top_of_set S) c) \<and> S \<subseteq> \<Union>C \<longrightarrow>
(\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D))"
proof safe
fix C
- assume "compact S" and "\<forall>c\<in>C. openin (subtopology euclidean S) c" and "S \<subseteq> \<Union>C"
+ assume "compact S" and "\<forall>c\<in>C. openin (top_of_set S) c" and "S \<subseteq> \<Union>C"
then have "\<forall>c\<in>{T. open T \<and> S \<inter> T \<in> C}. open c" and "S \<subseteq> \<Union>{T. open T \<and> S \<inter> T \<in> C}"
unfolding openin_open by force+
with \<open>compact S\<close> obtain D where "D \<subseteq> {T. open T \<and> S \<inter> T \<in> C}" and "finite D" and "S \<subseteq> \<Union>D"
@@ -223,14 +232,14 @@
by auto
then show "\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D" ..
next
- assume 1: "\<forall>C. (\<forall>c\<in>C. openin (subtopology euclidean S) c) \<and> S \<subseteq> \<Union>C \<longrightarrow>
+ assume 1: "\<forall>C. (\<forall>c\<in>C. openin (top_of_set S) c) \<and> S \<subseteq> \<Union>C \<longrightarrow>
(\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D)"
show "compact S"
proof (rule compactI)
fix C
let ?C = "image (\<lambda>T. S \<inter> T) C"
assume "\<forall>t\<in>C. open t" and "S \<subseteq> \<Union>C"
- then have "(\<forall>c\<in>?C. openin (subtopology euclidean S) c) \<and> S \<subseteq> \<Union>?C"
+ then have "(\<forall>c\<in>?C. openin (top_of_set S) c) \<and> S \<subseteq> \<Union>?C"
unfolding openin_open by auto
with 1 obtain D where "D \<subseteq> ?C" and "finite D" and "S \<subseteq> \<Union>D"
by metis
@@ -276,7 +285,7 @@
lemma continuous_closedin_preimage_constant:
fixes f :: "_ \<Rightarrow> 'b::t1_space"
- shows "continuous_on S f \<Longrightarrow> closedin (subtopology euclidean S) {x \<in> S. f x = a}"
+ shows "continuous_on S f \<Longrightarrow> closedin (top_of_set S) {x \<in> 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 \<open>Continuity relative to a union.\<close>
lemma continuous_on_Un_local:
- "\<lbrakk>closedin (subtopology euclidean (s \<union> t)) s; closedin (subtopology euclidean (s \<union> t)) t;
+ "\<lbrakk>closedin (top_of_set (s \<union> t)) s; closedin (top_of_set (s \<union> t)) t;
continuous_on s f; continuous_on t f\<rbrakk>
\<Longrightarrow> continuous_on (s \<union> 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:
- "\<lbrakk>closedin (subtopology euclidean (s \<union> t)) s; closedin (subtopology euclidean (s \<union> t)) t;
+ "\<lbrakk>closedin (top_of_set (s \<union> t)) s; closedin (top_of_set (s \<union> t)) t;
continuous_on s f; continuous_on t g;
\<And>x. \<lbrakk>x \<in> s \<and> \<not>P x \<or> x \<in> t \<and> P x\<rbrakk> \<Longrightarrow> f x = g x\<rbrakk>
\<Longrightarrow> continuous_on (s \<union> t) (\<lambda>x. if P x then f x else g x)"
@@ -358,9 +367,9 @@
proof -
have s: "s = (s \<inter> h -` atMost a) \<union> (s \<inter> h -` atLeast a)"
by force
- have 1: "closedin (subtopology euclidean s) (s \<inter> h -` atMost a)"
+ have 1: "closedin (top_of_set s) (s \<inter> h -` atMost a)"
by (rule continuous_closedin_preimage [OF h closed_atMost])
- have 2: "closedin (subtopology euclidean s) (s \<inter> h -` atLeast a)"
+ have 2: "closedin (top_of_set s) (s \<inter> h -` atLeast a)"
by (rule continuous_closedin_preimage [OF h closed_atLeast])
have eq: "s \<inter> h -` {..a} = {t \<in> s. h t \<le> a}" "s \<inter> h -` {a..} = {t \<in> s. a \<le> h t}"
by auto
@@ -388,7 +397,7 @@
assumes contf: "continuous_on S f"
and imf: "f ` S = T"
and injf: "\<And>x. x \<in> S \<Longrightarrow> g (f x) = x"
- and oo: "\<And>U. openin (subtopology euclidean S) U \<Longrightarrow> openin (subtopology euclidean T) (f ` U)"
+ and oo: "\<And>U. openin (top_of_set S) U \<Longrightarrow> 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: "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x"
- and oo: "\<And>U. closedin (subtopology euclidean S) U \<Longrightarrow> closedin (subtopology euclidean T) (f ` U)"
+ and oo: "\<And>U. closedin (top_of_set S) U \<Longrightarrow> 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: "\<And>U. openin (subtopology euclidean S) U \<Longrightarrow> openin (subtopology euclidean T) (f ` U)"
+ and oo: "\<And>U. openin (top_of_set S) U \<Longrightarrow> 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: "\<And>U. closedin (subtopology euclidean S) U \<Longrightarrow> closedin (subtopology euclidean T) (f ` U)"
+ and oo: "\<And>U. closedin (top_of_set S) U \<Longrightarrow> 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 \<inter> 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 \<inter> g -` U"
using closedin_subset by (fastforce simp: homeomorphism_def rev_image_eqI)
@@ -476,13 +485,13 @@
obtains \<B> :: "'a:: second_countable_topology set set"
where "countable \<B>"
"{} \<notin> \<B>"
- "\<And>C. C \<in> \<B> \<Longrightarrow> openin(subtopology euclidean S) C"
- "\<And>T. openin(subtopology euclidean S) T \<Longrightarrow> \<exists>\<U>. \<U> \<subseteq> \<B> \<and> T = \<Union>\<U>"
+ "\<And>C. C \<in> \<B> \<Longrightarrow> openin(top_of_set S) C"
+ "\<And>T. openin(top_of_set S) T \<Longrightarrow> \<exists>\<U>. \<U> \<subseteq> \<B> \<and> T = \<Union>\<U>"
proof -
obtain \<B> :: "'a set set"
where "countable \<B>"
- and opeB: "\<And>C. C \<in> \<B> \<Longrightarrow> openin(subtopology euclidean S) C"
- and \<B>: "\<And>T. openin(subtopology euclidean S) T \<Longrightarrow> \<exists>\<U>. \<U> \<subseteq> \<B> \<and> T = \<Union>\<U>"
+ and opeB: "\<And>C. C \<in> \<B> \<Longrightarrow> openin(top_of_set S) C"
+ and \<B>: "\<And>T. openin(top_of_set S) T \<Longrightarrow> \<exists>\<U>. \<U> \<subseteq> \<B> \<and> T = \<Union>\<U>"
proof -
obtain \<C> :: "'a set set"
where "countable \<C>" and ope: "\<And>C. C \<in> \<C> \<Longrightarrow> open C"
@@ -492,9 +501,9 @@
proof
show "countable ((\<lambda>C. S \<inter> C) ` \<C>)"
by (simp add: \<open>countable \<C>\<close>)
- show "\<And>C. C \<in> (\<inter>) S ` \<C> \<Longrightarrow> openin (subtopology euclidean S) C"
+ show "\<And>C. C \<in> (\<inter>) S ` \<C> \<Longrightarrow> openin (top_of_set S) C"
using ope by auto
- show "\<And>T. openin (subtopology euclidean S) T \<Longrightarrow> \<exists>\<U>\<subseteq>(\<inter>) S ` \<C>. T = \<Union>\<U>"
+ show "\<And>T. openin (top_of_set S) T \<Longrightarrow> \<exists>\<U>\<subseteq>(\<inter>) S ` \<C>. T = \<Union>\<U>"
by (metis \<C> image_mono inf_Sup openin_open)
qed
qed
@@ -502,9 +511,9 @@
proof
show "countable (\<B> - {{}})"
using \<open>countable \<B>\<close> by blast
- show "\<And>C. \<lbrakk>C \<in> \<B> - {{}}\<rbrakk> \<Longrightarrow> openin (subtopology euclidean S) C"
- by (simp add: \<open>\<And>C. C \<in> \<B> \<Longrightarrow> openin (subtopology euclidean S) C\<close>)
- show "\<exists>\<U>\<subseteq>\<B> - {{}}. T = \<Union>\<U>" if "openin (subtopology euclidean S) T" for T
+ show "\<And>C. \<lbrakk>C \<in> \<B> - {{}}\<rbrakk> \<Longrightarrow> openin (top_of_set S) C"
+ by (simp add: \<open>\<And>C. C \<in> \<B> \<Longrightarrow> openin (top_of_set S) C\<close>)
+ show "\<exists>\<U>\<subseteq>\<B> - {{}}. T = \<Union>\<U>" if "openin (top_of_set S) T" for T
using \<B> [OF that]
apply clarify
apply (rule_tac x="\<U> - {{}}" in exI, auto)
@@ -514,7 +523,7 @@
lemma Lindelof_openin:
fixes \<F> :: "'a::second_countable_topology set set"
- assumes "\<And>S. S \<in> \<F> \<Longrightarrow> openin (subtopology euclidean U) S"
+ assumes "\<And>S. S \<in> \<F> \<Longrightarrow> openin (top_of_set U) S"
obtains \<F>' where "\<F>' \<subseteq> \<F>" "countable \<F>'" "\<Union>\<F>' = \<Union>\<F>"
proof -
have "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>T. open T \<and> S = U \<inter> T"
@@ -535,16 +544,16 @@
lemma continuous_imp_closed_map:
fixes f :: "'a::t2_space \<Rightarrow> '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 \<inter> f -` T')) U"
- and cc: "\<And>U. closedin (subtopology euclidean S) U \<Longrightarrow> closedin (subtopology euclidean T) (f ` U)"
+ assumes cloU: "closedin (top_of_set (S \<inter> f -` T')) U"
+ and cc: "\<And>U. closedin (top_of_set S) U \<Longrightarrow> closedin (top_of_set T) (f ` U)"
and "T' \<subseteq> T"
- shows "closedin (subtopology euclidean T') (f ` U)"
+ shows "closedin (top_of_set T') (f ` U)"
proof -
obtain V where "closed V" "U = S \<inter> f -` T' \<inter> V"
using cloU by (auto simp: closedin_closed)
@@ -555,10 +564,10 @@
subsection%unimportant\<open>Open Maps\<close>
lemma open_map_restrict:
- assumes opeU: "openin (subtopology euclidean (S \<inter> f -` T')) U"
- and oo: "\<And>U. openin (subtopology euclidean S) U \<Longrightarrow> openin (subtopology euclidean T) (f ` U)"
+ assumes opeU: "openin (top_of_set (S \<inter> f -` T')) U"
+ and oo: "\<And>U. openin (top_of_set S) U \<Longrightarrow> openin (top_of_set T) (f ` U)"
and "T' \<subseteq> T"
- shows "openin (subtopology euclidean T') (f ` U)"
+ shows "openin (top_of_set T') (f ` U)"
proof -
obtain V where "open V" "U = S \<inter> f -` T' \<inter> V"
using opeU by (auto simp: openin_open)
@@ -572,8 +581,8 @@
lemma quotient_map_imp_continuous_open:
assumes T: "f ` S \<subseteq> T"
and ope: "\<And>U. U \<subseteq> T
- \<Longrightarrow> (openin (subtopology euclidean S) (S \<inter> f -` U) \<longleftrightarrow>
- openin (subtopology euclidean T) U)"
+ \<Longrightarrow> (openin (top_of_set S) (S \<inter> f -` U) \<longleftrightarrow>
+ openin (top_of_set T) U)"
shows "continuous_on S f"
proof -
have [simp]: "S \<inter> f -` f ` S = S" by auto
@@ -586,8 +595,8 @@
lemma quotient_map_imp_continuous_closed:
assumes T: "f ` S \<subseteq> T"
and ope: "\<And>U. U \<subseteq> T
- \<Longrightarrow> (closedin (subtopology euclidean S) (S \<inter> f -` U) \<longleftrightarrow>
- closedin (subtopology euclidean T) U)"
+ \<Longrightarrow> (closedin (top_of_set S) (S \<inter> f -` U) \<longleftrightarrow>
+ closedin (top_of_set T) U)"
shows "continuous_on S f"
proof -
have [simp]: "S \<inter> 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 \<subseteq> f ` S"
- and ope: "\<And>T. openin (subtopology euclidean S) T
- \<Longrightarrow> openin (subtopology euclidean (f ` S)) (f ` T)"
- shows "openin (subtopology euclidean S) (S \<inter> f -` T) =
- openin (subtopology euclidean (f ` S)) T"
+ and ope: "\<And>T. openin (top_of_set S) T
+ \<Longrightarrow> openin (top_of_set (f ` S)) (f ` T)"
+ shows "openin (top_of_set S) (S \<inter> f -` T) =
+ openin (top_of_set (f ` S)) T"
proof -
have "T = f ` (S \<inter> 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 \<subseteq> f ` S"
- and ope: "\<And>T. closedin (subtopology euclidean S) T
- \<Longrightarrow> closedin (subtopology euclidean (f ` S)) (f ` T)"
- shows "openin (subtopology euclidean S) (S \<inter> f -` T) \<longleftrightarrow>
- openin (subtopology euclidean (f ` S)) T"
+ and ope: "\<And>T. closedin (top_of_set S) T
+ \<Longrightarrow> closedin (top_of_set (f ` S)) (f ` T)"
+ shows "openin (top_of_set S) (S \<inter> f -` T) \<longleftrightarrow>
+ openin (top_of_set (f ` S)) T"
(is "?lhs = ?rhs")
proof
assume ?lhs
- then have *: "closedin (subtopology euclidean S) (S - (S \<inter> f -` T))"
+ then have *: "closedin (top_of_set S) (S - (S \<inter> f -` T))"
using closedin_diff by fastforce
have [simp]: "(f ` S - f ` (S - (S \<inter> f -` T))) = T"
using T by blast
@@ -638,14 +647,14 @@
and contg: "continuous_on T g" and img: "g ` T \<subseteq> S"
and fg [simp]: "\<And>y. y \<in> T \<Longrightarrow> f(g y) = y"
and U: "U \<subseteq> T"
- shows "openin (subtopology euclidean S) (S \<inter> f -` U) \<longleftrightarrow>
- openin (subtopology euclidean T) U"
+ shows "openin (top_of_set S) (S \<inter> f -` U) \<longleftrightarrow>
+ openin (top_of_set T) U"
(is "?lhs = ?rhs")
proof -
- have f: "\<And>Z. openin (subtopology euclidean (f ` S)) Z \<Longrightarrow>
- openin (subtopology euclidean S) (S \<inter> f -` Z)"
- and g: "\<And>Z. openin (subtopology euclidean (g ` T)) Z \<Longrightarrow>
- openin (subtopology euclidean T) (T \<inter> g -` Z)"
+ have f: "\<And>Z. openin (top_of_set (f ` S)) Z \<Longrightarrow>
+ openin (top_of_set S) (S \<inter> f -` Z)"
+ and g: "\<And>Z. openin (top_of_set (g ` T)) Z \<Longrightarrow>
+ openin (top_of_set T) (T \<inter> 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 \<inter> g -` (g ` T \<inter> (S \<inter> f -` U)) = U" .
assume ?lhs
- then have *: "openin (subtopology euclidean (g ` T)) (g ` T \<inter> (S \<inter> f -` U))"
+ then have *: "openin (top_of_set (g ` T)) (g ` T \<inter> (S \<inter> 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 "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x"
and "U \<subseteq> f ` S"
- shows "openin (subtopology euclidean S) (S \<inter> f -` U) \<longleftrightarrow>
- openin (subtopology euclidean (f ` S)) U"
+ shows "openin (top_of_set S) (S \<inter> f -` U) \<longleftrightarrow>
+ 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 \<Rightarrow> 'b::t2_space"
assumes "continuous_on S f" "f ` S = T" "compact S" "U \<subseteq> T"
- shows "openin (subtopology euclidean S) (S \<inter> f -` U) \<longleftrightarrow>
- openin (subtopology euclidean T) U"
+ shows "openin (top_of_set S) (S \<inter> f -` U) \<longleftrightarrow>
+ openin (top_of_set T) U"
by (metis (no_types, lifting) assms closed_map_imp_quotient_map continuous_imp_closed_map)
-subsection%unimportant\<open>Pasting functions together\<close>
+subsection%unimportant\<open>Pasting lemmas for functions, for of casewise definitions\<close>
-text\<open>on open sets\<close>
+subsubsection\<open>on open sets\<close>
lemma pasting_lemma:
- fixes f :: "'i \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::topological_space"
- assumes clo: "\<And>i. i \<in> I \<Longrightarrow> openin (subtopology euclidean S) (T i)"
- and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_on (T i) (f i)"
- and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> S \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
- and g: "\<And>x. x \<in> S \<Longrightarrow> \<exists>j. j \<in> I \<and> x \<in> T j \<and> 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: "\<And>i. i \<in> I \<Longrightarrow> (T i) \<subseteq> S"
- using clo openin_imp_subset by blast
- have *: "(S \<inter> g -` U) = (\<Union>i \<in> I. T i \<inter> f i -` U)"
- using S f g by fastforce
- show "openin (subtopology euclidean S) (S \<inter> g -` U)"
- apply (subst *)
- apply (rule openin_Union, clarify)
- using \<open>open U\<close> clo cont continuous_openin_preimage_gen openin_trans by blast
+ assumes ope: "\<And>i. i \<in> I \<Longrightarrow> openin X (T i)"
+ and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_map(subtopology X (T i)) Y (f i)"
+ and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> topspace X \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
+ and g: "\<And>x. x \<in> topspace X \<Longrightarrow> \<exists>j. j \<in> I \<and> x \<in> T j \<and> 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 \<subseteq> 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 \<subseteq> topspace X" if "i \<in> I" for i
+ using ope by (simp add: openin_subset that)
+ have *: "topspace X \<inter> g -` U = (\<Union>i \<in> I. T i \<inter> f i -` U)"
+ using f g T by fastforce
+ have "\<And>i. i \<in> I \<Longrightarrow> openin X (T i \<inter> 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 \<inter> g -` U)"
+ by (auto simp: *)
qed
lemma pasting_lemma_exists:
- fixes f :: "'i \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::topological_space"
- assumes S: "S \<subseteq> (\<Union>i \<in> I. T i)"
- and clo: "\<And>i. i \<in> I \<Longrightarrow> openin (subtopology euclidean S) (T i)"
- and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_on (T i) (f i)"
- and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> S \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
- obtains g where "continuous_on S g" "\<And>x i. \<lbrakk>i \<in> I; x \<in> S \<inter> T i\<rbrakk> \<Longrightarrow> g x = f i x"
+ assumes X: "topspace X \<subseteq> (\<Union>i \<in> I. T i)"
+ and ope: "\<And>i. i \<in> I \<Longrightarrow> openin X (T i)"
+ and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_map (subtopology X (T i)) Y (f i)"
+ and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> topspace X \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
+ obtains g where "continuous_map X Y g" "\<And>x i. \<lbrakk>i \<in> I; x \<in> topspace X \<inter> T i\<rbrakk> \<Longrightarrow> g x = f i x"
proof
- show "continuous_on S (\<lambda>x. f (SOME i. i \<in> I \<and> x \<in> T i) x)"
- apply (rule pasting_lemma [OF clo cont])
+ let ?h = "\<lambda>x. f (SOME i. i \<in> I \<and> x \<in> 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 \<in> I \<and> x \<in> T i) x = f i x" if "i \<in> I" "x \<in> topspace X \<inter> T i" for i x
+ by (metis (no_types, lifting) IntD2 IntI f someI_ex that)
+qed
+
+lemma pasting_lemma_locally_finite:
+ assumes fin: "\<And>x. x \<in> topspace X \<Longrightarrow> \<exists>V. openin X V \<and> x \<in> V \<and> finite {i \<in> I. T i \<inter> V \<noteq> {}}"
+ and clo: "\<And>i. i \<in> I \<Longrightarrow> closedin X (T i)"
+ and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_map(subtopology X (T i)) Y (f i)"
+ and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> topspace X \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
+ and g: "\<And>x. x \<in> topspace X \<Longrightarrow> \<exists>j. j \<in> I \<and> x \<in> T j \<and> 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 \<subseteq> 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 \<subseteq> topspace X" if "i \<in> I" for i
+ using clo by (simp add: closedin_subset that)
+ have *: "topspace X \<inter> g -` U = (\<Union>i \<in> I. T i \<inter> f i -` U)"
+ using f g T by fastforce
+ have cTf: "\<And>i. i \<in> I \<Longrightarrow> closedin X (T i \<inter> 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 \<in> (\<lambda>i. T i \<inter> f i -` U) ` I. Z \<inter> V \<noteq> {}}
+ \<subseteq> (\<lambda>i. T i \<inter> f i -` U) ` {i \<in> I. T i \<inter> V \<noteq> {}}" for V
+ by auto
+ have 1: "(\<Union>i\<in>I. T i \<inter> f i -` U) \<subseteq> topspace X"
+ using T by blast
+ then have lf: "locally_finite_in X ((\<lambda>i. T i \<inter> f i -` U) ` I)"
+ unfolding locally_finite_in_def
+ using finite_subset [OF sub] fin by force
+ show "closedin X (topspace X \<inter> g -` U)"
+ apply (subst *)
+ apply (rule closedin_locally_finite_Union)
+ apply (auto intro: cTf lf)
done
+qed
+
+subsubsection\<open>Likewise on closed sets, with a finiteness assumption\<close>
+
+lemma pasting_lemma_closed:
+ assumes fin: "finite I"
+ and clo: "\<And>i. i \<in> I \<Longrightarrow> closedin X (T i)"
+ and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_map(subtopology X (T i)) Y (f i)"
+ and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> topspace X \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
+ and g: "\<And>x. x \<in> topspace X \<Longrightarrow> \<exists>j. j \<in> I \<and> x \<in> T j \<and> 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: "\<And>x. x \<in> topspace X \<Longrightarrow> \<exists>V. openin X V \<and> x \<in> V \<and> finite {i \<in> I. T i \<inter> V \<noteq> {}}"
+ and X: "topspace X \<subseteq> \<Union>(T ` I)"
+ and clo: "\<And>i. i \<in> I \<Longrightarrow> closedin X (T i)"
+ and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_map(subtopology X (T i)) Y (f i)"
+ and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> topspace X \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
+ and g: "\<And>x. x \<in> topspace X \<Longrightarrow> \<exists>j. j \<in> I \<and> x \<in> T j \<and> g x = f j x"
+ obtains g where "continuous_map X Y g" "\<And>x i. \<lbrakk>i \<in> I; x \<in> topspace X \<inter> T i\<rbrakk> \<Longrightarrow> g x = f i x"
+proof
+ show "continuous_map X Y (\<lambda>x. f(@i. i \<in> I \<and> x \<in> 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 \<in> I" "x \<in> S \<inter> T i"
+ assume "i \<in> I" and "x \<in> topspace X \<inter> T i"
+ show "f (SOME i. i \<in> I \<and> x \<in> T i) x = f i x"
+ apply (rule someI2_ex)
+ using \<open>i \<in> I\<close> \<open>x \<in> topspace X \<inter> T i\<close> apply blast
+ by (meson Int_iff \<open>i \<in> I\<close> \<open>x \<in> topspace X \<inter> T i\<close> f)
+qed
+
+lemma pasting_lemma_exists_closed:
+ assumes fin: "finite I"
+ and X: "topspace X \<subseteq> \<Union>(T ` I)"
+ and clo: "\<And>i. i \<in> I \<Longrightarrow> closedin X (T i)"
+ and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_map(subtopology X (T i)) Y (f i)"
+ and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> topspace X \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
+ obtains g where "continuous_map X Y g" "\<And>x i. \<lbrakk>i \<in> I; x \<in> topspace X \<inter> T i\<rbrakk> \<Longrightarrow> g x = f i x"
+proof
+ show "continuous_map X Y (\<lambda>x. f (SOME i. i \<in> I \<and> x \<in> T i) x)"
+ apply (rule pasting_lemma_closed [OF \<open>finite I\<close> clo cont])
+ apply (blast intro: f)+
+ by (metis (mono_tags, lifting) UN_iff X someI_ex subset_iff)
+next
+ fix x i
+ assume "i \<in> I" "x \<in> topspace X \<inter> T i"
then show "f (SOME i. i \<in> I \<and> x \<in> T i) x = f i x"
by (metis (no_types, lifting) IntD2 IntI f someI_ex)
qed
-text\<open>Likewise on closed sets, with a finiteness assumption\<close>
+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. \<not> P x})) Y g"
+ and fg: "\<And>x. x \<in> X frontier_of {x. P x} \<Longrightarrow> f x = g x"
+ shows "continuous_map X Y (\<lambda>x. if P x then f x else g x)"
+proof (rule pasting_lemma_closed)
+ let ?f = "\<lambda>b. if b then f else g"
+ let ?g = "\<lambda>x. if P x then f x else g x"
+ let ?T = "\<lambda>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 \<inter> {x. \<not> P x}"
+ by blast
+ show "?f i x = ?f j x"
+ if "i \<in> {True,False}" "j \<in> {True,False}" and x: "x \<in> topspace X \<inter> ?T i \<inter> ?T j" for i j x
+ proof -
+ have "f x = g x"
+ if "i" "\<not> 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 \<in> X closure_of {x. \<not> P x}" "x \<in> X closure_of Collect P" "\<not> 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 "\<exists>j. j \<in> {True,False} \<and> x \<in> ?T j \<and> (if P x then f x else g x) = ?f j x"
+ if "x \<in> 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 \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::topological_space"
- assumes "finite I"
- and clo: "\<And>i. i \<in> I \<Longrightarrow> closedin (subtopology euclidean S) (T i)"
- and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_on (T i) (f i)"
- and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> S \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
- and g: "\<And>x. x \<in> S \<Longrightarrow> \<exists>j. j \<in> I \<and> x \<in> T j \<and> 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: "\<And>i. i \<in> I \<Longrightarrow> (T i) \<subseteq> S"
- using clo closedin_imp_subset by blast
- have *: "(S \<inter> g -` U) = (\<Union>i \<in> I. T i \<inter> f i -` U)"
- using S f g by fastforce
- show "closedin (subtopology euclidean S) (S \<inter> g -` U)"
- apply (subst *)
- apply (rule closedin_Union)
- using \<open>finite I\<close> apply simp
- apply (blast intro: \<open>closed U\<close> continuous_closedin_preimage cont clo closedin_trans)
- done
+lemma continuous_map_cases_alt:
+ assumes f: "continuous_map (subtopology X (X closure_of {x \<in> topspace X. P x})) Y f"
+ and g: "continuous_map (subtopology X (X closure_of {x \<in> topspace X. ~P x})) Y g"
+ and fg: "\<And>x. x \<in> X frontier_of {x \<in> topspace X. P x} \<Longrightarrow> f x = g x"
+ shows "continuous_map X Y (\<lambda>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 \<in> topspace X. p x \<in> Z closure_of U}) Y f"
+ and contg: "continuous_map (subtopology X {x \<in> topspace X. p x \<in> Z closure_of (topspace Z - U)}) Y g"
+ and fg: "\<And>x. \<lbrakk>x \<in> topspace X; p x \<in> Z frontier_of U\<rbrakk> \<Longrightarrow> f x = g x"
+ shows "continuous_map X Y (\<lambda>x. if p x \<in> U then f x else g x)"
+proof (rule continuous_map_cases_alt)
+ show "continuous_map (subtopology X (X closure_of {x \<in> topspace X. p x \<in> U})) Y f"
+ proof (rule continuous_map_from_subtopology_mono)
+ let ?T = "{x \<in> topspace X. p x \<in> Z closure_of U}"
+ show "continuous_map (subtopology X ?T) Y f"
+ by (simp add: contf)
+ show "X closure_of {x \<in> topspace X. p x \<in> U} \<subseteq> ?T"
+ by (rule continuous_map_closure_preimage_subset [OF contp])
+ qed
+ show "continuous_map (subtopology X (X closure_of {x \<in> topspace X. p x \<notin> U})) Y g"
+ proof (rule continuous_map_from_subtopology_mono)
+ let ?T = "{x \<in> topspace X. p x \<in> Z closure_of (topspace Z - U)}"
+ show "continuous_map (subtopology X ?T) Y g"
+ by (simp add: contg)
+ have "X closure_of {x \<in> topspace X. p x \<notin> U} \<subseteq> X closure_of {x \<in> topspace X. p x \<in> topspace Z - U}"
+ apply (rule closure_of_mono)
+ using continuous_map_closedin contp by fastforce
+ then show "X closure_of {x \<in> topspace X. p x \<notin> U} \<subseteq> ?T"
+ by (rule order_trans [OF _ continuous_map_closure_preimage_subset [OF contp]])
+ qed
+next
+ show "f x = g x" if "x \<in> X frontier_of {x \<in> topspace X. p x \<in> 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 \<Rightarrow> 'a::topological_space \<Rightarrow> 'b::topological_space"
- assumes "finite I"
- and S: "S \<subseteq> (\<Union>i \<in> I. T i)"
- and clo: "\<And>i. i \<in> I \<Longrightarrow> closedin (subtopology euclidean S) (T i)"
- and cont: "\<And>i. i \<in> I \<Longrightarrow> continuous_on (T i) (f i)"
- and f: "\<And>i j x. \<lbrakk>i \<in> I; j \<in> I; x \<in> S \<inter> T i \<inter> T j\<rbrakk> \<Longrightarrow> f i x = f j x"
- obtains g where "continuous_on S g" "\<And>x i. \<lbrakk>i \<in> I; x \<in> S \<inter> T i\<rbrakk> \<Longrightarrow> g x = f i x"
-proof
- show "continuous_on S (\<lambda>x. f (SOME i. i \<in> I \<and> x \<in> T i) x)"
- apply (rule pasting_lemma_closed [OF \<open>finite I\<close> clo cont])
- apply (blast intro: f)+
- apply (metis (mono_tags, lifting) S UN_iff subsetCE someI)
- done
-next
- fix x i
- assume "i \<in> I" "x \<in> S \<inter> T i"
- then show "f (SOME i. i \<in> I \<and> x \<in> T i) x = f i x"
- by (metis (no_types, lifting) IntD2 IntI f someI_ex)
-qed
-
-
subsection \<open>Retractions\<close>
definition%important retraction :: "('a::topological_space) set \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 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 \<subseteq> T" "continuous_on T r" "r ` T \<subseteq> S" "\<And>x. x \<in> S \<Longrightarrow> 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 \<inter> r -` U) \<longleftrightarrow> openin (subtopology euclidean T) U"
+ "openin (top_of_set S) (S \<inter> r -` U) \<longleftrightarrow> openin (top_of_set T) U"
if retraction: "retraction S T r" and "U \<subseteq> 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\<open>Retractions on a topological space\<close>
+
+definition retract_of_space :: "'a set \<Rightarrow> 'a topology \<Rightarrow> bool" (infix "retract'_of'_space" 50)
+ where "S retract_of_space X
+ \<equiv> S \<subseteq> topspace X \<and> (\<exists>r. continuous_map X (subtopology X S) r \<and> (\<forall>x \<in> S. r x = x))"
+
+lemma retract_of_space_retraction_maps:
+ "S retract_of_space X \<longleftrightarrow> S \<subseteq> topspace X \<and> (\<exists>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 \<longleftrightarrow> S \<subseteq> topspace X \<and> 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 \<Longrightarrow> S \<subseteq> 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 \<longleftrightarrow> topspace X = {}"
+ by (auto simp: continuous_map_def retract_of_space_def)
+
+lemma retract_of_space_singleton [simp]:
+ "{a} retract_of_space X \<longleftrightarrow> a \<in> topspace X"
+proof -
+ have "continuous_map X (subtopology X {a}) (\<lambda>x. a) \<and> (\<lambda>x. a) a = a" if "a \<in> 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 = {} \<Longrightarrow> topspace X = {}"
+ shows "S retract_of_space X"
+proof (cases "S = {}")
+ case False
+ then obtain a where "a \<in> S"
+ by blast
+ show ?thesis
+ unfolding retract_of_space_def
+ proof (intro exI conjI)
+ show "S \<subseteq> topspace X"
+ by (simp add: assms closedin_subset)
+ have "continuous_map X X (\<lambda>x. if x \<in> S then x else a)"
+ proof (rule continuous_map_cases)
+ show "continuous_map (subtopology X (X closure_of {x. x \<in> S})) X (\<lambda>x. x)"
+ by (simp add: continuous_map_from_subtopology)
+ show "continuous_map (subtopology X (X closure_of {x. x \<notin> S})) X (\<lambda>x. a)"
+ using \<open>S \<subseteq> topspace X\<close> \<open>a \<in> S\<close> by force
+ show "x = a" if "x \<in> X frontier_of {x. x \<in> S}" for x
+ using assms that clopenin_eq_frontier_of by fastforce
+ qed
+ then show "continuous_map X (subtopology X S) (\<lambda>x. if x \<in> S then x else a)"
+ using \<open>S \<subseteq> topspace X\<close> \<open>a \<in> S\<close> 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 \<union> T = topspace X" and "S = {} \<Longrightarrow> topspace X = {}"
+ shows "S retract_of_space X"
+proof (rule retract_of_space_clopen)
+ have "S \<inter> T = {}"
+ by (meson ST disjnt_def)
+ then have "S = topspace X - T"
+ using ST by auto
+ then show "closedin X S"
+ using \<open>openin X T\<close> 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 \<subseteq> 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
+ \<Longrightarrow> 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\<open>Paths and path-connectedness\<close>
+
+definition pathin :: "'a topology \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> bool" where
+ "pathin X g \<equiv> continuous_map (subtopology euclideanreal {0..1}) X g"
+
+lemma pathin_compose:
+ "\<lbrakk>pathin X g; continuous_map X Y f\<rbrakk> \<Longrightarrow> pathin Y (f \<circ> g)"
+ by (simp add: continuous_map_compose pathin_def)
+
+lemma pathin_subtopology:
+ "pathin (subtopology X S) g \<longleftrightarrow> pathin X g \<and> (\<forall>x \<in> {0..1}. g x \<in> S)"
+ by (auto simp: pathin_def continuous_map_in_subtopology)
+
+lemma pathin_const:
+ "pathin X (\<lambda>x. a) \<longleftrightarrow> a \<in> topspace X"
+ by (simp add: pathin_def)
+
+definition path_connected_space :: "'a topology \<Rightarrow> bool"
+ where "path_connected_space X \<equiv> \<forall>x \<in> topspace X. \<forall> y \<in> topspace X. \<exists>g. pathin X g \<and> g 0 = x \<and> g 1 = y"
+
+definition path_connectedin :: "'a topology \<Rightarrow> 'a set \<Rightarrow> bool"
+ where "path_connectedin X S \<equiv> S \<subseteq> topspace X \<and> path_connected_space(subtopology X S)"
+
+lemma path_connectedin_absolute [simp]:
+ "path_connectedin (subtopology X S) S \<longleftrightarrow> path_connectedin X S"
+ by (simp add: path_connectedin_def subtopology_subtopology topspace_subtopology)
+
+lemma path_connectedin_subset_topspace:
+ "path_connectedin X S \<Longrightarrow> S \<subseteq> topspace X"
+ by (simp add: path_connectedin_def)
+
+lemma path_connectedin_subtopology:
+ "path_connectedin (subtopology X S) T \<longleftrightarrow> path_connectedin X T \<and> T \<subseteq> S"
+ by (auto simp: path_connectedin_def subtopology_subtopology topspace_subtopology inf.absorb2)
+
+lemma path_connectedin:
+ "path_connectedin X S \<longleftrightarrow>
+ S \<subseteq> topspace X \<and>
+ (\<forall>x \<in> S. \<forall>y \<in> S. \<exists>g. pathin X g \<and> g ` {0..1} \<subseteq> S \<and> g 0 = x \<and> 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) \<longleftrightarrow> 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 *: "\<exists>S. connectedin X S \<and> g 0 \<in> S \<and> g 1 \<in> 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 \<Longrightarrow> connectedin X S"
+ by (simp add: connectedin_def path_connected_imp_connected_space path_connectedin_def)
+
+lemma path_connected_space_topspace_empty:
+ "topspace X = {} \<Longrightarrow> 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} \<longleftrightarrow> a \<in> topspace X"
+proof
+ show "path_connectedin X {a} \<Longrightarrow> a \<in> topspace X"
+ by (simp add: path_connectedin)
+ show "a \<in> topspace X \<Longrightarrow> 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) \<subseteq> topspace Y"
+ by (metis f continuous_map_image_subset_topspace)
+ show ?thesis
+ unfolding path_connectedin
+ proof (intro conjI ballI; clarify?)
+ fix x
+ assume "x \<in> S"
+ show "f x \<in> topspace Y"
+ by (meson S fX \<open>x \<in> S\<close> image_subset_iff path_connectedin_subset_topspace set_mp)
+ next
+ fix x y
+ assume "x \<in> S" and "y \<in> S"
+ then obtain g where g: "pathin X g" "g ` {0..1} \<subseteq> S" "g 0 = x" "g 1 = y"
+ using S by (force simp: path_connectedin)
+ show "\<exists>g. pathin Y g \<and> g ` {0..1} \<subseteq> f ` S \<and> g 0 = f x \<and> g 1 = f y"
+ proof (intro exI conjI)
+ show "pathin Y (f \<circ> g)"
+ using \<open>pathin X g\<close> f pathin_compose by auto
+ qed (use g in auto)
+ qed
+qed
+
+lemma homeomorphic_path_connected_space_imp:
+ "\<lbrakk>path_connected_space X; X homeomorphic_space Y\<rbrakk> \<Longrightarrow> 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 \<Longrightarrow> path_connected_space X \<longleftrightarrow> 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 \<subseteq> topspace X"
+ shows "path_connectedin Y (f ` U) \<longleftrightarrow> path_connectedin X U"
+ unfolding path_connectedin_def
+proof (intro conj_cong homeomorphic_path_connected_space)
+ show "(f ` U \<subseteq> topspace Y) = (U \<subseteq> topspace X)"
+ using assms homeomorphic_imp_surjective_map by blast
+next
+ assume "U \<subseteq> 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 \<Longrightarrow> path_connectedin X U \<longleftrightarrow> U \<subseteq> topspace X \<and> path_connectedin Y (f ` U)"
+ by (meson homeomorphic_map_path_connectedness path_connectedin_def)
+
+subsection\<open>Connected components\<close>
+
+definition connected_component_of :: "'a topology \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
+ where "connected_component_of X x y \<equiv>
+ \<exists>T. connectedin X T \<and> x \<in> T \<and> y \<in> T"
+
+abbreviation connected_component_of_set
+ where "connected_component_of_set X x \<equiv> Collect (connected_component_of X x)"
+
+definition connected_components_of :: "'a topology \<Rightarrow> ('a set) set"
+ where "connected_components_of X \<equiv> connected_component_of_set X ` topspace X"
+
+lemma connected_component_in_topspace:
+ "connected_component_of X x y \<Longrightarrow> x \<in> topspace X \<and> y \<in> topspace X"
+ by (meson connected_component_of_def connectedin_subset_topspace in_mono)
+
+lemma connected_component_of_refl:
+ "connected_component_of X x x \<longleftrightarrow> x \<in> 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 \<longleftrightarrow> connected_component_of X y x"
+ by (meson connected_component_of_def)
+
+lemma connected_component_of_trans:
+ "\<lbrakk>connected_component_of X x y; connected_component_of X y z\<rbrakk>
+ \<Longrightarrow> connected_component_of X x z"
+ unfolding connected_component_of_def
+ using connectedin_Un by fastforce
+
+lemma connected_component_of_mono:
+ "\<lbrakk>connected_component_of (subtopology X S) x y; S \<subseteq> T\<rbrakk>
+ \<Longrightarrow> 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. \<exists>T. connectedin X T \<and> x \<in> T \<and> y \<in> T}"
+ by (meson connected_component_of_def)
+
+lemma connected_component_of_subset_topspace:
+ "connected_component_of_set X x \<subseteq> topspace X"
+ using connected_component_in_topspace by force
+
+lemma connected_component_of_eq_empty:
+ "connected_component_of_set X x = {} \<longleftrightarrow> (x \<notin> topspace X)"
+ using connected_component_in_topspace connected_component_of_refl by fastforce
+
+lemma connected_space_iff_connected_component:
+ "connected_space X \<longleftrightarrow> (\<forall>x \<in> topspace X. \<forall>y \<in> 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:
+ "\<lbrakk>connected_space X; a \<in> topspace X; b \<in> topspace X\<rbrakk>
+ \<Longrightarrow> connected_component_of X a b"
+ by (simp add: connected_space_iff_connected_component)
+
+lemma connected_space_connected_component_set:
+ "connected_space X \<longleftrightarrow> (\<forall>x \<in> 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:
+ "\<lbrakk>connectedin X S; x \<in> S\<rbrakk> \<Longrightarrow> S \<subseteq> 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 \<longleftrightarrow>
+ x \<in> topspace X \<and> y \<in> topspace X \<and> 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)
+ \<longleftrightarrow> ~(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 \<longleftrightarrow>
+ (x \<notin> topspace X) \<and> (y \<notin> topspace X) \<or>
+ x \<in> topspace X \<and> y \<in> topspace X \<and>
+ 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 = \<Union> {T. connectedin X T \<and> x \<in> 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:
+ "\<Union>(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:
+ "\<lbrakk>C \<in> connected_components_of X; connectedin X S; ~disjnt C S\<rbrakk> \<Longrightarrow> S \<subseteq> 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 \<in> connected_components_of X
+ \<Longrightarrow> topspace X - C = \<Union> (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 \<in> connected_components_of X \<Longrightarrow> C \<noteq> {}"
+ unfolding connected_components_of_def
+ by (metis (no_types, lifting) connected_component_of_eq_empty imageE)
+
+lemma connected_components_of_subset:
+ "C \<in> connected_components_of X \<Longrightarrow> C \<subseteq> topspace X"
+ using Union_connected_components_of by fastforce
+
+lemma connectedin_connected_components_of:
+ assumes "C \<in> connected_components_of X"
+ shows "connectedin X C"
+proof -
+ have "C \<in> 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 \<in> connected_components_of X \<longleftrightarrow> a \<in> 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 \<longleftrightarrow> (\<forall>C \<in> connected_components_of X. \<forall>C' \<in> 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 = {} \<longleftrightarrow> topspace X = {}"
+ by (simp add: connected_components_of_def)
+
+lemma connected_components_of_empty_space:
+ "topspace X = {} \<Longrightarrow> connected_components_of X = {}"
+ by (simp add: connected_components_of_eq_empty)
+
+lemma connected_components_of_subset_sing:
+ "connected_components_of X \<subseteq> {S} \<longleftrightarrow> connected_space X \<and> (topspace X = {} \<or> 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 \<longleftrightarrow> (\<exists>a. connected_components_of X \<subseteq> {a})"
+ by (simp add: connected_components_of_subset_sing)
+
+lemma connected_components_of_eq_singleton:
+ "connected_components_of X = {S}
+\<longleftrightarrow> connected_space X \<and> topspace X \<noteq> {} \<and> 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 \<Longrightarrow> 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 \<noteq> {}"
+ shows "\<exists>C. C \<in> connected_components_of X \<and> S \<subseteq> 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 \<in> connected_components_of X"
+ shows "closedin X C"
+proof -
+ obtain x where "x \<in> 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 \<subseteq> topspace X"
+ by (simp add: connected_component_of_subset_topspace)
+ moreover have "X closure_of connected_component_of_set X x \<subseteq> 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 \<in> X closure_of connected_component_of_set X x"
+ by (simp add: \<open>x \<in> topspace X\<close> 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 \<longleftrightarrow>
+ (x \<notin> topspace X) \<and> (y \<notin> topspace X) \<or>
+ ~(connected_component_of_set X x \<inter> connected_component_of_set X y = {})"
+ using connected_component_of_equiv by fastforce
+
+lemma connected_component_of_nonoverlap:
+ "connected_component_of_set X x \<inter> connected_component_of_set X y = {} \<longleftrightarrow>
+ (x \<notin> topspace X) \<or> (y \<notin> topspace X) \<or>
+ ~(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 \<inter> connected_component_of_set X y = {}) \<longleftrightarrow>
+ x \<in> topspace X \<and> y \<in> topspace X \<and>
+ 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
--- a/src/HOL/Analysis/Analysis.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Analysis.thy Tue Mar 19 16:14:59 2019 +0000
@@ -27,7 +27,7 @@
Homeomorphism
Bounded_Continuous_Function
Abstract_Topology
- Function_Topology
+ Product_Topology
T1_Spaces
Infinite_Products
Infinite_Set_Sum
--- a/src/HOL/Analysis/Arcwise_Connected.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Arcwise_Connected.thy Tue Mar 19 16:14:59 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 \<noteq> {}"
+ assumes "open S" and opeSV: "openin (top_of_set (frontier S)) V" and "V \<noteq> {}"
obtains g where "arc g" "g ` {0..<1} \<subseteq> S" "pathstart g \<in> S" "pathfinish g \<in> V"
proof -
obtain z where "z \<in> 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 \<in> S" "V \<noteq> {}"
- 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} \<subseteq> S" "pathstart g = x" "pathfinish g \<in> V"
proof -
have "V \<subseteq> 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 \<noteq> {}" "\<not> U \<subseteq> V"
obtains g where "arc g" "pathstart g \<in> U" "pathfinish g \<in> V" "g ` {0<..<1} \<subseteq> S"
proof -
@@ -2150,7 +2150,7 @@
proof (rule dense_accessible_frontier_points_connected [OF S \<open>x \<in> S\<close>])
show "U - {pathfinish g} \<noteq> {}"
using \<open>pathfinish g \<in> V\<close> \<open>\<not> U \<subseteq> V\<close> 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 \<gamma> where "arc \<gamma>"
@@ -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 \<noteq> {}" "V \<noteq> {}" "U \<noteq> V"
obtains g where "arc g" "pathstart g \<in> U" "pathfinish g \<in> V" "g ` {0<..<1} \<subseteq> S"
proof -
--- a/src/HOL/Analysis/Brouwer_Fixpoint.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Brouwer_Fixpoint.thy Tue Mar 19 16:14:59 2019 +0000
@@ -85,7 +85,7 @@
qed
lemma retraction_imp_quotient_map:
- "openin (subtopology euclidean S) (S \<inter> r -` U) \<longleftrightarrow> openin (subtopology euclidean T) U"
+ "openin (top_of_set S) (S \<inter> r -` U) \<longleftrightarrow> openin (top_of_set T) U"
if retraction: "retraction S T r" and "U \<subseteq> 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 \<Rightarrow> bool" where
"AR S \<equiv> \<forall>U. \<forall>S'::('a * real) set.
- S homeomorphic S' \<and> closedin (subtopology euclidean U) S' \<longrightarrow> S' retract_of U"
+ S homeomorphic S' \<and> closedin (top_of_set U) S' \<longrightarrow> S' retract_of U"
definition%important ANR :: "'a::topological_space set \<Rightarrow> bool" where
"ANR S \<equiv> \<forall>U. \<forall>S'::('a * real) set.
- S homeomorphic S' \<and> closedin (subtopology euclidean U) S'
- \<longrightarrow> (\<exists>T. openin (subtopology euclidean U) T \<and> S' retract_of T)"
+ S homeomorphic S' \<and> closedin (top_of_set U) S'
+ \<longrightarrow> (\<exists>T. openin (top_of_set U) T \<and> S' retract_of T)"
definition%important ENR :: "'a::topological_space set \<Rightarrow> bool" where
"ENR S \<equiv> \<exists>U. open U \<and> S retract_of U"
@@ -248,14 +248,14 @@
lemma AR_imp_absolute_extensor:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "AR S" and contf: "continuous_on T f" and "f ` T \<subseteq> 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 \<subseteq> S" "\<And>x. x \<in> T \<Longrightarrow> g x = f x"
proof -
have "aff_dim S < int (DIM('b \<times> real))"
using aff_dim_le_DIM [of S] by simp
then obtain C and S' :: "('b * real) set"
where C: "convex C" "C \<noteq> {}"
- 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 "\<And>f :: 'a * real \<Rightarrow> 'a.
\<And>U T. \<lbrakk>continuous_on T f; f ` T \<subseteq> S;
- closedin (subtopology euclidean U) T\<rbrakk>
+ closedin (top_of_set U) T\<rbrakk>
\<Longrightarrow> \<exists>g. continuous_on U g \<and> g ` U \<subseteq> S \<and> (\<forall>x \<in> 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 \<subseteq> S"
@@ -369,7 +369,7 @@
shows "AR S \<longleftrightarrow>
(\<forall>f :: 'a * real \<Rightarrow> 'a.
\<forall>U T. continuous_on T f \<longrightarrow> f ` T \<subseteq> S \<longrightarrow>
- closedin (subtopology euclidean U) T \<longrightarrow>
+ closedin (top_of_set U) T \<longrightarrow>
(\<exists>g. continuous_on U g \<and> g ` U \<subseteq> S \<and> (\<forall>x \<in> 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 \<and> closedin (subtopology euclidean U) S"
+ assumes "AR S \<and> 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 \<Rightarrow> 'b::euclidean_space"
assumes "ANR S" and contf: "continuous_on T f" and "f ` T \<subseteq> S"
- and cloUT: "closedin (subtopology euclidean U) T"
- obtains V g where "T \<subseteq> V" "openin (subtopology euclidean U) V"
+ and cloUT: "closedin (top_of_set U) T"
+ obtains V g where "T \<subseteq> V" "openin (top_of_set U) V"
"continuous_on V g"
"g ` V \<subseteq> S" "\<And>x. x \<in> T \<Longrightarrow> 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 \<noteq> {}"
- 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 \<open>ANR S\<close> by (auto simp: ANR_def)
then obtain r where "S' \<subseteq> D" and contr: "continuous_on D r"
and "r ` D \<subseteq> S'" and rid: "\<And>x. x \<in> S' \<Longrightarrow> r x = x"
@@ -437,7 +437,7 @@
show "T \<subseteq> U \<inter> f' -` D"
using cloUT closedin_imp_subset \<open>S' \<subseteq> D\<close> \<open>f ` T \<subseteq> S\<close> eq homeomorphism_image1 homgh
by fastforce
- show ope: "openin (subtopology euclidean U) (U \<inter> f' -` D)"
+ show ope: "openin (top_of_set U) (U \<inter> f' -` D)"
using \<open>f' ` U \<subseteq> C\<close> by (auto simp: opD contf' continuous_openin_preimage)
have conth: "continuous_on (r ` f' ` (U \<inter> 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 \<open>ANR S\<close> h clo]
- obtain V h' where "S' \<subseteq> V" and opUV: "openin (subtopology euclidean U) V"
+ obtain V h' where "S' \<subseteq> V" and opUV: "openin (top_of_set U) V"
and h': "continuous_on V h'" "h' ` V \<subseteq> S"
and h'h:"\<And>x. x \<in> S' \<Longrightarrow> h' x = h x"
by (blast intro: ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> h clo])
@@ -508,20 +508,20 @@
fixes S :: "'a::euclidean_space set"
assumes "\<And>f :: 'a * real \<Rightarrow> 'a.
\<And>U T. \<lbrakk>continuous_on T f; f ` T \<subseteq> S;
- closedin (subtopology euclidean U) T\<rbrakk>
- \<Longrightarrow> \<exists>V g. T \<subseteq> V \<and> openin (subtopology euclidean U) V \<and>
+ closedin (top_of_set U) T\<rbrakk>
+ \<Longrightarrow> \<exists>V g. T \<subseteq> V \<and> openin (top_of_set U) V \<and>
continuous_on V g \<and> g ` V \<subseteq> S \<and> (\<forall>x \<in> 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 \<subseteq> S"
using hom homeomorphism_def apply blast
apply (metis hom equalityE homeomorphism_def)
done
- obtain V h' where "T \<subseteq> V" and opV: "openin (subtopology euclidean U) V"
+ obtain V h' where "T \<subseteq> V" and opV: "openin (top_of_set U) V"
and h': "continuous_on V h'" "h' ` V \<subseteq> S"
and h'h: "\<forall>x\<in>T. h' x = h x"
using assms [OF h clo] by blast
@@ -538,7 +538,7 @@
show "\<forall>x\<in>T. (g \<circ> h') x = x"
by clarsimp (metis h'h hom homeomorphism_def)
qed
- then show "\<exists>V. openin (subtopology euclidean U) V \<and> T retract_of V"
+ then show "\<exists>V. openin (top_of_set U) V \<and> T retract_of V"
using opV by blast
qed
@@ -547,8 +547,8 @@
shows "ANR S \<longleftrightarrow>
(\<forall>f :: 'a * real \<Rightarrow> 'a.
\<forall>U T. continuous_on T f \<longrightarrow> f ` T \<subseteq> S \<longrightarrow>
- closedin (subtopology euclidean U) T \<longrightarrow>
- (\<exists>V g. T \<subseteq> V \<and> openin (subtopology euclidean U) V \<and>
+ closedin (top_of_set U) T \<longrightarrow>
+ (\<exists>V g. T \<subseteq> V \<and> openin (top_of_set U) V \<and>
continuous_on V g \<and> g ` V \<subseteq> S \<and> (\<forall>x \<in> 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' \<subseteq> V" "V \<subseteq> 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' \<inter> (U - Z) = {}"
using \<open>S' retract_of Z\<close> 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' \<subseteq> V" "U - Z \<subseteq> W" "V \<inter> 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 \<subseteq> V" "V \<subseteq> 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' \<subseteq> 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 \<open>ENR S\<close> 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 \<open>S retract_of X\<close> \<open>open X\<close> 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 \<inter> W" in that, blast)
@@ -752,7 +752,7 @@
proof
assume ?lhs
obtain C and S' :: "('a * real) set"
- where "convex C" "C \<noteq> {}" "closedin (subtopology euclidean C) S'" "S homeomorphic S'"
+ where "convex C" "C \<noteq> {}" "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 \<open>AR S\<close> 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 "\<exists>g. continuous_on W g \<and> g ` W \<subseteq> S \<and> (\<forall>x\<in>T. g x = f x)"
if f: "continuous_on T f" "f ` T \<subseteq> S"
- and WT: "closedin (subtopology euclidean W) T"
+ and WT: "closedin (top_of_set W) T"
for W T and f :: "'a \<times> real \<Rightarrow> 'a"
proof -
obtain U g
- where "T \<subseteq> U" and WU: "openin (subtopology euclidean W) U"
+ where "T \<subseteq> U" and WU: "openin (top_of_set W) U"
and contg: "continuous_on U g"
and "g ` U \<subseteq> S" and gf: "\<And>x. x \<in> T \<Longrightarrow> g x = f x"
using iffD1 [OF ANR_eq_absolute_neighbourhood_extensor \<open>ANR S\<close>, 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) \<inter> T = {}"
using \<open>T \<subseteq> U\<close> 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 \<subseteq> V'" "T \<subseteq> V" "V' \<inter> V = {}"
using separation_normal_local [of W "W-U" T] WT by blast
then have WVT: "T \<inter> (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 \<times> real \<Rightarrow> 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 \<inter> T)"
shows "(S \<union> T) retract_of U"
proof -
@@ -943,10 +943,10 @@
define S' where "S' \<equiv> {x \<in> U. setdist {x} S \<le> setdist {x} T}"
define T' where "T' \<equiv> {x \<in> U. setdist {x} T \<le> setdist {x} S}"
define W where "W \<equiv> {x \<in> 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 "\<lambda>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 "\<lambda>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 \<subseteq> S'"
@@ -971,7 +971,7 @@
by (force simp: W_def setdist_sing_in_set)
have "S' \<inter> 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 \<equiv> \<lambda>x. if x \<in> W then r0 x else x"
have "r ` (W \<union> S) \<subseteq> S" "r ` (W \<union> T) \<subseteq> T"
@@ -979,14 +979,14 @@
have contr: "continuous_on (W \<union> (S \<union> T)) r"
unfolding r_def
proof (rule continuous_on_cases_local [OF _ _ contr0 continuous_on_id])
- show "closedin (subtopology euclidean (W \<union> (S \<union> T))) W"
- using \<open>S \<subseteq> U\<close> \<open>T \<subseteq> U\<close> \<open>W \<subseteq> U\<close> \<open>closedin (subtopology euclidean U) W\<close> closedin_subset_trans by fastforce
- show "closedin (subtopology euclidean (W \<union> (S \<union> T))) (S \<union> T)"
+ show "closedin (top_of_set (W \<union> (S \<union> T))) W"
+ using \<open>S \<subseteq> U\<close> \<open>T \<subseteq> U\<close> \<open>W \<subseteq> U\<close> \<open>closedin (top_of_set U) W\<close> closedin_subset_trans by fastforce
+ show "closedin (top_of_set (W \<union> (S \<union> T))) (S \<union> T)"
by (meson \<open>S \<subseteq> U\<close> \<open>T \<subseteq> U\<close> \<open>W \<subseteq> U\<close> assms closedin_Un closedin_subset_trans sup.bounded_iff sup.cobounded2)
show "\<And>x. x \<in> W \<and> x \<notin> W \<or> x \<in> S \<union> T \<and> x \<in> W \<Longrightarrow> r0 x = x"
by (auto simp: ST)
qed
- have cloUWS: "closedin (subtopology euclidean U) (W \<union> S)"
+ have cloUWS: "closedin (top_of_set U) (W \<union> S)"
by (simp add: cloUW assms closedin_Un)
obtain g where contg: "continuous_on U g"
and "g ` U \<subseteq> S" and geqr: "\<And>x. x \<in> W \<union> S \<Longrightarrow> g x = r x"
@@ -994,7 +994,7 @@
apply (rule continuous_on_subset [OF contr])
using \<open>r ` (W \<union> S) \<subseteq> S\<close> apply auto
done
- have cloUWT: "closedin (subtopology euclidean U) (W \<union> T)"
+ have cloUWT: "closedin (top_of_set U) (W \<union> T)"
by (simp add: cloUW assms closedin_Un)
obtain h where conth: "continuous_on U h"
and "h ` U \<subseteq> T" and heqr: "\<And>x. x \<in> W \<union> T \<Longrightarrow> 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 \<union> T)) S"
- and STT: "closedin (subtopology euclidean (S \<union> T)) T"
+ assumes STS: "closedin (top_of_set (S \<union> T)) S"
+ and STT: "closedin (top_of_set (S \<union> T)) T"
and "AR S" "AR T" "AR(S \<inter> T)"
shows "AR(S \<union> T)"
proof -
have "C retract_of U"
- if hom: "S \<union> T homeomorphic C" and UC: "closedin (subtopology euclidean U) C"
+ if hom: "S \<union> 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 \<union> T) C f g"
using hom by (force simp: homeomorphic_def)
- have US: "closedin (subtopology euclidean U) (C \<inter> g -` S)"
+ have US: "closedin (top_of_set U) (C \<inter> 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 \<inter> g -` T)"
+ have UT: "closedin (top_of_set U) (C \<inter> 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 \<inter> T)"
- obtains V where "openin (subtopology euclidean U) V" "(S \<union> T) retract_of V"
+ obtains V where "openin (top_of_set U) V" "(S \<union> T) retract_of V"
proof (cases "S = {} \<or> 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' \<equiv> {x \<in> U. setdist {x} S \<le> setdist {x} T}"
define T' where "T' \<equiv> {x \<in> U. setdist {x} T \<le> setdist {x} S}"
define W where "W \<equiv> {x \<in> 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 "\<lambda>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 "\<lambda>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 \<subseteq> S'"
@@ -1123,17 +1123,17 @@
using \<open>S \<subseteq> U\<close> by (force simp: W_def setdist_sing_in_set)+
have "S' \<inter> 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 \<inter> T \<subseteq> W'" "W' \<subseteq> W0"
and ret: "(S \<inter> T) retract_of W0"
apply (rule ANR_imp_closed_neighbourhood_retract [OF \<open>ANR(S \<inter> T)\<close>])
apply (rule closedin_subset_trans [of U, OF _ ST_W \<open>W \<subseteq> U\<close>])
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 \<inter> T \<subseteq> U0" "U0 \<inter> W \<subseteq> W0"
unfolding openin_open using \<open>W \<subseteq> U\<close> by blast
have "W0 \<subseteq> U"
@@ -1150,29 +1150,29 @@
have contr: "continuous_on (W0 \<union> (S \<union> T)) r"
unfolding r_def
proof (rule continuous_on_cases_local [OF _ _ contr0 continuous_on_id])
- show "closedin (subtopology euclidean (W0 \<union> (S \<union> T))) W0"
+ show "closedin (top_of_set (W0 \<union> (S \<union> T))) W0"
apply (rule closedin_subset_trans [of U])
using cloWW0 cloUW closedin_trans \<open>W0 \<subseteq> U\<close> \<open>S \<subseteq> U\<close> \<open>T \<subseteq> U\<close> apply blast+
done
- show "closedin (subtopology euclidean (W0 \<union> (S \<union> T))) (S \<union> T)"
+ show "closedin (top_of_set (W0 \<union> (S \<union> T))) (S \<union> T)"
by (meson \<open>S \<subseteq> U\<close> \<open>T \<subseteq> U\<close> \<open>W0 \<subseteq> U\<close> assms closedin_Un closedin_subset_trans sup.bounded_iff sup.cobounded2)
show "\<And>x. x \<in> W0 \<and> x \<notin> W0 \<or> x \<in> S \<union> T \<and> x \<in> W0 \<Longrightarrow> r0 x = x"
using ST cloWW0 closedin_subset by fastforce
qed
- have cloS'WS: "closedin (subtopology euclidean S') (W0 \<union> S)"
+ have cloS'WS: "closedin (top_of_set S') (W0 \<union> S)"
by (meson closedin_subset_trans US cloUS' \<open>S \<subseteq> S'\<close> \<open>W \<subseteq> S'\<close> cloUW cloWW0
closedin_Un closedin_imp_subset closedin_trans)
obtain W1 g where "W0 \<union> S \<subseteq> 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 \<subseteq> S" and geqr: "\<And>x. x \<in> W0 \<union> S \<Longrightarrow> g x = r x"
apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR S\<close> _ \<open>r ` (W0 \<union> S) \<subseteq> S\<close> cloS'WS])
apply (rule continuous_on_subset [OF contr], blast+)
done
- have cloT'WT: "closedin (subtopology euclidean T') (W0 \<union> T)"
+ have cloT'WT: "closedin (top_of_set T') (W0 \<union> T)"
by (meson closedin_subset_trans UT cloUT' \<open>T \<subseteq> T'\<close> \<open>W \<subseteq> T'\<close> cloUW cloWW0
closedin_Un closedin_imp_subset closedin_trans)
obtain W2 h where "W0 \<union> T \<subseteq> 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 \<subseteq> T" and heqr: "\<And>x. x \<in> W0 \<union> T \<Longrightarrow> h x = r x"
apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR T\<close> _ \<open>r ` (W0 \<union> T) \<subseteq> T\<close> cloT'WT])
apply (rule continuous_on_subset [OF contr], blast+)
@@ -1187,17 +1187,17 @@
((U - T') \<inter> O1 \<union> (U - S') \<inter> O2 \<union> U \<inter> O1 \<inter> O2) - (W - U0)"
using \<open>U0 \<inter> W \<subseteq> W0\<close> \<open>W0 \<union> S \<subseteq> W1\<close> \<open>W0 \<union> T \<subseteq> W2\<close>
by (auto simp: \<open>S' \<union> T' = U\<close> [symmetric] \<open>S' \<inter> T' = W\<close> [symmetric] \<open>W1 = S' \<inter> O1\<close> \<open>W2 = T' \<inter> O2\<close>)
- show "openin (subtopology euclidean U) (W1 - (W - U0) \<union> (W2 - (W - U0)))"
+ show "openin (top_of_set U) (W1 - (W - U0) \<union> (W2 - (W - U0)))"
apply (subst eq)
apply (intro openin_Un openin_Int_open openin_diff closedin_diff cloUW opeUU0 cloUS' cloUT' \<open>open O1\<close> \<open>open O2\<close>, simp_all)
done
- have cloW1: "closedin (subtopology euclidean (W1 - (W - U0) \<union> (W2 - (W - U0)))) (W1 - (W - U0))"
+ have cloW1: "closedin (top_of_set (W1 - (W - U0) \<union> (W2 - (W - U0)))) (W1 - (W - U0))"
using cloUS' apply (simp add: closedin_closed)
apply (erule ex_forward)
using U0 \<open>W0 \<union> S \<subseteq> W1\<close>
apply (auto simp: \<open>W1 = S' \<inter> O1\<close> \<open>W2 = T' \<inter> O2\<close> \<open>S' \<union> T' = U\<close> [symmetric]\<open>S' \<inter> T' = W\<close> [symmetric])
done
- have cloW2: "closedin (subtopology euclidean (W1 - (W - U0) \<union> (W2 - (W - U0)))) (W2 - (W - U0))"
+ have cloW2: "closedin (top_of_set (W1 - (W - U0) \<union> (W2 - (W - U0)))) (W2 - (W - U0))"
using cloUT' apply (simp add: closedin_closed)
apply (erule ex_forward)
using U0 \<open>W0 \<union> T \<subseteq> W2\<close>
@@ -1227,24 +1227,24 @@
lemma ANR_closed_Un_local:
fixes S :: "'a::euclidean_space set"
- assumes STS: "closedin (subtopology euclidean (S \<union> T)) S"
- and STT: "closedin (subtopology euclidean (S \<union> T)) T"
+ assumes STS: "closedin (top_of_set (S \<union> T)) S"
+ and STT: "closedin (top_of_set (S \<union> T)) T"
and "ANR S" "ANR T" "ANR(S \<inter> T)"
shows "ANR(S \<union> T)"
proof -
- have "\<exists>T. openin (subtopology euclidean U) T \<and> C retract_of T"
- if hom: "S \<union> T homeomorphic C" and UC: "closedin (subtopology euclidean U) C"
+ have "\<exists>T. openin (top_of_set U) T \<and> C retract_of T"
+ if hom: "S \<union> 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 \<union> T) C f g"
using hom by (force simp: homeomorphic_def)
- have US: "closedin (subtopology euclidean U) (C \<inter> g -` S)"
+ have US: "closedin (top_of_set U) (C \<inter> 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 \<inter> g -` T)"
+ have UT: "closedin (top_of_set U) (C \<inter> 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 \<times> real \<Rightarrow> 'a" and U C
assume contf: "continuous_on C f" and fim: "f ` C \<subseteq> S"
- and cloUC: "closedin (subtopology euclidean U) C"
+ and cloUC: "closedin (top_of_set U) C"
have "f ` C \<subseteq> T"
using fim opeTS openin_imp_subset by blast
obtain W g where "C \<subseteq> 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 \<subseteq> T"
and geq: "\<And>x. x \<in> C \<Longrightarrow> g x = f x"
apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR T\<close> contf \<open>f ` C \<subseteq> T\<close> cloUC])
using fim by auto
- show "\<exists>V g. C \<subseteq> V \<and> openin (subtopology euclidean U) V \<and> continuous_on V g \<and> g ` V \<subseteq> S \<and> (\<forall>x\<in>C. g x = f x)"
+ show "\<exists>V g. C \<subseteq> V \<and> openin (top_of_set U) V \<and> continuous_on V g \<and> g ` V \<subseteq> S \<and> (\<forall>x\<in>C. g x = f x)"
proof (intro exI conjI)
show "C \<subseteq> W \<inter> g -` S"
using \<open>C \<subseteq> W\<close> fim geq by blast
- show "openin (subtopology euclidean U) (W \<inter> g -` S)"
+ show "openin (top_of_set U) (W \<inter> g -` S)"
by (metis (mono_tags, lifting) UW contg continuous_openin_preimage gim opeTS openin_trans)
show "continuous_on (W \<inter> 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 \<times> real) set"
where "convex U" "U \<noteq> {}"
- 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 \<noteq> {}" for V
+ if "openin (top_of_set U) V" "T retract_of V" "U \<noteq> {}" for V
using \<open>S homeomorphic T\<close> 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 \<times> 'b) \<times> real \<Rightarrow> 'a \<times> 'b" and U C
assume "continuous_on C f" and fim: "f ` C \<subseteq> S \<times> T"
- and cloUC: "closedin (subtopology euclidean U) C"
+ and cloUC: "closedin (top_of_set U) C"
have contf1: "continuous_on C (fst \<circ> f)"
by (simp add: \<open>continuous_on C f\<close> continuous_on_fst)
obtain W1 g where "C \<subseteq> 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 \<subseteq> S"
and geq: "\<And>x. x \<in> C \<Longrightarrow> g x = (fst \<circ> f) x"
@@ -1508,7 +1508,7 @@
have contf2: "continuous_on C (snd \<circ> f)"
by (simp add: \<open>continuous_on C f\<close> continuous_on_snd)
obtain W2 h where "C \<subseteq> 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 \<subseteq> T"
and heq: "\<And>x. x \<in> C \<Longrightarrow> h x = (snd \<circ> f) x"
@@ -1516,12 +1516,12 @@
using fim apply auto
done
show "\<exists>V g. C \<subseteq> V \<and>
- openin (subtopology euclidean U) V \<and>
+ openin (top_of_set U) V \<and>
continuous_on V g \<and> g ` V \<subseteq> S \<times> T \<and> (\<forall>x\<in>C. g x = f x)"
proof (intro exI conjI)
show "C \<subseteq> W1 \<inter> W2"
by (simp add: \<open>C \<subseteq> W1\<close> \<open>C \<subseteq> W2\<close>)
- show "openin (subtopology euclidean U) (W1 \<inter> W2)"
+ show "openin (top_of_set U) (W1 \<inter> W2)"
by (simp add: UW1 UW2 openin_Int)
show "continuous_on (W1 \<inter> W2) (\<lambda>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 \<noteq> a}" .
- moreover have "openin (subtopology euclidean UNIV) (UNIV \<inter> closest_point (affine hull S) -` (- {a}))"
+ moreover have "openin (top_of_set UNIV) (UNIV \<inter> 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 "\<lbrakk>ENR S; ENR T; ENR(S \<inter> T);
- closedin (subtopology euclidean (S \<union> T)) S; closedin (subtopology euclidean (S \<union> T)) T\<rbrakk>
+ closedin (top_of_set (S \<union> T)) S; closedin (top_of_set (S \<union> T)) T\<rbrakk>
\<Longrightarrow> ENR(S \<union> 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 \<union> T)) S"
- and clT: "closedin (subtopology euclidean (S \<union> T)) T"
+ assumes clS: "closedin (top_of_set (S \<union> T)) S"
+ and clT: "closedin (top_of_set (S \<union> T)) T"
and Un: "(S \<union> T) retract_of U" and Int: "(S \<inter> 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 \<union> T)) S"
- and clT: "closedin (subtopology euclidean (S \<union> T)) T"
+ assumes clS: "closedin (top_of_set (S \<union> T)) S"
+ and clT: "closedin (top_of_set (S \<union> T)) T"
and Un: "AR(S \<union> T)" and Int: "AR(S \<inter> 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 \<union> T)) S"
- and "closedin (subtopology euclidean (S \<union> T)) T"
+ assumes "closedin (top_of_set (S \<union> T)) S"
+ and "closedin (top_of_set (S \<union> T)) T"
and "AR(S \<union> T)" "AR(S \<inter> 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 \<union> T)) S"
- and clT: "closedin (subtopology euclidean (S \<union> T)) T"
+ assumes clS: "closedin (top_of_set (S \<union> T)) S"
+ and clT: "closedin (top_of_set (S \<union> T)) T"
and Un: "ANR(S \<union> T)" and Int: "ANR(S \<inter> T)"
shows "ANR S"
proof -
- obtain V where clo: "closedin (subtopology euclidean (S \<union> T)) (S \<inter> T)"
- and ope: "openin (subtopology euclidean (S \<union> T)) V"
+ obtain V where clo: "closedin (top_of_set (S \<union> T)) (S \<inter> T)"
+ and ope: "openin (top_of_set (S \<union> T)) V"
and ret: "S \<inter> 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 \<subseteq> S \<inter> T" and req: "\<forall>x\<in>S \<inter> T. r x = x"
@@ -3808,9 +3808,9 @@
using Vsub by blast
have "continuous_on (S \<union> V \<inter> T) (\<lambda>x. if x \<in> S then x else r x)"
proof (rule continuous_on_cases_local)
- show "closedin (subtopology euclidean (S \<union> V \<inter> T)) S"
+ show "closedin (top_of_set (S \<union> V \<inter> T)) S"
using clS closedin_subset_trans inf.boundedE by blast
- show "closedin (subtopology euclidean (S \<union> V \<inter> T)) (V \<inter> T)"
+ show "closedin (top_of_set (S \<union> V \<inter> T)) (V \<inter> T)"
using clT Vsup by (auto simp: closedin_closed)
show "continuous_on (V \<inter> 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 \<noteq> {}" "S homeomorphic S'" "closedin (subtopology euclidean U) S'"
+ assumes "S retract_of T" "convex T" "T \<noteq> {}" "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 \<union> T)) S" "closedin (subtopology euclidean (S \<union> T)) T" "ENR(S \<union> T)" "ENR(S \<inter> T)"
+ assumes "closedin (top_of_set (S \<union> T)) S" "closedin (top_of_set (S \<union> T)) T" "ENR(S \<union> T)" "ENR(S \<inter> 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 \<Rightarrow> 'b::euclidean_space"
- assumes cloTS: "closedin (subtopology euclidean T) S"
+ assumes cloTS: "closedin (top_of_set T) S"
and anr: "(ANR S \<and> ANR T) \<or> ANR U"
and contf: "continuous_on T f"
and "f ` T \<subseteq> U"
@@ -4140,15 +4140,15 @@
using assms by (auto simp: homotopic_with_def)
define h' where "h' \<equiv> \<lambda>z. if snd z \<in> S then h z else (f \<circ> snd) z"
define B where "B \<equiv> {0::real} \<times> T \<union> {0..1} \<times> S"
- have clo0T: "closedin (subtopology euclidean ({0..1} \<times> T)) ({0::real} \<times> T)"
+ have clo0T: "closedin (top_of_set ({0..1} \<times> T)) ({0::real} \<times> T)"
by (simp add: closedin_subtopology_refl closedin_Times)
- moreover have cloT1S: "closedin (subtopology euclidean ({0..1} \<times> T)) ({0..1} \<times> S)"
+ moreover have cloT1S: "closedin (top_of_set ({0..1} \<times> T)) ({0..1} \<times> S)"
by (simp add: closedin_subtopology_refl closedin_Times assms)
- ultimately have clo0TB:"closedin (subtopology euclidean ({0..1} \<times> T)) B"
+ ultimately have clo0TB:"closedin (top_of_set ({0..1} \<times> T)) B"
by (auto simp: B_def)
- have cloBS: "closedin (subtopology euclidean B) ({0..1} \<times> S)"
+ have cloBS: "closedin (top_of_set B) ({0..1} \<times> 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} \<times> T)"
+ moreover have cloBT: "closedin (top_of_set B) ({0} \<times> T)"
using \<open>S \<subseteq> T\<close> closedin_subset_trans [OF clo0T]
by (metis B_def Un_upper1 clo0TB closedin_closed inf_le1)
moreover have "continuous_on ({0} \<times> T) (f \<circ> snd)"
@@ -4161,7 +4161,7 @@
done
have "image h' B \<subseteq> U"
using \<open>f ` T \<subseteq> U\<close> him by (auto simp: h'_def B_def)
- obtain V k where "B \<subseteq> V" and opeTV: "openin (subtopology euclidean ({0..1} \<times> T)) V"
+ obtain V k where "B \<subseteq> V" and opeTV: "openin (top_of_set ({0..1} \<times> T)) V"
and contk: "continuous_on V k" and kim: "k ` V \<subseteq> U"
and keq: "\<And>x. x \<in> B \<Longrightarrow> 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} \<times> T)) V"
+ have *: thesis if "openin (top_of_set ({0..1::real} \<times> 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' \<equiv> {x. \<exists>u::real. u \<in> {0..1} \<and> (u, x::'a) \<in> {0..1} \<times> 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. \<exists>u::real. (u, x::'a) \<in> {0..1} \<times> T - V}"
by (auto simp: S'_def)
- have cloTS': "closedin (subtopology euclidean T) S'"
- using S'_def \<open>closedin (subtopology euclidean T) S'\<close> by blast
+ have cloTS': "closedin (top_of_set T) S'"
+ using S'_def \<open>closedin (top_of_set T) S'\<close> by blast
have "S \<inter> S' = {}"
using S'_def B_def \<open>B \<subseteq> V\<close> by force
obtain a :: "'a \<Rightarrow> real" where conta: "continuous_on T a"
@@ -4269,7 +4269,7 @@
assume ?lhs
then obtain c where c: "homotopic_with (\<lambda>x. True) S T (\<lambda>x. c) f"
by (blast intro: homotopic_with_symD)
- have "closedin (subtopology euclidean UNIV) S"
+ have "closedin (top_of_set UNIV) S"
using \<open>closed S\<close> closed_closedin by fastforce
then obtain g where "continuous_on UNIV g" "range g \<subseteq> T"
"\<And>x. x \<in> S \<Longrightarrow> g x = f x"
@@ -4412,7 +4412,7 @@
"g ` (S \<union> connected_component_set (- S) a) \<subseteq> sphere 0 1"
"\<And>x. x \<in> S \<Longrightarrow> 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: \<open>compact S\<close> closed_subset compact_imp_closed)
show "continuous_on ?T (\<lambda>x. (x - b) /\<^sub>R norm (x - b))"
by (simp add: \<open>b \<notin> S\<close> notcc continuous_on_Borsuk_map)
@@ -4467,8 +4467,8 @@
subsubsection\<open>More extension theorems\<close>
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 \<subseteq> U" and null: "U = {} \<Longrightarrow> S = {}"
obtains g where "continuous_on S g" "g ` S \<subseteq> U" "\<And>x. x \<in> T \<Longrightarrow> g x = f x"
proof (cases "U = {}")
@@ -4503,8 +4503,8 @@
and C: "C \<in> components S" and contf: "continuous_on C f" and fim: "f ` C \<subseteq> U"
obtains g where "continuous_on S g" "g ` S \<subseteq> U" "\<And>x. x \<in> C \<Longrightarrow> 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 \<subseteq> T" and contg: "continuous_on T g" and gim: "g ` T \<subseteq> U"
and gf: "\<And>x. x \<in> C \<Longrightarrow> g x = f x"
using S
@@ -4514,7 +4514,7 @@
by (metis C \<open>locally connected S\<close> openin_components_locally_connected closedin_component contf fim order_refl that)
next
assume "compact S"
- then obtain W g where "C \<subseteq> W" and opeW: "openin (subtopology euclidean S) W"
+ then obtain W g where "C \<subseteq> W" and opeW: "openin (top_of_set S) W"
and contg: "continuous_on W g"
and gim: "g ` W \<subseteq> U" and gf: "\<And>x. x \<in> C \<Longrightarrow> g x = f x"
using ANR_imp_absolute_neighbourhood_extensor [of U C f S] C \<open>ANR U\<close> closedin_component contf fim by blast
@@ -4522,11 +4522,11 @@
by (auto simp: openin_open)
moreover have "locally compact S"
by (simp add: \<open>compact S\<close> closed_imp_locally_compact compact_imp_closed)
- ultimately obtain K where opeK: "openin (subtopology euclidean S) K" and "compact K" "C \<subseteq> K" "K \<subseteq> V"
+ ultimately obtain K where opeK: "openin (top_of_set S) K" and "compact K" "C \<subseteq> K" "K \<subseteq> V"
by (metis C Int_subset_iff \<open>C \<subseteq> W\<close> \<open>compact S\<close> compact_components Sura_Bura_clopen_subset)
show ?thesis
proof
- show "closedin (subtopology euclidean S) K"
+ show "closedin (top_of_set S) K"
by (meson \<open>compact K\<close> \<open>compact S\<close> closedin_compact_eq opeK openin_imp_subset)
show "continuous_on K g"
by (metis Int_subset_iff V \<open>K \<subseteq> V\<close> 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 \<noteq> {}" "(\<lambda>x. (x,a)) ` S \<subseteq> U"
- and ope: "openin (subtopology euclidean (S \<times> T)) U"
- obtains V where "openin (subtopology euclidean T) V" "a \<in> V" "S \<times> V \<subseteq> U"
+ and ope: "openin (top_of_set (S \<times> T)) U"
+ obtains V where "openin (top_of_set T) V" "a \<in> V" "S \<times> V \<subseteq> U"
proof -
let ?W = "{y. \<exists>x. x \<in> S \<and> (x, y) \<in> (S \<times> T - U)}"
- have "U \<subseteq> S \<times> T" "closedin (subtopology euclidean (S \<times> T)) (S \<times> T - U)"
+ have "U \<subseteq> S \<times> T" "closedin (top_of_set (S \<times> T)) (S \<times> 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 \<open>compact S\<close> closedin_compact_projection by blast
moreover have "a \<in> T - ?W"
using \<open>U \<subseteq> S \<times> T\<close> 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 \<noteq> {}" "T \<subseteq> T'" "S \<times> T \<subseteq> U"
- and ope: "openin (subtopology euclidean (S \<times> T')) U"
- obtains V where "openin (subtopology euclidean T') V" "T \<subseteq> V" "S \<times> V \<subseteq> U"
+ and ope: "openin (top_of_set (S \<times> T')) U"
+ obtains V where "openin (top_of_set T') V" "T \<subseteq> V" "S \<times> V \<subseteq> U"
proof -
- have "\<And>x. x \<in> T \<Longrightarrow> \<exists>V. openin (subtopology euclidean T') V \<and> x \<in> V \<and> S \<times> V \<subseteq> U"
+ have "\<And>x. x \<in> T \<Longrightarrow> \<exists>V. openin (top_of_set T') V \<and> x \<in> V \<and> S \<times> V \<subseteq> U"
using assms by (auto intro: tube_lemma [OF \<open>compact S\<close>])
- then obtain F where F: "\<And>x. x \<in> T \<Longrightarrow> openin (subtopology euclidean T') (F x) \<and> x \<in> F x \<and> S \<times> F x \<subseteq> U"
+ then obtain F where F: "\<And>x. x \<in> T \<Longrightarrow> openin (top_of_set T') (F x) \<and> x \<in> F x \<and> S \<times> F x \<subseteq> U"
by metis
show ?thesis
proof
- show "openin (subtopology euclidean T') (\<Union>(F ` T))"
+ show "openin (top_of_set T') (\<Union>(F ` T))"
using F by blast
show "T \<subseteq> \<Union>(F ` T)"
using F by blast
@@ -4586,9 +4586,9 @@
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes contf: "continuous_on S f" and fim: "f ` S \<subseteq> U"
and contg: "continuous_on S g" and gim: "g ` S \<subseteq> U"
- and clo: "closedin (subtopology euclidean S) T"
+ and clo: "closedin (top_of_set S) T"
and "ANR U" and hom: "homotopic_with (\<lambda>x. True) T U f g"
- obtains V where "T \<subseteq> V" "openin (subtopology euclidean S) V"
+ obtains V where "T \<subseteq> V" "openin (top_of_set S) V"
"homotopic_with (\<lambda>x. True) V U f g"
proof -
have "T \<subseteq> S"
@@ -4604,11 +4604,11 @@
have "continuous_on(?S0 \<union> (?S1 \<union> {0..1} \<times> T)) h'"
unfolding h'_def
proof (intro continuous_on_cases_local)
- show "closedin (subtopology euclidean (?S0 \<union> (?S1 \<union> {0..1} \<times> T))) ?S0"
- "closedin (subtopology euclidean (?S1 \<union> {0..1} \<times> T)) ?S1"
+ show "closedin (top_of_set (?S0 \<union> (?S1 \<union> {0..1} \<times> T))) ?S0"
+ "closedin (top_of_set (?S1 \<union> {0..1} \<times> T)) ?S1"
using \<open>T \<subseteq> S\<close> by (force intro: closedin_Times closedin_subset_trans [of "{0..1} \<times> S"])+
- show "closedin (subtopology euclidean (?S0 \<union> (?S1 \<union> {0..1} \<times> T))) (?S1 \<union> {0..1} \<times> T)"
- "closedin (subtopology euclidean (?S1 \<union> {0..1} \<times> T)) ({0..1} \<times> T)"
+ show "closedin (top_of_set (?S0 \<union> (?S1 \<union> {0..1} \<times> T))) (?S1 \<union> {0..1} \<times> T)"
+ "closedin (top_of_set (?S1 \<union> {0..1} \<times> T)) ({0..1} \<times> T)"
using \<open>T \<subseteq> S\<close> by (force intro: clo closedin_Times closedin_subset_trans [of "{0..1} \<times> S"])+
show "continuous_on (?S0) (\<lambda>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} \<times> S \<union> {0..1} \<times> T) \<subseteq> U"
using fim gim him \<open>T \<subseteq> S\<close> unfolding h'_def by force
- moreover have "closedin (subtopology euclidean ({0..1::real} \<times> S)) ({0,1} \<times> S \<union> {0..1::real} \<times> T)"
+ moreover have "closedin (top_of_set ({0..1::real} \<times> S)) ({0,1} \<times> S \<union> {0..1::real} \<times> T)"
by (intro closedin_Times closedin_Un clo) (simp_all add: closed_subset)
ultimately
obtain W k where W: "({0,1} \<times> S) \<union> ({0..1} \<times> T) \<subseteq> W"
- and opeW: "openin (subtopology euclidean ({0..1} \<times> S)) W"
+ and opeW: "openin (top_of_set ({0..1} \<times> S)) W"
and contk: "continuous_on W k"
and kim: "k ` W \<subseteq> U"
and kh': "\<And>x. x \<in> ({0,1} \<times> S) \<union> ({0..1} \<times> T) \<Longrightarrow> k x = h' x"
by (metis ANR_imp_absolute_neighbourhood_extensor [OF \<open>ANR U\<close>, of "({0,1} \<times> S) \<union> ({0..1} \<times> T)" h' "{0..1} \<times> S"])
- obtain T' where opeT': "openin (subtopology euclidean S) T'"
+ obtain T' where opeT': "openin (top_of_set S) T'"
and "T \<subseteq> T'" and TW: "{0..1} \<times> T' \<subseteq> W"
using tube_lemma_gen [of "{0..1::real}" T S W] W \<open>T \<subseteq> S\<close> opeW by auto
moreover have "homotopic_with (\<lambda>x. True) T' U f g"
@@ -4649,8 +4649,8 @@
text\<open> Homotopy on a union of closed-open sets.\<close>
proposition%unimportant homotopic_on_clopen_Union:
fixes \<F> :: "'a::euclidean_space set set"
- assumes "\<And>S. S \<in> \<F> \<Longrightarrow> closedin (subtopology euclidean (\<Union>\<F>)) S"
- and "\<And>S. S \<in> \<F> \<Longrightarrow> openin (subtopology euclidean (\<Union>\<F>)) S"
+ assumes "\<And>S. S \<in> \<F> \<Longrightarrow> closedin (top_of_set (\<Union>\<F>)) S"
+ and "\<And>S. S \<in> \<F> \<Longrightarrow> openin (top_of_set (\<Union>\<F>)) S"
and "\<And>S. S \<in> \<F> \<Longrightarrow> homotopic_with (\<lambda>x. True) S T f g"
shows "homotopic_with (\<lambda>x. True) (\<Union>\<F>) T f g"
proof -
@@ -4665,8 +4665,8 @@
case False
then obtain V :: "nat \<Rightarrow> 'a set" where V: "range V = \<V>"
using range_from_nat_into \<open>countable \<V>\<close> by metis
- with \<open>\<V> \<subseteq> \<F>\<close> have clo: "\<And>n. closedin (subtopology euclidean (\<Union>\<F>)) (V n)"
- and ope: "\<And>n. openin (subtopology euclidean (\<Union>\<F>)) (V n)"
+ with \<open>\<V> \<subseteq> \<F>\<close> have clo: "\<And>n. closedin (top_of_set (\<Union>\<F>)) (V n)"
+ and ope: "\<And>n. openin (top_of_set (\<Union>\<F>)) (V n)"
and hom: "\<And>n. homotopic_with (\<lambda>x. True) (V n) T f g"
using assms by auto
then obtain h where conth: "\<And>n. continuous_on ({0..1::real} \<times> V n) (h n)"
@@ -4680,20 +4680,20 @@
and eq: "\<And>x i. \<lbrakk>x \<in> {0..1} \<times> \<Union>(V ` UNIV) \<inter>
{0..1} \<times> (V i - (\<Union>m<i. V m))\<rbrakk> \<Longrightarrow> \<zeta> x = h i x"
proof (rule pasting_lemma_exists)
- show "{0..1} \<times> \<Union>(V ` UNIV) \<subseteq> (\<Union>i. {0..1::real} \<times> (V i - (\<Union>m<i. V m)))"
+ let ?X = "top_of_set ({0..1::real} \<times> \<Union>(range V))"
+ show "topspace ?X \<subseteq> (\<Union>i. {0..1::real} \<times> (V i - (\<Union>m<i. V m)))"
by (force simp: Ball_def dest: wop)
- show "openin (subtopology euclidean ({0..1} \<times> \<Union>(V ` UNIV)))
+ show "openin (top_of_set ({0..1} \<times> \<Union>(V ` UNIV)))
({0..1::real} \<times> (V i - (\<Union>m<i. V m)))" for i
proof (intro openin_Times openin_subtopology_self openin_diff)
- show "openin (subtopology euclidean (\<Union>(V ` UNIV))) (V i)"
+ show "openin (top_of_set (\<Union>(V ` UNIV))) (V i)"
using ope V eqU by auto
- show "closedin (subtopology euclidean (\<Union>(V ` UNIV))) (\<Union>m<i. V m)"
+ show "closedin (top_of_set (\<Union>(V ` UNIV))) (\<Union>m<i. V m)"
using V clo eqU by (force intro: closedin_Union)
qed
- show "continuous_on ({0..1} \<times> (V i - (\<Union>m<i. V m))) (h i)" for i
- by (rule continuous_on_subset [OF conth]) auto
- show "\<And>i j x. x \<in> {0..1} \<times> \<Union>(V ` UNIV) \<inter>
- {0..1} \<times> (V i - (\<Union>m<i. V m)) \<inter> {0..1} \<times> (V j - (\<Union>m<j. V m))
+ show "continuous_map (subtopology ?X ({0..1} \<times> (V i - \<Union> (V ` {..<i})))) euclidean (h i)" for i
+ by (auto simp add: subtopology_subtopology intro!: continuous_on_subset [OF conth])
+ show "\<And>i j x. x \<in> topspace ?X \<inter> {0..1} \<times> (V i - (\<Union>m<i. V m)) \<inter> {0..1} \<times> (V j - (\<Union>m<j. V m))
\<Longrightarrow> 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 "\<exists>U. C \<subseteq> U \<and> closedin (subtopology euclidean S) U \<and>
- openin (subtopology euclidean S) U \<and>
+ have "\<exists>U. C \<subseteq> U \<and> closedin (top_of_set S) U \<and>
+ openin (top_of_set S) U \<and>
homotopic_with (\<lambda>x. True) U T f g" if C: "C \<in> components S" for C
proof -
have "C \<subseteq> 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: \<open>locally connected S\<close> openin_components_locally_connected that)
show "homotopic_with (\<lambda>x. True) C T f g"
by (simp add: R that)
@@ -4761,7 +4761,7 @@
assume "compact S"
have hom: "homotopic_with (\<lambda>x. True) C T f g"
using R that by blast
- obtain U where "C \<subseteq> U" and opeU: "openin (subtopology euclidean S) U"
+ obtain U where "C \<subseteq> U" and opeU: "openin (top_of_set S) U"
and hom: "homotopic_with (\<lambda>x. True) U T f g"
using homotopic_neighbourhood_extension [OF contf fim contg gim _ \<open>ANR T\<close> hom]
\<open>C \<in> components S\<close> closedin_component by blast
@@ -4769,11 +4769,11 @@
by (auto simp: openin_open)
moreover have "locally compact S"
by (simp add: \<open>compact S\<close> closed_imp_locally_compact compact_imp_closed)
- ultimately obtain K where opeK: "openin (subtopology euclidean S) K" and "compact K" "C \<subseteq> K" "K \<subseteq> V"
+ ultimately obtain K where opeK: "openin (top_of_set S) K" and "compact K" "C \<subseteq> K" "K \<subseteq> V"
by (metis C Int_subset_iff Sura_Bura_clopen_subset \<open>C \<subseteq> U\<close> \<open>compact S\<close> compact_components)
show ?thesis
proof (intro exI conjI)
- show "closedin (subtopology euclidean S) K"
+ show "closedin (top_of_set S) K"
by (meson \<open>compact K\<close> \<open>compact S\<close> closedin_compact_eq opeK openin_imp_subset)
show "homotopic_with (\<lambda>x. True) K T f g"
using V \<open>K \<subseteq> V\<close> hom homotopic_with_subset_left opeK openin_imp_subset by fastforce
@@ -4781,8 +4781,8 @@
qed
qed
then obtain \<phi> where \<phi>: "\<And>C. C \<in> components S \<Longrightarrow> C \<subseteq> \<phi> C"
- and clo\<phi>: "\<And>C. C \<in> components S \<Longrightarrow> closedin (subtopology euclidean S) (\<phi> C)"
- and ope\<phi>: "\<And>C. C \<in> components S \<Longrightarrow> openin (subtopology euclidean S) (\<phi> C)"
+ and clo\<phi>: "\<And>C. C \<in> components S \<Longrightarrow> closedin (top_of_set S) (\<phi> C)"
+ and ope\<phi>: "\<And>C. C \<in> components S \<Longrightarrow> openin (top_of_set S) (\<phi> C)"
and hom\<phi>: "\<And>C. C \<in> components S \<Longrightarrow> homotopic_with (\<lambda>x. True) (\<phi> C) T f g"
by metis
have Seq: "S = \<Union> (\<phi> ` components S)"
--- a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Tue Mar 19 16:14:59 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) (\<Inter>n. {w \<in> S. (deriv ^^ n) f w = 0})"
+ have 1: "openin (top_of_set S) (\<Inter>n. {w \<in> S. (deriv ^^ n) f w = 0})"
apply (rule open_subset, force)
using \<open>open S\<close>
apply (simp add: open_contains_ball Ball_def)
apply (erule all_forward)
using "*" by auto blast+
- have 2: "closedin (subtopology euclidean S) (\<Inter>n. {w \<in> S. (deriv ^^ n) f w = 0})"
+ have 2: "closedin (top_of_set S) (\<Inter>n. {w \<in> 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 \<subseteq> S" using openE [OF \<open>open S\<close> \<open>w \<in> S\<close>] .
--- a/src/HOL/Analysis/Change_Of_Vars.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Change_Of_Vars.thy Tue Mar 19 16:14:59 2019 +0000
@@ -1567,24 +1567,24 @@
show "countable {A. \<forall>i j. A $ i $ j \<in> \<rat>}"
using countable_vector [OF countable_vector, of "\<lambda>i j. \<rat>"] by (simp add: countable_rat)
qed blast
- have *: "\<lbrakk>U \<noteq> {} \<Longrightarrow> closedin (subtopology euclidean S) (S \<inter> \<Inter> U)\<rbrakk>
- \<Longrightarrow> closedin (subtopology euclidean S) (S \<inter> \<Inter> U)" for U
+ have *: "\<lbrakk>U \<noteq> {} \<Longrightarrow> closedin (top_of_set S) (S \<inter> \<Inter> U)\<rbrakk>
+ \<Longrightarrow> closedin (top_of_set S) (S \<inter> \<Inter> U)" for U
by fastforce
have eq: "{x::(real,'m)vec. P x \<and> (Q x \<longrightarrow> R x)} = {x. P x \<and> \<not> Q x} \<union> {x. P x \<and> R x}" for P Q R
by auto
have sets: "S \<inter> (\<Inter>y\<in>S. {x \<in> S. norm (y - x) < d \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> e * norm (y - x)})
\<in> sets lebesgue" for e A d
proof -
- have clo: "closedin (subtopology euclidean S)
+ have clo: "closedin (top_of_set S)
{x \<in> S. norm (y - x) < d \<longrightarrow> norm (f y - f x - A *v (y - x)) \<le> e * norm (y - x)}"
for y
proof -
have cont1: "continuous_on S (\<lambda>x. norm (y - x))"
and cont2: "continuous_on S (\<lambda>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 \<in> S. d \<le> norm (y - x)}"
+ have clo1: "closedin (top_of_set S) {x \<in> S. d \<le> 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 \<in> S. norm (f y - f x - (A *v y - A *v x)) \<le> e * norm (y - x)}"
using continuous_closedin_preimage [OF cont2, of "{0..}"] by (simp add: vimage_def Int_def)
show ?thesis
--- a/src/HOL/Analysis/Connected.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Connected.thy Tue Mar 19 16:14:59 2019 +0000
@@ -14,8 +14,8 @@
lemma connected_local:
"connected S \<longleftrightarrow>
\<not> (\<exists>e1 e2.
- openin (subtopology euclidean S) e1 \<and>
- openin (subtopology euclidean S) e2 \<and>
+ openin (top_of_set S) e1 \<and>
+ openin (top_of_set S) e2 \<and>
S \<subseteq> e1 \<union> e2 \<and>
e1 \<inter> e2 = {} \<and>
e1 \<noteq> {} \<and>
@@ -39,8 +39,8 @@
qed
lemma connected_clopen: "connected S \<longleftrightarrow>
- (\<forall>T. openin (subtopology euclidean S) T \<and>
- closedin (subtopology euclidean S) T \<longrightarrow> T = {} \<or> T = S)" (is "?lhs \<longleftrightarrow> ?rhs")
+ (\<forall>T. openin (top_of_set S) T \<and>
+ closedin (top_of_set S) T \<longrightarrow> T = {} \<or> T = S)" (is "?lhs \<longleftrightarrow> ?rhs")
proof -
have "\<not> connected S \<longleftrightarrow>
(\<exists>e1 e2. open e1 \<and> open (- e2) \<and> S \<subseteq> e1 \<union> (- e2) \<and> e1 \<inter> (- e2) \<inter> S = {} \<and> e1 \<inter> S \<noteq> {} \<and> (- e2) \<inter> S \<noteq> {})"
@@ -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 \<in> components s \<Longrightarrow> closedin (subtopology euclidean s) C"
+ "C \<in> components s \<Longrightarrow> closedin (top_of_set s) C"
using closedin_connected_component componentsE by blast
@@ -433,7 +433,7 @@
lemma continuous_levelset_openin_cases:
fixes f :: "_ \<Rightarrow> 'b::t1_space"
shows "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>
- openin (subtopology euclidean s) {x \<in> s. f x = a}
+ openin (top_of_set s) {x \<in> s. f x = a}
\<Longrightarrow> (\<forall>x \<in> s. f x \<noteq> a) \<or> (\<forall>x \<in> s. f x = a)"
unfolding connected_clopen
using continuous_closedin_preimage_constant by auto
@@ -441,7 +441,7 @@
lemma continuous_levelset_openin:
fixes f :: "_ \<Rightarrow> 'b::t1_space"
shows "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>
- openin (subtopology euclidean s) {x \<in> s. f x = a} \<Longrightarrow>
+ openin (top_of_set s) {x \<in> s. f x = a} \<Longrightarrow>
(\<exists>x \<in> s. f x = a) \<Longrightarrow> (\<forall>x \<in> 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: "\<And>U. U \<subseteq> T
- \<Longrightarrow> openin (subtopology euclidean S) (S \<inter> f -` U) \<longleftrightarrow>
- openin (subtopology euclidean T) U"
+ \<Longrightarrow> openin (top_of_set S) (S \<inter> f -` U) \<longleftrightarrow>
+ openin (top_of_set T) U"
and connT: "\<And>y. y \<in> T \<Longrightarrow> connected (S \<inter> f -` {y})"
shows "connected S"
proof (rule connectedI)
@@ -492,9 +492,9 @@
qed
ultimately have UU: "(S \<inter> f -` f ` (S \<inter> U)) = S \<inter> U" and VV: "(S \<inter> f -` f ` (S \<inter> V)) = S \<inter> V"
by auto
- have opeU: "openin (subtopology euclidean T) (f ` (S \<inter> U))"
+ have opeU: "openin (top_of_set T) (f ` (S \<inter> U))"
by (metis UU \<open>open U\<close> fim image_Int_subset le_inf_iff opT openin_open_Int)
- have opeV: "openin (subtopology euclidean T) (f ` (S \<inter> V))"
+ have opeV: "openin (top_of_set T) (f ` (S \<inter> V))"
by (metis opT fim VV \<open>open V\<close> openin_open_Int image_Int_subset inf.bounded_iff)
have "T \<subseteq> f ` (S \<inter> U) \<union> f ` (S \<inter> V)"
using \<open>S \<subseteq> U \<union> V\<close> 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: "\<And>C. openin (subtopology euclidean S) C \<Longrightarrow> openin (subtopology euclidean T) (f ` C)"
+ and ST: "\<And>C. openin (top_of_set S) C \<Longrightarrow> openin (top_of_set T) (f ` C)"
and connT: "\<And>y. y \<in> T \<Longrightarrow> connected (S \<inter> f -` {y})"
and "connected C" "C \<subseteq> T"
shows "connected (S \<inter> f -` C)"
@@ -525,12 +525,12 @@
ultimately show ?thesis
by metis
qed
- have "\<And>U. openin (subtopology euclidean (S \<inter> f -` C)) U
- \<Longrightarrow> openin (subtopology euclidean C) (f ` U)"
+ have "\<And>U. openin (top_of_set (S \<inter> f -` C)) U
+ \<Longrightarrow> openin (top_of_set C) (f ` U)"
using open_map_restrict [OF _ ST \<open>C \<subseteq> T\<close>] by metis
then show "\<And>D. D \<subseteq> C
- \<Longrightarrow> openin (subtopology euclidean (S \<inter> f -` C)) (S \<inter> f -` C \<inter> f -` D) =
- openin (subtopology euclidean C) D"
+ \<Longrightarrow> openin (top_of_set (S \<inter> f -` C)) (S \<inter> f -` C \<inter> f -` D) =
+ openin (top_of_set C) D"
using open_map_imp_quotient_map [of "(S \<inter> 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: "\<And>C. closedin (subtopology euclidean S) C \<Longrightarrow> closedin (subtopology euclidean T) (f ` C)"
+ and ST: "\<And>C. closedin (top_of_set S) C \<Longrightarrow> closedin (top_of_set T) (f ` C)"
and connT: "\<And>y. y \<in> T \<Longrightarrow> connected (S \<inter> f -` {y})"
and "connected C" "C \<subseteq> T"
shows "connected (S \<inter> f -` C)"
@@ -558,12 +558,12 @@
ultimately show ?thesis
by metis
qed
- have "\<And>U. closedin (subtopology euclidean (S \<inter> f -` C)) U
- \<Longrightarrow> closedin (subtopology euclidean C) (f ` U)"
+ have "\<And>U. closedin (top_of_set (S \<inter> f -` C)) U
+ \<Longrightarrow> closedin (top_of_set C) (f ` U)"
using closed_map_restrict [OF _ ST \<open>C \<subseteq> T\<close>] by metis
then show "\<And>D. D \<subseteq> C
- \<Longrightarrow> openin (subtopology euclidean (S \<inter> f -` C)) (S \<inter> f -` C \<inter> f -` D) =
- openin (subtopology euclidean C) D"
+ \<Longrightarrow> openin (top_of_set (S \<inter> f -` C)) (S \<inter> f -` C \<inter> f -` D) =
+ openin (top_of_set C) D"
using closed_map_imp_quotient_map [of "(S \<inter> 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 \<subseteq> 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 \<union> T)"
proof -
have *: "\<lbrakk>\<And>x y. P x y \<longleftrightarrow> P y x; \<And>x y. P x y \<Longrightarrow> S \<subseteq> x \<or> S \<subseteq> y;
@@ -587,11 +587,11 @@
unfolding connected_closedin_eq
proof (rule *)
fix H1 H2
- assume H: "closedin (subtopology euclidean (S \<union> T)) H1 \<and>
- closedin (subtopology euclidean (S \<union> T)) H2 \<and>
+ assume H: "closedin (top_of_set (S \<union> T)) H1 \<and>
+ closedin (top_of_set (S \<union> T)) H2 \<and>
H1 \<union> H2 = S \<union> T \<and> H1 \<inter> H2 = {} \<and> H1 \<noteq> {} \<and> H2 \<noteq> {}"
- then have clo: "closedin (subtopology euclidean S) (S \<inter> H1)"
- "closedin (subtopology euclidean S) (S \<inter> H2)"
+ then have clo: "closedin (top_of_set S) (S \<inter> H1)"
+ "closedin (top_of_set S) (S \<inter> H2)"
by (metis Un_upper1 closedin_closed_subset inf_commute)+
have Seq: "S \<inter> (H1 \<union> H2) = S"
by (simp add: H)
@@ -606,8 +606,8 @@
using H \<open>connected S\<close> unfolding connected_closedin by blast
next
fix H1 H2
- assume H: "closedin (subtopology euclidean (S \<union> T)) H1 \<and>
- closedin (subtopology euclidean (S \<union> T)) H2 \<and>
+ assume H: "closedin (top_of_set (S \<union> T)) H1 \<and>
+ closedin (top_of_set (S \<union> T)) H2 \<and>
H1 \<union> H2 = S \<union> T \<and> H1 \<inter> H2 = {} \<and> H1 \<noteq> {} \<and> H2 \<noteq> {}"
and "S \<subseteq> H1"
then have H2T: "H2 \<subseteq> T"
@@ -616,17 +616,17 @@
using Diff_iff opeT openin_imp_subset by auto
with \<open>S \<subseteq> U\<close> have Ueq: "U = (U - S) \<union> (S \<union> T)"
by auto
- have "openin (subtopology euclidean ((U - S) \<union> (S \<union> T))) H2"
+ have "openin (top_of_set ((U - S) \<union> (S \<union> T))) H2"
proof (rule openin_subtopology_Un)
- show "openin (subtopology euclidean (S \<union> T)) H2"
+ show "openin (top_of_set (S \<union> T)) H2"
using \<open>H2 \<subseteq> T\<close> 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) \<union> (S \<union> T))) H2"
+ moreover have "closedin (top_of_set ((U - S) \<union> (S \<union> 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 \<open>connected S\<close> 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 \<union> H4 = U - C" and "H3 \<inter> H4 = {}" and "H3 \<noteq> {}" and "H4 \<noteq> {}"
and * [rule_format]:
- "\<forall>H1 H2. \<not> closedin (subtopology euclidean S) H1 \<or>
- \<not> closedin (subtopology euclidean S) H2 \<or>
+ "\<forall>H1 H2. \<not> closedin (top_of_set S) H1 \<or>
+ \<not> closedin (top_of_set S) H2 \<or>
H1 \<union> H2 \<noteq> S \<or> H1 \<inter> H2 \<noteq> {} \<or> \<not> H1 \<noteq> {} \<or> \<not> H2 \<noteq> {}"
- then have "H3 \<subseteq> U-C" and ope3: "openin (subtopology euclidean (U - C)) (U - C - H3)"
- and "H4 \<subseteq> U-C" and ope4: "openin (subtopology euclidean (U - C)) (U - C - H4)"
+ then have "H3 \<subseteq> U-C" and ope3: "openin (top_of_set (U - C)) (U - C - H3)"
+ and "H4 \<subseteq> U-C" and ope4: "openin (top_of_set (U - C)) (U - C - H4)"
by (auto simp: closedin_def)
have "C \<noteq> {}" "C \<subseteq> U-S" "connected C"
using C in_components_nonempty in_components_subset in_components_maximal by blast+
have cCH3: "connected (C \<union> H3)"
proof (rule connected_Un_clopen_in_complement [OF \<open>connected C\<close> \<open>connected U\<close> _ _ clo3])
- show "openin (subtopology euclidean (U - C)) H3"
+ show "openin (top_of_set (U - C)) H3"
apply (simp add: openin_closedin_eq \<open>H3 \<subseteq> U - C\<close>)
apply (simp add: closedin_subtopology)
by (metis Diff_cancel Diff_triv Un_Diff clo4 \<open>H3 \<inter> H4 = {}\<close> \<open>H3 \<union> H4 = U - C\<close> closedin_closed inf_commute sup_bot.left_neutral)
qed (use clo3 \<open>C \<subseteq> U - S\<close> in auto)
have cCH4: "connected (C \<union> H4)"
proof (rule connected_Un_clopen_in_complement [OF \<open>connected C\<close> \<open>connected U\<close> _ _ clo4])
- show "openin (subtopology euclidean (U - C)) H4"
+ show "openin (top_of_set (U - C)) H4"
apply (simp add: openin_closedin_eq \<open>H4 \<subseteq> U - C\<close>)
apply (simp add: closedin_subtopology)
by (metis Diff_cancel Int_commute Un_Diff Un_Diff_Int \<open>H3 \<inter> H4 = {}\<close> \<open>H3 \<union> H4 = U - C\<close> clo3 closedin_closed)
qed (use clo4 \<open>C \<subseteq> U - S\<close> in auto)
- have "closedin (subtopology euclidean S) (S \<inter> H3)" "closedin (subtopology euclidean S) (S \<inter> H4)"
+ have "closedin (top_of_set S) (S \<inter> H3)" "closedin (top_of_set S) (S \<inter> H4)"
using clo3 clo4 \<open>S \<subseteq> U\<close> \<open>C \<subseteq> U - S\<close> by (auto simp: closedin_closed)
moreover have "S \<inter> H3 \<noteq> {}"
using components_maximal [OF C cCH3] \<open>C \<noteq> {}\<close> \<open>C \<subseteq> U - S\<close> \<open>H3 \<noteq> {}\<close> \<open>H3 \<subseteq> U - C\<close> 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 \<inter> u = {}" and tus: "t \<union> u = S"
have conif: "continuous_on S (\<lambda>x. if x \<in> t then 0 else 1)"
apply (subst tus [symmetric])
--- a/src/HOL/Analysis/Continuous_Extension.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Continuous_Extension.thy Tue Mar 19 16:14:59 2019 +0000
@@ -33,7 +33,7 @@
by (simp add: supp_sum_def sum_nonneg)
have sd_pos: "0 < setdist {x} (S - V)" if "V \<in> \<C>" "x \<in> S" "x \<in> 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 \<open>V \<in> \<C>\<close>)
with that False setdist_pos_le [of "{x}" "S - V"]
show ?thesis
@@ -67,7 +67,7 @@
fix x assume "x \<in> S"
then obtain X where "open X" and x: "x \<in> S \<inter> X" and finX: "finite {U \<in> \<C>. U \<inter> X \<noteq> {}}"
using assms by blast
- then have OSX: "openin (subtopology euclidean S) (S \<inter> X)" by blast
+ then have OSX: "openin (top_of_set S) (S \<inter> X)" by blast
have sumeq: "\<And>x. x \<in> S \<inter> X \<Longrightarrow>
(\<Sum>V | V \<in> \<C> \<and> V \<inter> X \<noteq> {}. setdist {x} (S - V))
= supp_sum (\<lambda>V. setdist {x} (S - V)) \<C>"
@@ -114,8 +114,8 @@
text \<open>For Euclidean spaces the proof is easy using distances.\<close>
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 \<inter> T = {}" "S \<noteq> {}" "T \<noteq> {}" "a \<noteq> b"
obtains f :: "'a::euclidean_space \<Rightarrow> '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 \<inter> T = {}" "a \<noteq> b"
obtains f :: "'a::euclidean_space \<Rightarrow> '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 \<inter> T = {}"
obtains f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
where "continuous_on U f"
@@ -301,7 +301,7 @@
theorem Dugundji:
fixes f :: "'a::{metric_space,second_countable_topology} \<Rightarrow> 'b::real_inner"
assumes "convex C" "C \<noteq> {}"
- and cloin: "closedin (subtopology euclidean U) S"
+ and cloin: "closedin (top_of_set U) S"
and contf: "continuous_on S f" and "f ` S \<subseteq> C"
obtains g where "continuous_on U g" "g ` U \<subseteq> C"
"\<And>x. x \<in> S \<Longrightarrow> g x = f x"
@@ -427,7 +427,7 @@
obtain N where N: "open N" "a \<in> N"
and finN: "finite {U \<in> \<C>. \<exists>a\<in>N. H U a \<noteq> 0}"
using Hfin False \<open>a \<in> U\<close> 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 \<in> \<C>" for T
using Hcont [OF \<open>T \<in> \<C>\<close>] False \<open>a \<in> U\<close> \<open>T \<in> \<C>\<close>
@@ -444,7 +444,7 @@
(\<lambda>x. \<Sum>T\<in>{U \<in> \<C>. \<exists>x\<in>N. H U x \<noteq> 0}. H T x *\<^sub>R f (\<A> T))"
by (force intro: continuous_intros HcontU)+
next
- show "openin (subtopology euclidean U) ((U - S) \<inter> N)"
+ show "openin (top_of_set U) ((U - S) \<inter> N)"
using N oUS openin_trans by blast
next
show "a \<in> (U - S) \<inter> N" using False \<open>a \<in> U\<close> N by blast
@@ -475,7 +475,7 @@
corollary Tietze:
fixes f :: "'a::{metric_space,second_countable_topology} \<Rightarrow> 'b::real_inner"
assumes "continuous_on S f"
- and "closedin (subtopology euclidean U) S"
+ and "closedin (top_of_set U) S"
and "0 \<le> B"
and "\<And>x. x \<in> S \<Longrightarrow> norm(f x) \<le> B"
obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
@@ -485,7 +485,7 @@
corollary%unimportant Tietze_closed_interval:
fixes f :: "'a::{metric_space,second_countable_topology} \<Rightarrow> '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 \<noteq> {}"
and "\<And>x. x \<in> S \<Longrightarrow> f x \<in> cbox a b"
obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
@@ -496,7 +496,7 @@
corollary%unimportant Tietze_closed_interval_1:
fixes f :: "'a::{metric_space,second_countable_topology} \<Rightarrow> real"
assumes "continuous_on S f"
- and "closedin (subtopology euclidean U) S"
+ and "closedin (top_of_set U) S"
and "a \<le> b"
and "\<And>x. x \<in> S \<Longrightarrow> f x \<in> cbox a b"
obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
@@ -507,7 +507,7 @@
corollary%unimportant Tietze_open_interval:
fixes f :: "'a::{metric_space,second_countable_topology} \<Rightarrow> '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 \<noteq> {}"
and "\<And>x. x \<in> S \<Longrightarrow> f x \<in> box a b"
obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
@@ -518,7 +518,7 @@
corollary%unimportant Tietze_open_interval_1:
fixes f :: "'a::{metric_space,second_countable_topology} \<Rightarrow> real"
assumes "continuous_on S f"
- and "closedin (subtopology euclidean U) S"
+ and "closedin (top_of_set U) S"
and "a < b"
and no: "\<And>x. x \<in> S \<Longrightarrow> f x \<in> box a b"
obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
@@ -529,7 +529,7 @@
corollary%unimportant Tietze_unbounded:
fixes f :: "'a::{metric_space,second_countable_topology} \<Rightarrow> '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" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
apply (rule Dugundji [of UNIV U S f])
using assms by auto
--- a/src/HOL/Analysis/Convex_Euclidean_Space.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy Tue Mar 19 16:14:59 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 \<open>Relative interior of a set\<close>
definition%important "rel_interior S =
- {x. \<exists>T. openin (subtopology euclidean (affine hull S)) T \<and> x \<in> T \<and> T \<subseteq> S}"
+ {x. \<exists>T. openin (top_of_set (affine hull S)) T \<and> x \<in> T \<and> T \<subseteq> S}"
lemma rel_interior_mono:
"\<lbrakk>S \<subseteq> T; affine hull S = affine hull T\<rbrakk>
@@ -327,7 +327,7 @@
by (auto simp: rel_interior_def)
lemma rel_interior_maximal:
- "\<lbrakk>T \<subseteq> S; openin(subtopology euclidean (affine hull S)) T\<rbrakk> \<Longrightarrow> T \<subseteq> (rel_interior S)"
+ "\<lbrakk>T \<subseteq> S; openin(top_of_set (affine hull S)) T\<rbrakk> \<Longrightarrow> T \<subseteq> (rel_interior S)"
by (auto simp: rel_interior_def)
lemma rel_interior:
@@ -651,20 +651,20 @@
definition%important "rel_open S \<longleftrightarrow> rel_interior S = S"
-lemma rel_open: "rel_open S \<longleftrightarrow> openin (subtopology euclidean (affine hull S)) S"
+lemma rel_open: "rel_open S \<longleftrightarrow> 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 \<longleftrightarrow> openin(subtopology euclidean (affine hull s)) s"
+ "rel_interior s = s \<longleftrightarrow> 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 \<Longrightarrow> rel_interior s = s"
+ "openin(top_of_set (affine hull s)) s \<Longrightarrow> 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 \<subseteq> T"
using \<open>open T\<close> \<open>a \<in> T\<close> by (auto simp: open_contains_cball)
- have "openin (subtopology euclidean S) {a}"
+ have "openin (top_of_set S) {a}"
unfolding openin_open using that \<open>a \<in> S\<close> 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 \<open>connected S\<close> connected_clopen S by blast
--- a/src/HOL/Analysis/Elementary_Metric_Spaces.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Elementary_Metric_Spaces.thy Tue Mar 19 16:14:59 2019 +0000
@@ -2718,7 +2718,7 @@
subsection \<open>With Abstract Topology (TODO: move and remove dependency?)\<close>
lemma openin_contains_ball:
- "openin (subtopology euclidean t) s \<longleftrightarrow>
+ "openin (top_of_set t) s \<longleftrightarrow>
s \<subseteq> t \<and> (\<forall>x \<in> s. \<exists>e. 0 < e \<and> ball x e \<inter> t \<subseteq> s)"
(is "?lhs = ?rhs")
proof
@@ -2735,7 +2735,7 @@
qed
lemma openin_contains_cball:
- "openin (subtopology euclidean t) s \<longleftrightarrow>
+ "openin (top_of_set t) s \<longleftrightarrow>
s \<subseteq> t \<and>
(\<forall>x \<in> s. \<exists>e. 0 < e \<and> cball x e \<inter> t \<subseteq> s)"
apply (simp add: openin_contains_ball)
--- a/src/HOL/Analysis/Elementary_Normed_Spaces.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Elementary_Normed_Spaces.thy Tue Mar 19 16:14:59 2019 +0000
@@ -915,7 +915,7 @@
theorem Baire:
fixes S::"'a::{real_normed_vector,heine_borel} set"
assumes "closed S" "countable \<G>"
- and ope: "\<And>T. T \<in> \<G> \<Longrightarrow> openin (subtopology euclidean S) T \<and> S \<subseteq> closure T"
+ and ope: "\<And>T. T \<in> \<G> \<Longrightarrow> openin (top_of_set S) T \<and> S \<subseteq> closure T"
shows "S \<subseteq> closure(\<Inter>\<G>)"
proof (cases "\<G> = {}")
case True
@@ -930,31 +930,31 @@
proof (clarsimp simp: closure_approachable)
fix x and e::real
assume "x \<in> S" "0 < e"
- obtain TF where opeF: "\<And>n. openin (subtopology euclidean S) (TF n)"
+ obtain TF where opeF: "\<And>n. openin (top_of_set S) (TF n)"
and ne: "\<And>n. TF n \<noteq> {}"
and subg: "\<And>n. S \<inter> closure(TF n) \<subseteq> ?g n"
and subball: "\<And>n. closure(TF n) \<subseteq> ball x e"
and decr: "\<And>n. TF(Suc n) \<subseteq> TF n"
proof -
- have *: "\<exists>Y. (openin (subtopology euclidean S) Y \<and> Y \<noteq> {} \<and>
+ have *: "\<exists>Y. (openin (top_of_set S) Y \<and> Y \<noteq> {} \<and>
S \<inter> closure Y \<subseteq> ?g n \<and> closure Y \<subseteq> ball x e) \<and> Y \<subseteq> U"
- if opeU: "openin (subtopology euclidean S) U" and "U \<noteq> {}" and cloU: "closure U \<subseteq> ball x e" for U n
+ if opeU: "openin (top_of_set S) U" and "U \<noteq> {}" and cloU: "closure U \<subseteq> ball x e" for U n
proof -
obtain T where T: "open T" "U = T \<inter> S"
- using \<open>openin (subtopology euclidean S) U\<close> by (auto simp: openin_subtopology)
+ using \<open>openin (top_of_set S) U\<close> by (auto simp: openin_subtopology)
with \<open>U \<noteq> {}\<close> have "T \<inter> closure (?g n) \<noteq> {}"
using gin ope by fastforce
then have "T \<inter> ?g n \<noteq> {}"
using \<open>open T\<close> open_Int_closure_eq_empty by blast
then obtain y where "y \<in> U" "y \<in> ?g n"
using T ope [of "?g n", OF gin] by (blast dest: openin_imp_subset)
- moreover have "openin (subtopology euclidean S) (U \<inter> ?g n)"
+ moreover have "openin (top_of_set S) (U \<inter> ?g n)"
using gin ope opeU by blast
ultimately obtain d where U: "U \<inter> ?g n \<subseteq> S" and "d > 0" and d: "ball y d \<inter> S \<subseteq> U \<inter> ?g n"
by (force simp: openin_contains_ball)
show ?thesis
proof (intro exI conjI)
- show "openin (subtopology euclidean S) (S \<inter> ball y (d/2))"
+ show "openin (top_of_set S) (S \<inter> ball y (d/2))"
by (simp add: openin_open_Int)
show "S \<inter> ball y (d/2) \<noteq> {}"
using \<open>0 < d\<close> \<open>y \<in> U\<close> opeU openin_imp_subset by fastforce
@@ -979,14 +979,14 @@
using ball_divide_subset_numeral d by blast
qed
qed
- let ?\<Phi> = "\<lambda>n X. openin (subtopology euclidean S) X \<and> X \<noteq> {} \<and>
+ let ?\<Phi> = "\<lambda>n X. openin (top_of_set S) X \<and> X \<noteq> {} \<and>
S \<inter> closure X \<subseteq> ?g n \<and> closure X \<subseteq> ball x e"
have "closure (S \<inter> ball x (e / 2)) \<subseteq> closure(ball x (e/2))"
by (simp add: closure_mono)
also have "... \<subseteq> ball x e"
using \<open>e > 0\<close> by auto
finally have "closure (S \<inter> ball x (e / 2)) \<subseteq> ball x e" .
- moreover have"openin (subtopology euclidean S) (S \<inter> ball x (e / 2))" "S \<inter> ball x (e / 2) \<noteq> {}"
+ moreover have"openin (top_of_set S) (S \<inter> ball x (e / 2))" "S \<inter> ball x (e / 2) \<noteq> {}"
using \<open>0 < e\<close> \<open>x \<in> S\<close> by auto
ultimately obtain Y where Y: "?\<Phi> 0 Y \<and> Y \<subseteq> S \<inter> ball x (e / 2)"
using * [of "S \<inter> ball x (e/2)" 0] by metis
--- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Tue Mar 19 16:14:59 2019 +0000
@@ -2032,7 +2032,7 @@
subsection\<open>Negligibility is a local property\<close>
lemma locally_negligible_alt:
- "negligible S \<longleftrightarrow> (\<forall>x \<in> S. \<exists>U. openin (subtopology euclidean S) U \<and> x \<in> U \<and> negligible U)"
+ "negligible S \<longleftrightarrow> (\<forall>x \<in> S. \<exists>U. openin (top_of_set S) U \<and> x \<in> U \<and> 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: "\<And>x. x \<in> S \<Longrightarrow> openin (subtopology euclidean S) (U x)"
+ then obtain U where ope: "\<And>x. x \<in> S \<Longrightarrow> openin (top_of_set S) (U x)"
and cov: "\<And>x. x \<in> S \<Longrightarrow> x \<in> U x"
and neg: "\<And>x. x \<in> S \<Longrightarrow> negligible (U x)"
by metis
@@ -4493,7 +4493,7 @@
fix a
assume a: "a \<in> U n" and fa: "f a \<in> T"
define V where "V \<equiv> {x. x \<in> U n \<and> f x \<in> T} \<inter> ball a (1 / n / 2)"
- show "\<exists>V. openin (subtopology euclidean {x \<in> U n. f x \<in> T}) V \<and> a \<in> V \<and> negligible V"
+ show "\<exists>V. openin (top_of_set {x \<in> U n. f x \<in> T}) V \<and> a \<in> V \<and> negligible V"
proof (intro exI conjI)
have noxy: "norm(x - y) \<le> n * norm(f x - f y)" if "x \<in> V" "y \<in> V" for x y
using that unfolding U_def V_def mem_Collect_eq Int_iff mem_ball dist_norm
--- a/src/HOL/Analysis/Function_Topology.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Function_Topology.thy Tue Mar 19 16:14:59 2019 +0000
@@ -7,9 +7,9 @@
begin
-section \<open>Product Topology\<close> (* FIX see comments by the author *)
+section \<open>Function Topology\<close>
-text \<open>We want to define the product topology.
+text \<open>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 \<open>T1\<close> to the topology \<open>T2\<close>. Then
the current \<open>continuous_on\<close> (with type classes) can be redefined as
@{text [display] \<open>continuous_on s f =
- continuous_on_topo (subtopology euclidean s) (topology euclidean) f\<close>}
+ continuous_on_topo (top_of_set s) (topology euclidean) f\<close>}
In fact, I need \<open>continuous_on_topo\<close> to express the continuity of the projection on subfactors
for the product topology, in Lemma~\<open>continuous_on_restrict_product_topology\<close>, and I show
@@ -61,369 +61,6 @@
I realized afterwards that this theory has a lot in common with \<^file>\<open>~~/src/HOL/Library/Finite_Map.thy\<close>.
\<close>
-subsection \<open>Topology without type classes\<close>
-
-subsubsection%important \<open>The topology generated by some (open) subsets\<close>
-
-text \<open>In the definition below of a generated topology, the \<open>Empty\<close> case is not necessary,
-as it follows from \<open>UN\<close> taking for \<open>K\<close> 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 \<open>UNIV\<close> to be an open set, as this will not be the case in applications. (We are
-thinking of a topology on a subset of \<open>UNIV\<close>, the remaining part of \<open>UNIV\<close> being irrelevant.)\<close>
-
-inductive generate_topology_on for S where
- Empty: "generate_topology_on S {}"
-| Int: "generate_topology_on S a \<Longrightarrow> generate_topology_on S b \<Longrightarrow> generate_topology_on S (a \<inter> b)"
-| UN: "(\<And>k. k \<in> K \<Longrightarrow> generate_topology_on S k) \<Longrightarrow> generate_topology_on S (\<Union>K)"
-| Basis: "s \<in> S \<Longrightarrow> 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 \<open>The basic property of the topology generated by a set \<open>S\<close> is that it is the
-smallest topology containing all the elements of \<open>S\<close>:\<close>
-
-lemma generate_topology_on_coarsest:
- assumes "istopology T"
- "\<And>s. s \<in> S \<Longrightarrow> 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) \<Rightarrow> ('a topology)"
- where "topology_generated_by S \<equiv> topology (generate_topology_on S)"
-
-lemma openin_topology_generated_by_iff:
- "openin (topology_generated_by S) s \<longleftrightarrow> 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 \<Longrightarrow> generate_topology_on S s"
-using openin_topology_generated_by_iff by auto
-
-lemma topology_generated_by_topspace:
- "topspace (topology_generated_by S) = (\<Union>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 \<subseteq> (\<Union>S)" by (induct, auto)
- }
- then show "topspace (topology_generated_by S) \<subseteq> (\<Union>S)"
- unfolding topspace_def by auto
-next
- have "generate_topology_on S (\<Union>S)"
- using generate_topology_on.UN[OF generate_topology_on.Basis, of S S] by simp
- then show "(\<Union>S) \<subseteq> topspace (topology_generated_by S)"
- unfolding topspace_def using openin_topology_generated_by_iff by auto
-qed
-
-lemma topology_generated_by_Basis:
- "s \<in> S \<Longrightarrow> 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:
- "\<lbrakk>finite \<F>; \<And>K. K \<in> \<F> \<Longrightarrow> generate_topology_on \<S> K; \<F> \<noteq> {}\<rbrakk> \<Longrightarrow> generate_topology_on \<S> (\<Inter>\<F>)"
- by (induction \<F> rule: finite_induct; force intro: generate_topology_on.intros)
-
-subsection\<open>Topology bases and sub-bases\<close>
-
-lemma istopology_base_alt:
- "istopology (arbitrary union_of P) \<longleftrightarrow>
- (\<forall>S T. (arbitrary union_of P) S \<and> (arbitrary union_of P) T
- \<longrightarrow> (arbitrary union_of P) (S \<inter> T))"
- by (simp add: istopology_def) (blast intro: arbitrary_union_of_Union)
-
-lemma istopology_base_eq:
- "istopology (arbitrary union_of P) \<longleftrightarrow>
- (\<forall>S T. P S \<and> P T \<longrightarrow> (arbitrary union_of P) (S \<inter> T))"
- by (simp add: istopology_base_alt arbitrary_union_of_Int_eq)
-
-lemma istopology_base:
- "(\<And>S T. \<lbrakk>P S; P T\<rbrakk> \<Longrightarrow> P(S \<inter> T)) \<Longrightarrow> 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 \<longleftrightarrow>
- (\<forall>V. P V \<longrightarrow> openin X V) \<and> (\<forall>U x. openin X U \<and> x \<in> U \<longrightarrow> (\<exists>V. P V \<and> x \<in> V \<and> V \<subseteq> U))"
- (is "?lhs = ?rhs")
-proof
- assume ?lhs
- then show ?rhs
- by (auto simp: union_of_def arbitrary_def)
-next
- assume R: ?rhs
- then have *: "\<exists>\<U>\<subseteq>Collect P. \<Union>\<U> = S" if "openin X S" for S
- using that by (rule_tac x="{V. P V \<and> V \<subseteq> S}" in exI) fastforce
- from R show ?lhs
- by (fastforce simp add: union_of_def arbitrary_def intro: *)
-qed
-
-lemma topology_base_unique:
- "\<lbrakk>\<And>S. P S \<Longrightarrow> openin X S;
- \<And>U x. \<lbrakk>openin X U; x \<in> U\<rbrakk> \<Longrightarrow> \<exists>B. P B \<and> x \<in> B \<and> B \<subseteq> U\<rbrakk>
- \<Longrightarrow> topology(arbitrary union_of P) = X"
- by (metis openin_topology_base_unique openin_inverse [of X])
-
-lemma topology_bases_eq_aux:
- "\<lbrakk>(arbitrary union_of P) S;
- \<And>U x. \<lbrakk>P U; x \<in> U\<rbrakk> \<Longrightarrow> \<exists>V. Q V \<and> x \<in> V \<and> V \<subseteq> U\<rbrakk>
- \<Longrightarrow> (arbitrary union_of Q) S"
- by (metis arbitrary_union_of_alt arbitrary_union_of_idempot)
-
-lemma topology_bases_eq:
- "\<lbrakk>\<And>U x. \<lbrakk>P U; x \<in> U\<rbrakk> \<Longrightarrow> \<exists>V. Q V \<and> x \<in> V \<and> V \<subseteq> U;
- \<And>V x. \<lbrakk>Q V; x \<in> V\<rbrakk> \<Longrightarrow> \<exists>U. P U \<and> x \<in> U \<and> U \<subseteq> V\<rbrakk>
- \<Longrightarrow> 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
- \<longleftrightarrow> (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 \<subseteq> U"
- by (metis arbitrary_union_of_relative_to openin_subbase openin_topspace relative_to_imp_subset)
- show "U \<subseteq> ?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:
- "\<lbrakk>\<And>S. P S \<Longrightarrow> openin X S; openin X U;
- openin(topology(arbitrary union_of (finite intersection_of P relative_to U))) S\<rbrakk>
- \<Longrightarrow> 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 (\<lambda>x. x \<in> 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 \<U> where \<U>: "\<And>T. T \<in> \<U> \<Longrightarrow> \<exists>\<F>. finite' \<F> \<and> \<F> \<subseteq> S \<and> \<Inter>\<F> = T" and eq: "A = \<Union>\<U>"
- unfolding union_of_def intersection_of_def by auto
- show "?lhs A"
- unfolding eq
- proof (rule generate_topology_on.UN)
- fix T
- assume "T \<in> \<U>"
- with \<U> obtain \<F> where "finite' \<F>" "\<F> \<subseteq> S" "\<Inter>\<F> = T"
- by blast
- have "generate_topology_on S (\<Inter>\<F>)"
- proof (rule generate_topology_on_Inter)
- show "finite \<F>" "\<F> \<noteq> {}"
- by (auto simp: \<open>finite' \<F>\<close>)
- show "\<And>K. K \<in> \<F> \<Longrightarrow> generate_topology_on S K"
- by (metis \<open>\<F> \<subseteq> S\<close> generate_topology_on.simps subset_iff)
- qed
- then show "generate_topology_on S T"
- using \<open>\<Inter>\<F> = T\<close> by blast
- qed
-qed
-
-subsubsection%important \<open>Continuity\<close>
-
-text \<open>We will need to deal with continuous maps in terms of topologies and not in terms
-of type classes, as defined below.\<close>
-
-definition%important continuous_on_topo::"'a topology \<Rightarrow> 'b topology \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
- where "continuous_on_topo T1 T2 f = ((\<forall> U. openin T2 U \<longrightarrow> openin T1 (f-`U \<inter> topspace(T1)))
- \<and> (f`(topspace T1) \<subseteq> (topspace T2)))"
-
-lemma continuous_on_continuous_on_topo:
- "continuous_on s f \<longleftrightarrow> 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 \<longleftrightarrow> 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 \<Longrightarrow> openin T2 U \<Longrightarrow> openin T1 (f-`U \<inter> topspace(T1))"
- unfolding continuous_on_topo_def by auto
-
-lemma continuous_on_topo_topspace [intro]:
- "continuous_on_topo T1 T2 f \<Longrightarrow> f`(topspace T1) \<subseteq> (topspace T2)"
-unfolding continuous_on_topo_def by auto
-
-lemma continuous_on_generated_topo_iff:
- "continuous_on_topo T1 (topology_generated_by S) f \<longleftrightarrow>
- ((\<forall>U. U \<in> S \<longrightarrow> openin T1 (f-`U \<inter> topspace(T1))) \<and> (f`(topspace T1) \<subseteq> (\<Union> S)))"
-unfolding continuous_on_topo_def topology_generated_by_topspace
-proof (auto simp add: topology_generated_by_Basis)
- assume H: "\<forall>U. U \<in> S \<longrightarrow> openin T1 (f -` U \<inter> 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 \<inter> topspace T1)"
- proof (induct)
- fix a b
- assume H: "openin T1 (f -` a \<inter> topspace T1)" "openin T1 (f -` b \<inter> topspace T1)"
- have "f -` (a \<inter> b) \<inter> topspace T1 = (f-`a \<inter> topspace T1) \<inter> (f-`b \<inter> topspace T1)"
- by auto
- then show "openin T1 (f -` (a \<inter> b) \<inter> topspace T1)" using H by auto
- next
- fix K
- assume H: "openin T1 (f -` k \<inter> topspace T1)" if "k\<in> K" for k
- define L where "L = {f -` k \<inter> topspace T1|k. k \<in> K}"
- have *: "openin T1 l" if "l \<in>L" for l using that H unfolding L_def by auto
- have "openin T1 (\<Union>L)" using openin_Union[OF *] by simp
- moreover have "(\<Union>L) = (f -` \<Union>K \<inter> topspace T1)" unfolding L_def by auto
- ultimately show "openin T1 (f -` \<Union>K \<inter> topspace T1)" by simp
- qed (auto simp add: H)
-qed
-
-lemma continuous_on_generated_topo:
- assumes "\<And>U. U \<in>S \<Longrightarrow> openin T1 (f-`U \<inter> topspace(T1))"
- "f`(topspace T1) \<subseteq> (\<Union> 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 \<inter> topspace T2) \<inter> topspace T1)"
- using H assms by blast
- moreover have "f -` (g -` U \<inter> topspace T2) \<inter> topspace T1 = (g \<circ> f) -` U \<inter> topspace T1"
- using H assms continuous_on_topo_topspace by fastforce
- ultimately show "openin T1 ((g \<circ> f) -` U \<inter> topspace T1)"
- by simp
-qed (blast)
-
-lemma continuous_on_topo_preimage_topspace [intro]:
- assumes "continuous_on_topo T1 T2 f"
- shows "f-`(topspace T2) \<inter> topspace T1 = topspace T1"
-using assms unfolding continuous_on_topo_def by auto
-
-
-subsubsection%important \<open>Pullback topology\<close>
-
-text \<open>Pulling back a topology by map gives again a topology. \<open>subtopology\<close> 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.\<close>
-
-text \<open>\<open>pullback_topology A f T\<close> is the pullback of the topology \<open>T\<close> by the map \<open>f\<close> on
-the set \<open>A\<close>.\<close>
-
-definition%important pullback_topology::"('a set) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('b topology) \<Rightarrow> ('a topology)"
- where "pullback_topology A f T = topology (\<lambda>S. \<exists>U. openin T U \<and> S = f-`U \<inter> A)"
-
-lemma istopology_pullback_topology:
- "istopology (\<lambda>S. \<exists>U. openin T U \<and> S = f-`U \<inter> A)"
- unfolding istopology_def proof (auto)
- fix K assume "\<forall>S\<in>K. \<exists>U. openin T U \<and> S = f -` U \<inter> A"
- then have "\<exists>U. \<forall>S\<in>K. openin T (U S) \<and> S = f-`(U S) \<inter> A"
- by (rule bchoice)
- then obtain U where U: "\<forall>S\<in>K. openin T (U S) \<and> S = f-`(U S) \<inter> A"
- by blast
- define V where "V = (\<Union>S\<in>K. U S)"
- have "openin T V" "\<Union>K = f -` V \<inter> A" unfolding V_def using U by auto
- then show "\<exists>V. openin T V \<and> \<Union>K = f -` V \<inter> A" by auto
-qed
-
-lemma openin_pullback_topology:
- "openin (pullback_topology A f T) S \<longleftrightarrow> (\<exists>U. openin T U \<and> S = f-`U \<inter> 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) \<inter> 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 \<inter> topspace T1)"
- using assms unfolding continuous_on_topo_def by auto
- have "(g o f)-`U \<inter> topspace (pullback_topology A f T1) = (g o f)-`U \<inter> A \<inter> f-`(topspace T1)"
- unfolding topspace_pullback_topology by auto
- also have "... = f-`(g-`U \<inter> topspace T1) \<inter> A "
- by auto
- also have "openin (pullback_topology A f T1) (...)"
- unfolding openin_pullback_topology using \<open>openin T1 (g-`U \<inter> topspace T1)\<close> by auto
- finally show "openin (pullback_topology A f T1) ((g \<circ> f) -` U \<inter> topspace (pullback_topology A f T1))"
- by auto
-next
- fix x assume "x \<in> topspace (pullback_topology A f T1)"
- then have "f x \<in> topspace T1"
- unfolding topspace_pullback_topology by auto
- then show "g (f x) \<in> 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 \<subseteq> 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 "\<exists>V. openin T2 V \<and> U = f-`V \<inter> A"
- unfolding openin_pullback_topology by auto
- then obtain V where "openin T2 V" "U = f-`V \<inter> A"
- by blast
- then have "g -` U \<inter> topspace T1 = g-`(f-`V \<inter> A) \<inter> topspace T1"
- by blast
- also have "... = (f o g)-`V \<inter> (g-`A \<inter> topspace T1)"
- by auto
- also have "... = (f o g)-`V \<inter> topspace T1"
- using assms(2) by auto
- also have "openin T1 (...)"
- using assms(1) \<open>openin T2 V\<close> by auto
- finally show "openin T1 (g -` U \<inter> topspace T1)" by simp
-next
- fix x assume "x \<in> topspace T1"
- have "(f o g) x \<in> topspace T2"
- using assms(1) \<open>x \<in> topspace T1\<close> unfolding continuous_on_topo_def by auto
- then have "g x \<in> f-`(topspace T2)"
- unfolding comp_def by blast
- moreover have "g x \<in> A" using assms(2) \<open>x \<in> topspace T1\<close> by blast
- ultimately show "g x \<in> topspace (pullback_topology A f T2)"
- unfolding topspace_pullback_topology by blast
-qed
-
-
subsection \<open>The product topology\<close>
text \<open>We can now define the product topology, as generated by
@@ -605,6 +242,9 @@
then show ?thesis using \<open>x \<in> U\<close> by auto
qed
+lemma product_topology_empty_discrete:
+ "product_topology T {} = discrete_topology {(\<lambda>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 (\<lambda>S. S \<in> {{},{\<lambda>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 \<longleftrightarrow> S \<in> {{},{\<lambda>k. undefined}}"
+ unfolding union_of_def intersection_of_def arbitrary_def relative_to_def openin_product_topology
+ by auto
+
subsubsection \<open>The basic property of the product topology is the continuity of projections:\<close>
@@ -949,10 +598,10 @@
by (rule continuous_on_topo_product_coordinates, simp)
lemma continuous_on_coordinatewise_then_product [intro, continuous_intros]:
- fixes f :: "('a \<Rightarrow> real) \<Rightarrow> 'b \<Rightarrow> real"
+ fixes f :: "'a::topological_space \<Rightarrow> 'b \<Rightarrow> 'c::topological_space"
assumes "\<And>i. continuous_on S (\<lambda>x. f x i)"
shows "continuous_on S f"
- using continuous_on_topo_coordinatewise_then_product [of UNIV, where T = "\<lambda>i. euclideanreal"]
+ using continuous_on_topo_coordinatewise_then_product [of UNIV, where T = "\<lambda>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\<open>The Alexander subbase theorem\<close>
+
+theorem Alexander_subbase:
+ assumes X: "topology (arbitrary union_of (finite intersection_of (\<lambda>x. x \<in> \<B>) relative_to \<Union>\<B>)) = X"
+ and fin: "\<And>C. \<lbrakk>C \<subseteq> \<B>; \<Union>C = topspace X\<rbrakk> \<Longrightarrow> \<exists>C'. finite C' \<and> C' \<subseteq> C \<and> \<Union>C' = topspace X"
+ shows "compact_space X"
+proof -
+ have U\<B>: "\<Union>\<B> = topspace X"
+ by (simp flip: X)
+ have False if \<U>: "\<forall>U\<in>\<U>. openin X U" and sub: "topspace X \<subseteq> \<Union>\<U>"
+ and neg: "\<And>\<F>. \<lbrakk>\<F> \<subseteq> \<U>; finite \<F>\<rbrakk> \<Longrightarrow> \<not> topspace X \<subseteq> \<Union>\<F>" for \<U>
+ proof -
+ define \<A> where "\<A> \<equiv> {\<C>. (\<forall>U \<in> \<C>. openin X U) \<and> topspace X \<subseteq> \<Union>\<C> \<and> (\<forall>\<F>. finite \<F> \<longrightarrow> \<F> \<subseteq> \<C> \<longrightarrow> ~(topspace X \<subseteq> \<Union>\<F>))}"
+ have 1: "\<A> \<noteq> {}"
+ unfolding \<A>_def using sub \<U> neg by force
+ have 2: "\<Union>\<C> \<in> \<A>" if "\<C>\<noteq>{}" and \<C>: "subset.chain \<A> \<C>" for \<C>
+ unfolding \<A>_def
+ proof (intro CollectI conjI ballI allI impI notI)
+ show "openin X U" if U: "U \<in> \<Union>\<C>" for U
+ using U \<C> unfolding \<A>_def subset_chain_def by force
+ have "\<C> \<subseteq> \<A>"
+ using subset_chain_def \<C> by blast
+ with that \<A>_def show UUC: "topspace X \<subseteq> \<Union>(\<Union>\<C>)"
+ by blast
+ show "False" if "finite \<F>" and "\<F> \<subseteq> \<Union>\<C>" and "topspace X \<subseteq> \<Union>\<F>" for \<F>
+ proof -
+ obtain \<B> where "\<B> \<in> \<C>" "\<F> \<subseteq> \<B>"
+ by (metis Sup_empty \<C> \<open>\<F> \<subseteq> \<Union>\<C>\<close> \<open>finite \<F>\<close> UUC empty_subsetI finite.emptyI finite_subset_Union_chain neg)
+ then show False
+ using \<A>_def \<open>\<C> \<subseteq> \<A>\<close> \<open>finite \<F>\<close> \<open>topspace X \<subseteq> \<Union>\<F>\<close> by blast
+ qed
+ qed
+ obtain \<K> where "\<K> \<in> \<A>" and "\<And>X. \<lbrakk>X\<in>\<A>; \<K> \<subseteq> X\<rbrakk> \<Longrightarrow> X = \<K>"
+ using subset_Zorn_nonempty [OF 1 2] by metis
+ then have *: "\<And>\<W>. \<lbrakk>\<And>W. W\<in>\<W> \<Longrightarrow> openin X W; topspace X \<subseteq> \<Union>\<W>; \<K> \<subseteq> \<W>;
+ \<And>\<F>. \<lbrakk>finite \<F>; \<F> \<subseteq> \<W>; topspace X \<subseteq> \<Union>\<F>\<rbrakk> \<Longrightarrow> False\<rbrakk>
+ \<Longrightarrow> \<W> = \<K>"
+ and ope: "\<forall>U\<in>\<K>. openin X U" and top: "topspace X \<subseteq> \<Union>\<K>"
+ and non: "\<And>\<F>. \<lbrakk>finite \<F>; \<F> \<subseteq> \<K>; topspace X \<subseteq> \<Union>\<F>\<rbrakk> \<Longrightarrow> False"
+ unfolding \<A>_def by simp_all metis+
+ then obtain x where "x \<in> topspace X" "x \<notin> \<Union>(\<B> \<inter> \<K>)"
+ proof -
+ have "\<Union>(\<B> \<inter> \<K>) \<noteq> \<Union>\<B>"
+ by (metis \<open>\<Union>\<B> = topspace X\<close> fin inf.bounded_iff non order_refl)
+ then have "\<exists>a. a \<notin> \<Union>(\<B> \<inter> \<K>) \<and> a \<in> \<Union>\<B>"
+ by blast
+ then show ?thesis
+ using that by (metis U\<B>)
+ qed
+ obtain C where C: "openin X C" "C \<in> \<K>" "x \<in> C"
+ using \<open>x \<in> topspace X\<close> ope top by auto
+ then have "C \<subseteq> topspace X"
+ by (metis openin_subset)
+ then have "(arbitrary union_of (finite intersection_of (\<lambda>x. x \<in> \<B>) relative_to \<Union>\<B>)) C"
+ using openin_subbase C unfolding X [symmetric] by blast
+ moreover have "C \<noteq> topspace X"
+ using \<open>\<K> \<in> \<A>\<close> \<open>C \<in> \<K>\<close> unfolding \<A>_def by blast
+ ultimately obtain \<V> W where W: "(finite intersection_of (\<lambda>x. x \<in> \<B>) relative_to topspace X) W"
+ and "x \<in> W" "W \<in> \<V>" "\<Union>\<V> \<noteq> topspace X" "C = \<Union>\<V>"
+ using C by (auto simp: union_of_def U\<B>)
+ then have "\<Union>\<V> \<subseteq> topspace X"
+ by (metis \<open>C \<subseteq> topspace X\<close>)
+ then have "topspace X \<notin> \<V>"
+ using \<open>\<Union>\<V> \<noteq> topspace X\<close> by blast
+ then obtain \<B>' where \<B>': "finite \<B>'" "\<B>' \<subseteq> \<B>" "x \<in> \<Inter>\<B>'" "W = topspace X \<inter> \<Inter>\<B>'"
+ using W \<open>x \<in> W\<close> unfolding relative_to_def intersection_of_def by auto
+ then have "\<Inter>\<B>' \<subseteq> \<Union>\<B>"
+ using \<open>W \<in> \<V>\<close> \<open>\<Union>\<V> \<noteq> topspace X\<close> \<open>\<Union>\<V> \<subseteq> topspace X\<close> by blast
+ then have "\<Inter>\<B>' \<subseteq> C"
+ using U\<B> \<open>C = \<Union>\<V>\<close> \<open>W = topspace X \<inter> \<Inter>\<B>'\<close> \<open>W \<in> \<V>\<close> by auto
+ have "\<forall>b \<in> \<B>'. \<exists>C'. finite C' \<and> C' \<subseteq> \<K> \<and> topspace X \<subseteq> \<Union>(insert b C')"
+ proof
+ fix b
+ assume "b \<in> \<B>'"
+ have "insert b \<K> = \<K>" if neg: "\<not> (\<exists>C'. finite C' \<and> C' \<subseteq> \<K> \<and> topspace X \<subseteq> \<Union>(insert b C'))"
+ proof (rule *)
+ show "openin X W" if "W \<in> insert b \<K>" for W
+ using that
+ proof
+ have "b \<in> \<B>"
+ using \<open>b \<in> \<B>'\<close> \<open>\<B>' \<subseteq> \<B>\<close> by blast
+ then have "\<exists>\<U>. finite \<U> \<and> \<U> \<subseteq> \<B> \<and> \<Inter>\<U> = b"
+ by (rule_tac x="{b}" in exI) auto
+ moreover have "\<Union>\<B> \<inter> b = b"
+ using \<B>'(2) \<open>b \<in> \<B>'\<close> by auto
+ ultimately show "openin X W" if "W = b"
+ using that \<open>b \<in> \<B>'\<close>
+ 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 \<in> \<K>"
+ by (simp add: \<open>W \<in> \<K>\<close> ope)
+ qed
+ next
+ show "topspace X \<subseteq> \<Union> (insert b \<K>)"
+ using top by auto
+ next
+ show False if "finite \<F>" and "\<F> \<subseteq> insert b \<K>" "topspace X \<subseteq> \<Union>\<F>" for \<F>
+ proof -
+ have "insert b (\<F> \<inter> \<K>) = \<F>"
+ using non that by blast
+ then show False
+ by (metis Int_lower2 finite_insert neg that(1) that(3))
+ qed
+ qed auto
+ then show "\<exists>C'. finite C' \<and> C' \<subseteq> \<K> \<and> topspace X \<subseteq> \<Union>(insert b C')"
+ using \<open>b \<in> \<B>'\<close> \<open>x \<notin> \<Union>(\<B> \<inter> \<K>)\<close> \<B>'
+ by (metis IntI InterE Union_iff subsetD insertI1)
+ qed
+ then obtain F where F: "\<forall>b \<in> \<B>'. finite (F b) \<and> F b \<subseteq> \<K> \<and> topspace X \<subseteq> \<Union>(insert b (F b))"
+ by metis
+ let ?\<D> = "insert C (\<Union>(F ` \<B>'))"
+ show False
+ proof (rule non)
+ have "topspace X \<subseteq> (\<Inter>b \<in> \<B>'. \<Union>(insert b (F b)))"
+ using F by (simp add: INT_greatest)
+ also have "\<dots> \<subseteq> \<Union>?\<D>"
+ using \<open>\<Inter>\<B>' \<subseteq> C\<close> by force
+ finally show "topspace X \<subseteq> \<Union>?\<D>" .
+ show "?\<D> \<subseteq> \<K>"
+ using \<open>C \<in> \<K>\<close> F by auto
+ show "finite ?\<D>"
+ using \<open>finite \<B>'\<close> F by auto
+ qed
+ qed
+ then show ?thesis
+ by (force simp: compact_space_def compactin_def)
+qed
+
+
+corollary Alexander_subbase_alt:
+ assumes "U \<subseteq> \<Union>\<B>"
+ and fin: "\<And>C. \<lbrakk>C \<subseteq> \<B>; U \<subseteq> \<Union>C\<rbrakk> \<Longrightarrow> \<exists>C'. finite C' \<and> C' \<subseteq> C \<and> U \<subseteq> \<Union>C'"
+ and X: "topology
+ (arbitrary union_of
+ (finite intersection_of (\<lambda>x. x \<in> \<B>) relative_to U)) = X"
+ shows "compact_space X"
+proof -
+ have "topspace X = U"
+ using X topspace_subbase by fastforce
+ have eq: "\<Union> (Collect ((\<lambda>x. x \<in> \<B>) relative_to U)) = U"
+ unfolding relative_to_def
+ using \<open>U \<subseteq> \<Union>\<B>\<close> by blast
+ have *: "\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<C> \<and> \<Union>\<F> = topspace X"
+ if "\<C> \<subseteq> Collect ((\<lambda>x. x \<in> \<B>) relative_to topspace X)" and UC: "\<Union>\<C> = topspace X" for \<C>
+ proof -
+ have "\<C> \<subseteq> (\<lambda>U. topspace X \<inter> U) ` \<B>"
+ using that by (auto simp: relative_to_def)
+ then obtain \<B>' where "\<B>' \<subseteq> \<B>" and \<B>': "\<C> = (\<inter>) (topspace X) ` \<B>'"
+ by (auto simp: subset_image_iff)
+ moreover have "U \<subseteq> \<Union>\<B>'"
+ using \<B>' \<open>topspace X = U\<close> UC by auto
+ ultimately obtain \<C>' where "finite \<C>'" "\<C>' \<subseteq> \<B>'" "U \<subseteq> \<Union>\<C>'"
+ using fin [of \<B>'] \<open>topspace X = U\<close> \<open>U \<subseteq> \<Union>\<B>\<close> by blast
+ then show ?thesis
+ unfolding \<B>' exists_finite_subset_image \<open>topspace X = U\<close> by auto
+ qed
+ show ?thesis
+ apply (rule Alexander_subbase [where \<B> = "Collect ((\<lambda>x. x \<in> \<B>) 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 \<longleftrightarrow>
+ f ` (topspace X) \<subseteq> extensional I \<and> (\<forall>k \<in> I. continuous_map X (Y k) (\<lambda>x. f x k))"
+ (is "?lhs \<longleftrightarrow> _ \<and> ?rhs")
+proof (cases "\<forall>x \<in> topspace X. f x \<in> extensional I")
+ case True
+ then have "f ` (topspace X) \<subseteq> extensional I"
+ by force
+ moreover have ?rhs if L: ?lhs
+ proof -
+ have "openin X {x \<in> topspace X. f x k \<in> U}" if "k \<in> I" and "openin (Y k) U" for k U
+ proof -
+ have "openin (product_topology Y I) ({Y. Y k \<in> U} \<inter> (\<Pi>\<^sub>E i\<in>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 \<in> topspace X. f x \<in> ({Y. Y k \<in> U} \<inter> (\<Pi>\<^sub>E i\<in>I. topspace (Y i)))}"
+ using L unfolding continuous_map_def by blast
+ moreover have "{x \<in> topspace X. f x \<in> ({Y. Y k \<in> U} \<inter> (\<Pi>\<^sub>E i\<in>I. topspace (Y i)))} = {x \<in> topspace X. f x k \<in> 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: "\<And>x. x \<in> topspace X \<Longrightarrow> f x \<in> (\<Pi>\<^sub>E i\<in>I. topspace (Y i))"
+ using that True by (auto simp: continuous_map_def PiE_iff)
+ have 2: "{x \<in> S. \<exists>T\<in>\<T>. f x \<in> T} = (\<Union>T\<in>\<T>. {x \<in> S. f x \<in> T})" for S \<T>
+ by blast
+ have 3: "{x \<in> S. \<forall>U\<in>\<U>. f x \<in> U} = (\<Inter> (insert S ((\<lambda>U. {x \<in> S. f x \<in> U}) ` \<U>)))" for S \<U>
+ by blast
+ show ?thesis
+ unfolding continuous_map_def openin_product_topology arbitrary_def
+ proof (clarsimp simp: all_union_of 1 2)
+ fix \<T>
+ assume \<T>: "\<T> \<subseteq> Collect (finite intersection_of (\<lambda>F. \<exists>i U. F = {f. f i \<in> U} \<and> i \<in> I \<and> openin (Y i) U)
+ relative_to (\<Pi>\<^sub>E i\<in>I. topspace (Y i)))"
+ show "openin X (\<Union>T\<in>\<T>. {x \<in> topspace X. f x \<in> T})"
+ proof (rule openin_Union; clarify)
+ fix S T
+ assume "T \<in> \<T>"
+ obtain \<U> where "T = (\<Pi>\<^sub>E i\<in>I. topspace (Y i)) \<inter> \<Inter>\<U>" and "finite \<U>"
+ "\<U> \<subseteq> {{f. f i \<in> U} |i U. i \<in> I \<and> openin (Y i) U}"
+ using subsetD [OF \<T> \<open>T \<in> \<T>\<close>] by (auto simp: intersection_of_def relative_to_def)
+ with that show "openin X {x \<in> topspace X. f x \<in> 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 \<longleftrightarrow> (\<forall>k. continuous_map X (Y k) (\<lambda>x. f x k))"
+ by (simp add: continuous_map_componentwise)
+
+lemma continuous_map_product_projection [continuous_intros]:
+ "k \<in> I \<Longrightarrow> continuous_map (product_topology X I) (X k) (\<lambda>x. x k)"
+ using continuous_map_componentwise [of "product_topology X I" X I id] by simp
+
+proposition open_map_product_projection:
+ assumes "i \<in> I"
+ shows "open_map (product_topology Y I) (Y i) (\<lambda>f. f i)"
+ unfolding openin_product_topology all_union_of arbitrary_def open_map_def image_Union
+proof clarify
+ fix \<V>
+ assume \<V>: "\<V> \<subseteq> Collect
+ (finite intersection_of
+ (\<lambda>F. \<exists>i U. F = {f. f i \<in> U} \<and> i \<in> I \<and> openin (Y i) U) relative_to
+ topspace (product_topology Y I))"
+ show "openin (Y i) (\<Union>x\<in>\<V>. (\<lambda>f. f i) ` x)"
+ proof (rule openin_Union, clarify)
+ fix S V
+ assume "V \<in> \<V>"
+ obtain \<F> where "finite \<F>"
+ and V: "V = (\<Pi>\<^sub>E i\<in>I. topspace (Y i)) \<inter> \<Inter>\<F>"
+ and \<F>: "\<F> \<subseteq> {{f. f i \<in> U} |i U. i \<in> I \<and> openin (Y i) U}"
+ using subsetD [OF \<V> \<open>V \<in> \<V>\<close>]
+ by (auto simp: intersection_of_def relative_to_def)
+ show "openin (Y i) ((\<lambda>f. f i) ` V)"
+ proof (subst openin_subopen; clarify)
+ fix x f
+ assume "f \<in> V"
+ let ?T = "{a \<in> topspace(Y i).
+ (\<lambda>j. if j = i then a
+ else if j \<in> I then f j else undefined) \<in> (\<Pi>\<^sub>E i\<in>I. topspace (Y i)) \<inter> \<Inter>\<F>}"
+ show "\<exists>T. openin (Y i) T \<and> f i \<in> T \<and> T \<subseteq> (\<lambda>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) (\<lambda>x. if k = i then x else f k)" if "k \<in> 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 \<open>f \<in> V\<close> that)
+ qed
+ then show "continuous_map (Y i) (product_topology Y I)
+ (\<lambda>x j. if j = i then x else if j \<in> I then f j else undefined)"
+ by (auto simp: continuous_map_componentwise assms extensional_def)
+ next
+ have "openin (product_topology Y I) (\<Pi>\<^sub>E i\<in>I. topspace (Y i))"
+ by (metis openin_topspace topspace_product_topology)
+ moreover have "openin (product_topology Y I) (\<Inter>B\<in>\<F>. (\<Pi>\<^sub>E i\<in>I. topspace (Y i)) \<inter> B)"
+ if "\<F> \<noteq> {}"
+ proof -
+ show ?thesis
+ proof (rule openin_Inter)
+ show "\<And>X. X \<in> (\<inter>) (\<Pi>\<^sub>E i\<in>I. topspace (Y i)) ` \<F> \<Longrightarrow> 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 \<F>]
+ apply (force intro: finite_intersection_of_inc)
+ done
+ qed (use \<open>finite \<F>\<close> \<open>\<F> \<noteq> {}\<close> in auto)
+ qed
+ ultimately show "openin (product_topology Y I) ((\<Pi>\<^sub>E i\<in>I. topspace (Y i)) \<inter> \<Inter>\<F>)"
+ by (auto simp only: Int_Inter_eq split: if_split)
+ qed
+ next
+ have eqf: "(\<lambda>j. if j = i then f i else if j \<in> I then f j else undefined) = f"
+ using PiE_arb V \<open>f \<in> V\<close> by force
+ show "f i \<in> ?T"
+ using V assms \<open>f \<in> V\<close> by (auto simp: PiE_iff eqf)
+ next
+ show "?T \<subseteq> (\<lambda>f. f i) ` V"
+ unfolding V by (auto simp: intro!: rev_image_eqI)
+ qed
+ qed
+ qed
+qed
+
+lemma retraction_map_product_projection:
+ assumes "i \<in> I"
+ shows "(retraction_map (product_topology X I) (X i) (\<lambda>x. x i) \<longleftrightarrow>
+ (topspace (product_topology X I) = {}) \<longrightarrow> 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 *: "\<exists>g. continuous_map (X i) (product_topology X I) g \<and> (\<forall>x\<in>topspace (X i). g x i = x)"
+ if z: "z \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))" for z
+ proof -
+ have cm: "continuous_map (X i) (X j) (\<lambda>x. if j = i then x else z j)" if "j \<in> I" for j
+ using \<open>j \<in> I\<close> z by (case_tac "j = i") auto
+ show ?thesis
+ using \<open>i \<in> I\<close> that
+ by (rule_tac x="\<lambda>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 \<open>i \<in> I\<close> False
+ by (auto simp: retraction_map_def retraction_maps_def assms continuous_map_product_projection *)
+ qed
+qed
+
end
--- a/src/HOL/Analysis/Further_Topology.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Further_Topology.thy Tue Mar 19 16:14:59 2019 +0000
@@ -1149,7 +1149,7 @@
assumes "finite K" "affine U" and contf: "continuous_on (U - K) f"
and fim: "f ` (U - K) \<subseteq> T"
and comps: "\<And>C. \<lbrakk>C \<in> components(U - S); C \<inter> K \<noteq> {}\<rbrakk> \<Longrightarrow> C \<inter> L \<noteq> {}"
- and clo: "closedin (subtopology euclidean U) S" and K: "disjnt K S" "K \<subseteq> U"
+ and clo: "closedin (top_of_set U) S" and K: "disjnt K S" "K \<subseteq> U"
obtains g where "continuous_on (U - L) g" "g ` (U - L) \<subseteq> T" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
proof (cases "K = {}")
case True
@@ -1176,11 +1176,11 @@
have "C \<subseteq> U-S" "C \<inter> L \<noteq> {}"
by (simp_all add: in_components_subset comps that)
then obtain a where a: "a \<in> C" "a \<in> 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: \<open>locally connected U\<close> 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 \<subseteq> U" "0 < d" and d: "cball a d \<inter> U \<subseteq> C"
@@ -1192,10 +1192,10 @@
and bou: "bounded {x. (\<not> (h x = x \<and> k x = x))}"
and hin: "\<And>x. x \<in> C \<inter> K \<Longrightarrow> h x \<in> ball a d \<inter> U"
proof (rule homeomorphism_grouping_points_exists_gen [of C "ball a d \<inter> U" "C \<inter> K" "S \<union> C"])
- show "openin (subtopology euclidean C) (ball a d \<inter> U)"
+ show "openin (top_of_set C) (ball a d \<inter> U)"
by (metis open_ball \<open>C \<subseteq> U\<close> \<open>ball a d \<inter> U \<subseteq> C\<close> inf.absorb_iff2 inf.orderE inf_assoc open_openin openin_subtopology)
- show "openin (subtopology euclidean (affine hull C)) C"
- by (metis \<open>a \<in> C\<close> \<open>openin (subtopology euclidean U) C\<close> affine_hull_eq affine_hull_openin all_not_in_conv \<open>affine U\<close>)
+ show "openin (top_of_set (affine hull C)) C"
+ by (metis \<open>a \<in> C\<close> \<open>openin (top_of_set U) C\<close> affine_hull_eq affine_hull_openin all_not_in_conv \<open>affine U\<close>)
show "ball a d \<inter> U \<noteq> {}"
using \<open>0 < d\<close> \<open>C \<subseteq> U\<close> \<open>a \<in> C\<close> by force
show "finite (C \<inter> K)"
@@ -1323,13 +1323,14 @@
obtain g where contg: "continuous_on (S \<union> UF) g"
and gh: "\<And>x i. \<lbrakk>i \<in> F; x \<in> (S \<union> UF) \<inter> (S \<union> (i - {a i}))\<rbrakk>
\<Longrightarrow> g x = h i x"
- proof (rule pasting_lemma_exists_closed [OF \<open>finite F\<close>, of "S \<union> UF" "\<lambda>C. S \<union> (C - {a C})" h])
- show "S \<union> UF \<subseteq> (\<Union>C\<in>F. S \<union> (C - {a C}))"
+ proof (rule pasting_lemma_exists_closed [OF \<open>finite F\<close>])
+ let ?X = "top_of_set (S \<union> UF)"
+ show "topspace ?X \<subseteq> (\<Union>C\<in>F. S \<union> (C - {a C}))"
using \<open>C0 \<in> F\<close> by (force simp: UF_def)
- show "closedin (subtopology euclidean (S \<union> UF)) (S \<union> (C - {a C}))"
+ show "closedin (top_of_set (S \<union> UF)) (S \<union> (C - {a C}))"
if "C \<in> F" for C
proof (rule closedin_closed_subset [of U "S \<union> C"])
- show "closedin (subtopology euclidean U) (S \<union> C)"
+ show "closedin (top_of_set U) (S \<union> C)"
apply (rule closedin_Un_complement_component [OF \<open>locally connected U\<close> clo])
using F_def that by blast
next
@@ -1346,22 +1347,26 @@
show "S \<union> (C - {a C}) = (S \<union> C) \<inter> (S \<union> UF)"
using F_def UF_def components_nonoverlap that by auto
qed
- next
- show "continuous_on (S \<union> (C' - {a C'})) (h C')" if "C' \<in> F" for C'
- using ah F_def that by blast
+ show "continuous_map (subtopology ?X (S \<union> (C' - {a C'}))) euclidean (h C')" if "C' \<in> F" for C'
+ proof -
+ have C': "C' \<in> components (U - S)" "C' \<inter> K \<noteq> {}"
+ 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 "\<And>i j x. \<lbrakk>i \<in> F; j \<in> F;
- x \<in> (S \<union> UF) \<inter> (S \<union> (i - {a i})) \<inter> (S \<union> (j - {a j}))\<rbrakk>
+ x \<in> topspace ?X \<inter> (S \<union> (i - {a i})) \<inter> (S \<union> (j - {a j}))\<rbrakk>
\<Longrightarrow> h i x = h j x"
using components_eq by (fastforce simp: components_eq F_def ah)
- qed blast
+ qed auto
have SU': "S \<union> \<Union>G \<union> (S \<union> UF) \<subseteq> U"
using \<open>S \<subseteq> U\<close> in_components_subset by (auto simp: F_def G_def UF_def)
- have clo1: "closedin (subtopology euclidean (S \<union> \<Union>G \<union> (S \<union> UF))) (S \<union> \<Union>G)"
+ have clo1: "closedin (top_of_set (S \<union> \<Union>G \<union> (S \<union> UF))) (S \<union> \<Union>G)"
proof (rule closedin_closed_subset [OF _ SU'])
- have *: "\<And>C. C \<in> F \<Longrightarrow> openin (subtopology euclidean U) C"
+ have *: "\<And>C. C \<in> F \<Longrightarrow> openin (top_of_set U) C"
unfolding F_def
by clarify (metis (no_types, lifting) \<open>locally connected U\<close> 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 \<union> \<Union>G = (U - UF) \<inter> (S \<union> \<Union>G \<union> (S \<union> 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 \<union> \<Union>G \<union> (S \<union> UF))) (S \<union> UF)"
+ have clo2: "closedin (top_of_set (S \<union> \<Union>G \<union> (S \<union> UF))) (S \<union> UF)"
proof (rule closedin_closed_subset [OF _ SU'])
- show "closedin (subtopology euclidean U) (\<Union>C\<in>F. S \<union> C)"
+ show "closedin (top_of_set U) (\<Union>C\<in>F. S \<union> C)"
apply (rule closedin_Union)
apply (simp add: \<open>finite F\<close>)
using F_def \<open>locally connected U\<close> clo closedin_Un_complement_component by blast
@@ -1467,7 +1472,7 @@
and him: "h ` (T - \<xi> ` K) \<subseteq> rel_frontier U"
and hg: "\<And>x. x \<in> S \<Longrightarrow> h x = g x"
proof (rule extend_map_affine_to_sphere1 [OF \<open>finite K\<close> \<open>affine T\<close> contg gim, of S "\<xi> ` K"])
- show cloTS: "closedin (subtopology euclidean T) S"
+ show cloTS: "closedin (top_of_set T) S"
by (simp add: \<open>compact S\<close> \<open>S \<subseteq> T\<close> closed_subset compact_imp_closed)
show "\<And>C. \<lbrakk>C \<in> components (T - S); C \<inter> K \<noteq> {}\<rbrakk> \<Longrightarrow> C \<inter> \<xi> ` K \<noteq> {}"
using \<xi> components_eq by blast
@@ -2008,8 +2013,8 @@
fixes f :: "'a \<Rightarrow> 'a::euclidean_space"
assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f ` U \<subseteq> 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 \<subseteq> 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: "\<And>T. open T \<longleftrightarrow> openin (subtopology euclidean (k ` S)) T"
+ have ope_iff: "\<And>T. open T \<longleftrightarrow> 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 \<Rightarrow> 'a::euclidean_space"
assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f ` U \<subseteq> 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' \<equiv> {y. \<forall>x \<in> S. orthogonal x y}"
have "subspace S'"
by (simp add: S'_def subspace_orthogonal_to_vectors)
define g where "g \<equiv> \<lambda>z::'a*'a. ((f \<circ> fst)z, snd z)"
- have "openin (subtopology euclidean (S \<times> S')) (g ` (U \<times> S'))"
+ have "openin (top_of_set (S \<times> S')) (g ` (U \<times> S'))"
proof (rule inv_of_domain_ss0)
show "continuous_on (U \<times> S') g"
apply (simp add: g_def)
@@ -2072,7 +2077,7 @@
using injf by (auto simp: g_def inj_on_def)
show "subspace (S \<times> S')"
by (simp add: \<open>subspace S'\<close> \<open>subspace S\<close> subspace_Times)
- show "openin (subtopology euclidean (S \<times> S')) (U \<times> S')"
+ show "openin (top_of_set (S \<times> S')) (U \<times> S')"
by (simp add: openin_Times [OF ope])
have "dim (S \<times> S') = dim S + dim S'"
by (simp add: \<open>subspace S'\<close> \<open>subspace S\<close> dim_Times)
@@ -2092,11 +2097,11 @@
corollary invariance_of_domain_subspaces:
fixes f :: "'a::euclidean_space \<Rightarrow> '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 \<le> dim U"
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> 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' \<subseteq> 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 \<circ> f) ` S \<subseteq> V'"
using fim homeomorphism_image1 homhk by fastforce
- moreover have "openin (subtopology euclidean U) ((h \<circ> f) ` S)"
+ moreover have "openin (top_of_set U) ((h \<circ> f) ` S)"
proof (rule inv_of_domain_ss1)
show "continuous_on S (h \<circ> f)"
by (meson contf continuous_on_compose continuous_on_subset fim homeomorphism_cont1 homhk)
@@ -2129,14 +2134,14 @@
show "(h \<circ> f) ` S \<subseteq> U"
using \<open>V' \<subseteq> U\<close> hfV' by auto
qed (auto simp: assms)
- ultimately show "openin (subtopology euclidean V') ((h \<circ> f) ` S)"
+ ultimately show "openin (top_of_set V') ((h \<circ> f) ` S)"
using openin_subset_trans \<open>V' \<subseteq> U\<close> by force
qed
qed
corollary invariance_of_dimension_subspaces:
fixes f :: "'a::euclidean_space \<Rightarrow> '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 \<subseteq> V"
and injf: "inj_on f S" and "S \<noteq> {}"
@@ -2158,7 +2163,7 @@
moreover have "inj_on (h \<circ> 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 \<circ> f) ` S)"
+ ultimately have ope_hf: "openin (top_of_set U) ((h \<circ> f) ` S)"
using invariance_of_domain_subspaces [OF ope \<open>subspace U\<close> \<open>subspace U\<close>] by blast
have "(h \<circ> f) ` S \<subseteq> T"
using fim homeomorphism_image1 homhk by fastforce
@@ -2176,11 +2181,11 @@
corollary invariance_of_domain_affine_sets:
fixes f :: "'a::euclidean_space \<Rightarrow> '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 \<le> aff_dim U"
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> 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 \<in> S" "a \<in> U" "b \<in> V"
using False fim ope openin_contains_cball by fastforce
- have "openin (subtopology euclidean ((+) (- b) ` V)) (((+) (- b) \<circ> f \<circ> (+) a) ` (+) (- a) ` S)"
+ have "openin (top_of_set ((+) (- b) ` V)) (((+) (- b) \<circ> f \<circ> (+) 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: \<open>a \<in> U\<close> affine_diffs_subspace_subtract \<open>affine U\<close> cong: image_cong_simp)
@@ -2211,7 +2216,7 @@
corollary invariance_of_dimension_affine_sets:
fixes f :: "'a::euclidean_space \<Rightarrow> '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 \<subseteq> V"
and injf: "inj_on f S" and "S \<noteq> {}"
@@ -2221,7 +2226,7 @@
using \<open>S \<noteq> {}\<close> fim ope openin_contains_cball by fastforce
have "dim ((+) (- a) ` U) \<le> 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: \<open>a \<in> U\<close> affine_diffs_subspace_subtract \<open>affine U\<close> cong: image_cong_simp)
@@ -2269,7 +2274,7 @@
case False
have "aff_dim (affine hull S) \<le> 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 \<inter> g -` T = f ` (S \<inter> T)"
by (auto simp: gf)
- show "openin (subtopology euclidean (f ` S)) (f ` S \<inter> g -` T)"
+ show "openin (top_of_set (f ` S)) (f ` S \<inter> 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 \<subseteq> 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) \<le> 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 \<subseteq> 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 \<le> aff_dim V"
by (simp add: affTU affV)
- have "openin (subtopology euclidean T) ((f \<circ> h) ` g ` (S - {b}))"
+ have "openin (top_of_set T) ((f \<circ> h) ` g ` (S - {b}))"
proof (rule invariance_of_domain_affine_sets [OF _ \<open>affine V\<close>])
- 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 \<circ> 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 \<circ> k) ` j ` (S - {c}))"
+ have "openin (top_of_set T) ((f \<circ> k) ` j ` (S - {c}))"
proof (rule invariance_of_domain_affine_sets [OF _ \<open>affine V\<close>])
- 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 \<circ> k)"
@@ -2890,7 +2895,7 @@
show "(f \<circ> k) ` j ` (S - {c}) \<subseteq> T"
by (metis fim image_comp image_mono kjsub subset_trans)
qed (auto simp: assms)
- ultimately have "openin (subtopology euclidean T) ((f \<circ> h) ` g ` (S - {b}) \<union> ((f \<circ> k) ` j ` (S - {c})))"
+ ultimately have "openin (top_of_set T) ((f \<circ> h) ` g ` (S - {b}) \<union> ((f \<circ> k) ` j ` (S - {c})))"
by (rule openin_Un)
moreover have "(f \<circ> h) ` g ` (S - {b}) = f ` (S - {b})"
proof -
@@ -2931,8 +2936,8 @@
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \<subseteq> T"
and "r \<noteq> 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 \<open>affine T\<close>])
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: \<open>r \<noteq> 0\<close> 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 "\<exists>T. z \<in> T \<and> openin (subtopology euclidean (- {0})) T \<and>
+ show "\<exists>T. z \<in> T \<and> openin (top_of_set (- {0})) T \<and>
(\<exists>v. \<Union>v = exp -` T \<and> (\<forall>u\<in>v. open u) \<and> disjoint v \<and>
(\<forall>u\<in>v. \<exists>q. homeomorphism u T exp q))"
if "z \<in> - {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 "\<Union>\<V> = exp -` exp ` ball (Ln z) 1"
by (force simp: \<V>_def Complex_Transcendental.exp_eq image_iff)
@@ -3647,150 +3652,150 @@
text\<open>Many similar proofs below.\<close>
lemma upper_hemicontinuous:
assumes "\<And>x. x \<in> S \<Longrightarrow> f x \<subseteq> T"
- shows "((\<forall>U. openin (subtopology euclidean T) U
- \<longrightarrow> openin (subtopology euclidean S) {x \<in> S. f x \<subseteq> U}) \<longleftrightarrow>
- (\<forall>U. closedin (subtopology euclidean T) U
- \<longrightarrow> closedin (subtopology euclidean S) {x \<in> S. f x \<inter> U \<noteq> {}}))"
+ shows "((\<forall>U. openin (top_of_set T) U
+ \<longrightarrow> openin (top_of_set S) {x \<in> S. f x \<subseteq> U}) \<longleftrightarrow>
+ (\<forall>U. closedin (top_of_set T) U
+ \<longrightarrow> closedin (top_of_set S) {x \<in> S. f x \<inter> U \<noteq> {}}))"
(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 \<in> S. f x \<subseteq> T - U}"
+ then have "openin (top_of_set S) {x \<in> S. f x \<subseteq> T - U}"
using * [of "T-U"] by blast
moreover have "S - {x \<in> S. f x \<subseteq> T - U} = {x \<in> S. f x \<inter> U \<noteq> {}}"
using assms by blast
- ultimately show "closedin (subtopology euclidean S) {x \<in> S. f x \<inter> U \<noteq> {}}"
+ ultimately show "closedin (top_of_set S) {x \<in> S. f x \<inter> U \<noteq> {}}"
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 \<in> S. f x \<inter> (T - U) \<noteq> {}}"
+ then have "closedin (top_of_set S) {x \<in> S. f x \<inter> (T - U) \<noteq> {}}"
using * [of "T-U"] by blast
moreover have "{x \<in> S. f x \<inter> (T - U) \<noteq> {}} = S - {x \<in> S. f x \<subseteq> U}"
using assms by auto
- ultimately show "openin (subtopology euclidean S) {x \<in> S. f x \<subseteq> U}"
+ ultimately show "openin (top_of_set S) {x \<in> S. f x \<subseteq> U}"
by (simp add: openin_closedin_eq)
qed
lemma lower_hemicontinuous:
assumes "\<And>x. x \<in> S \<Longrightarrow> f x \<subseteq> T"
- shows "((\<forall>U. closedin (subtopology euclidean T) U
- \<longrightarrow> closedin (subtopology euclidean S) {x \<in> S. f x \<subseteq> U}) \<longleftrightarrow>
- (\<forall>U. openin (subtopology euclidean T) U
- \<longrightarrow> openin (subtopology euclidean S) {x \<in> S. f x \<inter> U \<noteq> {}}))"
+ shows "((\<forall>U. closedin (top_of_set T) U
+ \<longrightarrow> closedin (top_of_set S) {x \<in> S. f x \<subseteq> U}) \<longleftrightarrow>
+ (\<forall>U. openin (top_of_set T) U
+ \<longrightarrow> openin (top_of_set S) {x \<in> S. f x \<inter> U \<noteq> {}}))"
(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 \<in> S. f x \<subseteq> T-U}"
+ then have "closedin (top_of_set S) {x \<in> S. f x \<subseteq> T-U}"
using * [of "T-U"] by blast
moreover have "{x \<in> S. f x \<subseteq> T-U} = S - {x \<in> S. f x \<inter> U \<noteq> {}}"
using assms by auto
- ultimately show "openin (subtopology euclidean S) {x \<in> S. f x \<inter> U \<noteq> {}}"
+ ultimately show "openin (top_of_set S) {x \<in> S. f x \<inter> U \<noteq> {}}"
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 \<in> S. f x \<inter> (T - U) \<noteq> {}}"
+ then have "openin (top_of_set S) {x \<in> S. f x \<inter> (T - U) \<noteq> {}}"
using * [of "T-U"] by blast
moreover have "S - {x \<in> S. f x \<inter> (T - U) \<noteq> {}} = {x \<in> S. f x \<subseteq> U}"
using assms by blast
- ultimately show "closedin (subtopology euclidean S) {x \<in> S. f x \<subseteq> U}"
+ ultimately show "closedin (top_of_set S) {x \<in> S. f x \<subseteq> U}"
by (simp add: openin_closedin_eq)
qed
lemma open_map_iff_lower_hemicontinuous_preimage:
assumes "f ` S \<subseteq> T"
- shows "((\<forall>U. openin (subtopology euclidean S) U
- \<longrightarrow> openin (subtopology euclidean T) (f ` U)) \<longleftrightarrow>
- (\<forall>U. closedin (subtopology euclidean S) U
- \<longrightarrow> closedin (subtopology euclidean T) {y \<in> T. {x. x \<in> S \<and> f x = y} \<subseteq> U}))"
+ shows "((\<forall>U. openin (top_of_set S) U
+ \<longrightarrow> openin (top_of_set T) (f ` U)) \<longleftrightarrow>
+ (\<forall>U. closedin (top_of_set S) U
+ \<longrightarrow> closedin (top_of_set T) {y \<in> T. {x. x \<in> S \<and> f x = y} \<subseteq> 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 \<in> T. {x \<in> S. f x = y} \<subseteq> U}"
using assms by blast
- ultimately show "closedin (subtopology euclidean T) {y \<in> T. {x \<in> S. f x = y} \<subseteq> U}"
+ ultimately show "closedin (top_of_set T) {y \<in> T. {x \<in> S. f x = y} \<subseteq> 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 \<in> T. {x \<in> S. f x = y} \<subseteq> S - U}"
+ then have "closedin (top_of_set T) {y \<in> T. {x \<in> S. f x = y} \<subseteq> S - U}"
using * [of "S-U"] by blast
moreover have "{y \<in> T. {x \<in> S. f x = y} \<subseteq> 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 \<subseteq> T"
- shows "((\<forall>U. closedin (subtopology euclidean S) U
- \<longrightarrow> closedin (subtopology euclidean T) (f ` U)) \<longleftrightarrow>
- (\<forall>U. openin (subtopology euclidean S) U
- \<longrightarrow> openin (subtopology euclidean T) {y \<in> T. {x. x \<in> S \<and> f x = y} \<subseteq> U}))"
+ shows "((\<forall>U. closedin (top_of_set S) U
+ \<longrightarrow> closedin (top_of_set T) (f ` U)) \<longleftrightarrow>
+ (\<forall>U. openin (top_of_set S) U
+ \<longrightarrow> openin (top_of_set T) {y \<in> T. {x. x \<in> S \<and> f x = y} \<subseteq> 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 \<in> T. {x. x \<in> S \<and> f x = y} \<subseteq> U}"
using assms openin_imp_subset [OF opeSU] by auto
- ultimately show "openin (subtopology euclidean T) {y \<in> T. {x. x \<in> S \<and> f x = y} \<subseteq> U}"
+ ultimately show "openin (top_of_set T) {y \<in> T. {x. x \<in> S \<and> f x = y} \<subseteq> 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 \<in> T. {x \<in> S. f x = y} \<subseteq> S - U}"
+ then have "openin (top_of_set T) {y \<in> T. {x \<in> S. f x = y} \<subseteq> S - U}"
using * [of "S-U"] by blast
moreover have "(f ` U) = T - {y \<in> T. {x \<in> S. f x = y} \<subseteq> 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: "\<And>x. x \<in> S \<Longrightarrow> f x \<subseteq> T"
- and ope: "\<And>U. openin (subtopology euclidean T) U
- \<Longrightarrow> openin (subtopology euclidean S) {x \<in> S. f x \<subseteq> U}"
- and clo: "\<And>U. closedin (subtopology euclidean T) U
- \<Longrightarrow> closedin (subtopology euclidean S) {x \<in> S. f x \<subseteq> U}"
+ and ope: "\<And>U. openin (top_of_set T) U
+ \<Longrightarrow> openin (top_of_set S) {x \<in> S. f x \<subseteq> U}"
+ and clo: "\<And>U. closedin (top_of_set T) U
+ \<Longrightarrow> closedin (top_of_set S) {x \<in> S. f x \<subseteq> U}"
and "x \<in> S" "0 < e" and bofx: "bounded(f x)" and fx_ne: "f x \<noteq> {}"
obtains d where "0 < d"
"\<And>x'. \<lbrakk>x' \<in> S; dist x x' < d\<rbrakk>
\<Longrightarrow> (\<forall>y \<in> f x. \<exists>y'. y' \<in> f x' \<and> dist y y' < e) \<and>
(\<forall>y' \<in> f x'. \<exists>y. y \<in> f x \<and> dist y' y < e)"
proof -
- have "openin (subtopology euclidean T) (T \<inter> (\<Union>a\<in>f x. \<Union>b\<in>ball 0 e. {a + b}))"
+ have "openin (top_of_set T) (T \<inter> (\<Union>a\<in>f x. \<Union>b\<in>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 \<in> S. f u \<subseteq> T \<inter> (\<Union>a\<in>f x. \<Union>b\<in>ball 0 e. {a + b})}" by blast
with \<open>0 < e\<close> \<open>x \<in> S\<close> obtain d1 where "d1 > 0" and
d1: "\<And>x'. \<lbrakk>x' \<in> S; dist x' x < d1\<rbrakk> \<Longrightarrow> f x' \<subseteq> T \<and> f x' \<subseteq> (\<Union>a \<in> f x. \<Union>b \<in> ball 0 e. {a + b})"
by (force simp: openin_euclidean_subtopology_iff dest: fST)
- have oo: "\<And>U. openin (subtopology euclidean T) U \<Longrightarrow>
- openin (subtopology euclidean S) {x \<in> S. f x \<inter> U \<noteq> {}}"
+ have oo: "\<And>U. openin (top_of_set T) U \<Longrightarrow>
+ openin (top_of_set S) {x \<in> S. f x \<inter> U \<noteq> {}}"
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 \<in> (\<Inter>a \<in> C. {x \<in> S. f x \<inter> T \<inter> ball a (e/2) \<noteq> {}})"
using \<open>x \<in> S\<close> \<open>0 < e\<close> fST \<open>C \<subseteq> f x\<close> by force
- have "openin (subtopology euclidean S) {x \<in> S. f x \<inter> (T \<inter> ball a (e/2)) \<noteq> {}}" for a
+ have "openin (top_of_set S) {x \<in> S. f x \<inter> (T \<inter> ball a (e/2)) \<noteq> {}}" for a
by (simp add: openin_open_Int oo)
- then have "openin (subtopology euclidean S) (\<Inter>a \<in> C. {x \<in> S. f x \<inter> T \<inter> ball a (e/2) \<noteq> {}})"
+ then have "openin (top_of_set S) (\<Inter>a \<in> C. {x \<in> S. f x \<inter> T \<inter> ball a (e/2) \<noteq> {}})"
by (simp add: Int_assoc openin_INT2 [OF \<open>finite C\<close> \<open>C \<noteq> {}\<close>])
with xin obtain d2 where "d2>0"
and d2: "\<And>u v. \<lbrakk>u \<in> S; dist u x < d2; v \<in> C\<rbrakk> \<Longrightarrow> f u \<inter> T \<inter> ball v (e/2) \<noteq> {}"
@@ -4375,8 +4380,8 @@
proposition Borsukian_open_Un:
fixes S :: "'a::real_normed_vector set"
- assumes opeS: "openin (subtopology euclidean (S \<union> T)) S"
- and opeT: "openin (subtopology euclidean (S \<union> T)) T"
+ assumes opeS: "openin (top_of_set (S \<union> T)) S"
+ and opeT: "openin (top_of_set (S \<union> T)) T"
and BS: "Borsukian S" and BT: "Borsukian T" and ST: "connected(S \<inter> T)"
shows "Borsukian(S \<union> T)"
proof (clarsimp simp add: Borsukian_continuous_logarithm)
@@ -4445,8 +4450,8 @@
text\<open>The proof is a duplicate of that of \<open>Borsukian_open_Un\<close>.\<close>
lemma Borsukian_closed_Un:
fixes S :: "'a::real_normed_vector set"
- assumes cloS: "closedin (subtopology euclidean (S \<union> T)) S"
- and cloT: "closedin (subtopology euclidean (S \<union> T)) T"
+ assumes cloS: "closedin (top_of_set (S \<union> T)) S"
+ and cloT: "closedin (top_of_set (S \<union> T)) T"
and BS: "Borsukian S" and BT: "Borsukian T" and ST: "connected(S \<inter> T)"
shows "Borsukian(S \<union> T)"
proof (clarsimp simp add: Borsukian_continuous_logarithm)
@@ -4586,8 +4591,8 @@
lemma Borsukian_open_map_image_compact:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "Borsukian S" and contf: "continuous_on S f" and fim: "f ` S = T" and "compact S"
- and ope: "\<And>U. openin (subtopology euclidean S) U
- \<Longrightarrow> openin (subtopology euclidean T) (f ` U)"
+ and ope: "\<And>U. openin (top_of_set S) U
+ \<Longrightarrow> openin (top_of_set T) (f ` U)"
shows "Borsukian T"
proof (clarsimp simp add: Borsukian_continuous_logarithm_circle_real)
fix g :: "'b \<Rightarrow> complex"
@@ -4636,12 +4641,12 @@
\<Longrightarrow> (\<forall>v \<in> {z \<in> S. f z = y}. \<exists>v'. v' \<in> {z \<in> S. f z = x'} \<and> dist v v' < d) \<and>
(\<forall>v' \<in> {z \<in> S. f z = x'}. \<exists>v. v \<in> {z \<in> S. f z = y} \<and> dist v' v < d)"
proof (rule upper_lower_hemicontinuous_explicit [of T "\<lambda>y. {z \<in> S. f z = y}" S])
- show "\<And>U. openin (subtopology euclidean S) U
- \<Longrightarrow> openin (subtopology euclidean T) {x \<in> T. {z \<in> S. f z = x} \<subseteq> U}"
+ show "\<And>U. openin (top_of_set S) U
+ \<Longrightarrow> openin (top_of_set T) {x \<in> T. {z \<in> S. f z = x} \<subseteq> U}"
using continuous_imp_closed_map closed_map_iff_upper_hemicontinuous_preimage [OF fim [THEN equalityD1]]
by (simp add: continuous_imp_closed_map \<open>compact S\<close> contf fim)
- show "\<And>U. closedin (subtopology euclidean S) U \<Longrightarrow>
- closedin (subtopology euclidean T) {x \<in> T. {z \<in> S. f z = x} \<subseteq> U}"
+ show "\<And>U. closedin (top_of_set S) U \<Longrightarrow>
+ closedin (top_of_set T) {x \<in> T. {z \<in> S. f z = x} \<subseteq> U}"
using ope open_map_iff_lower_hemicontinuous_preimage [OF fim [THEN equalityD1]]
by meson
show "bounded {z \<in> S. f z = y}"
@@ -4812,17 +4817,17 @@
definition%important unicoherent where
"unicoherent U \<equiv>
\<forall>S T. connected S \<and> connected T \<and> S \<union> T = U \<and>
- closedin (subtopology euclidean U) S \<and> closedin (subtopology euclidean U) T
+ closedin (top_of_set U) S \<and> closedin (top_of_set U) T
\<longrightarrow> connected (S \<inter> T)"
lemma unicoherentI [intro?]:
- assumes "\<And>S T. \<lbrakk>connected S; connected T; U = S \<union> T; closedin (subtopology euclidean U) S; closedin (subtopology euclidean U) T\<rbrakk>
+ assumes "\<And>S T. \<lbrakk>connected S; connected T; U = S \<union> T; closedin (top_of_set U) S; closedin (top_of_set U) T\<rbrakk>
\<Longrightarrow> connected (S \<inter> T)"
shows "unicoherent U"
using assms unfolding unicoherent_def by blast
lemma unicoherentD:
- assumes "unicoherent U" "connected S" "connected T" "U = S \<union> T" "closedin (subtopology euclidean U) S" "closedin (subtopology euclidean U) T"
+ assumes "unicoherent U" "connected S" "connected T" "U = S \<union> T" "closedin (top_of_set U) S" "closedin (top_of_set U) T"
shows "connected (S \<inter> 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 \<union> 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 \<inter> g ` V) \<subseteq> U" "f ` (g ` U \<inter> g ` V) \<subseteq> V"
using gf fim T by auto (metis UnCI image_iff)+
moreover have "U \<inter> V \<subseteq> f ` (g ` U \<inter> 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 \<union> T"
- and cloS: "closedin (subtopology euclidean (S \<union> T)) S"
- and cloT: "closedin (subtopology euclidean (S \<union> T)) T"
+ and cloS: "closedin (top_of_set (S \<union> T)) S"
+ and cloT: "closedin (top_of_set (S \<union> T)) T"
show "connected (S \<inter> T)"
unfolding connected_closedin_eq
proof clarify
fix V W
- assume "closedin (subtopology euclidean (S \<inter> T)) V"
- and "closedin (subtopology euclidean (S \<inter> T)) W"
+ assume "closedin (top_of_set (S \<inter> T)) V"
+ and "closedin (top_of_set (S \<inter> T)) W"
and VW: "V \<union> W = S \<inter> T" "V \<inter> W = {}" and "V \<noteq> {}" "W \<noteq> {}"
- 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 \<open>U = S \<union> T\<close> cloS cloT closedin_trans by blast+
obtain q where contq: "continuous_on U q"
and q01: "\<And>x. x \<in> U \<Longrightarrow> q x \<in> {0..1::real}"
@@ -5020,8 +5025,8 @@
proof
fix U V
assume UV: "connected U" "connected V" "T = U \<union> 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 \<open>compact S\<close> 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 \<open>unicoherent S\<close>, of "S \<inter> f -` U" "S \<inter> f -` V"])
- have "\<And>C. closedin (subtopology euclidean S) C \<Longrightarrow> closedin (subtopology euclidean T) (f ` C)"
+ have "\<And>C. closedin (top_of_set S) C \<Longrightarrow> closedin (top_of_set T) (f ` C)"
by (metis \<open>compact S\<close> closed_subset closedin_compact closedin_imp_subset compact_continuous_image compact_imp_closed contf continuous_on_subset fim image_mono)
then show "connected (S \<inter> f -` U)" "connected (S \<inter> f -` V)"
using UV by (auto simp: conn intro: connected_closed_monotone_preimage [OF contf fim])
show "S = (S \<inter> f -` U) \<union> (S \<inter> f -` V)"
using UV fim by blast
- show "closedin (subtopology euclidean S) (S \<inter> f -` U)"
- "closedin (subtopology euclidean S) (S \<inter> f -` V)"
+ show "closedin (top_of_set S) (S \<inter> f -` U)"
+ "closedin (top_of_set S) (S \<inter> 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 \<subseteq> -{0}"
and gf: "\<And>x. x \<in> S \<Longrightarrow> 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 (\<lambda>t. a) \<subseteq> - {0}"
using a homotopic_with_imp_subset2 False by blast
@@ -5396,8 +5401,8 @@
lemma inessential_on_clopen_Union:
fixes \<F> :: "'a::euclidean_space set set"
assumes T: "path_connected T"
- and "\<And>S. S \<in> \<F> \<Longrightarrow> closedin (subtopology euclidean (\<Union>\<F>)) S"
- and "\<And>S. S \<in> \<F> \<Longrightarrow> openin (subtopology euclidean (\<Union>\<F>)) S"
+ and "\<And>S. S \<in> \<F> \<Longrightarrow> closedin (top_of_set (\<Union>\<F>)) S"
+ and "\<And>S. S \<in> \<F> \<Longrightarrow> openin (top_of_set (\<Union>\<F>)) S"
and hom: "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>a. homotopic_with (\<lambda>x. True) S T f (\<lambda>x. a)"
obtains a where "homotopic_with (\<lambda>x. True) (\<Union>\<F>) T f (\<lambda>x. a)"
proof (cases "\<Union>\<F> = {}")
@@ -5408,16 +5413,16 @@
case False
then obtain C where "C \<in> \<F>" "C \<noteq> {}"
by blast
- then obtain a where clo: "closedin (subtopology euclidean (\<Union>\<F>)) C"
- and ope: "openin (subtopology euclidean (\<Union>\<F>)) C"
+ then obtain a where clo: "closedin (top_of_set (\<Union>\<F>)) C"
+ and ope: "openin (top_of_set (\<Union>\<F>)) C"
and "homotopic_with (\<lambda>x. True) C T f (\<lambda>x. a)"
using assms by blast
with \<open>C \<noteq> {}\<close> have "f ` C \<subseteq> T" "a \<in> T"
using homotopic_with_imp_subset1 homotopic_with_imp_subset2 by blast+
have "homotopic_with (\<lambda>x. True) (\<Union>\<F>) T f (\<lambda>x. a)"
proof (rule homotopic_on_clopen_Union)
- show "\<And>S. S \<in> \<F> \<Longrightarrow> closedin (subtopology euclidean (\<Union>\<F>)) S"
- "\<And>S. S \<in> \<F> \<Longrightarrow> openin (subtopology euclidean (\<Union>\<F>)) S"
+ show "\<And>S. S \<in> \<F> \<Longrightarrow> closedin (top_of_set (\<Union>\<F>)) S"
+ "\<And>S. S \<in> \<F> \<Longrightarrow> openin (top_of_set (\<Union>\<F>)) S"
by (simp_all add: assms)
show "homotopic_with (\<lambda>x. True) S T f (\<lambda>x. a)" if "S \<in> \<F>" for S
proof (cases "S = {}")
--- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Tue Mar 19 16:14:59 2019 +0000
@@ -4415,9 +4415,9 @@
then show "\<exists>e>0. ball x e \<subseteq> (S \<inter> f -` {f x})"
using \<open>0 < e\<close> by blast
qed
- then have "openin (subtopology euclidean S) (S \<inter> f -` {y})"
+ then have "openin (top_of_set S) (S \<inter> f -` {y})"
by (auto intro!: open_openin_trans[OF \<open>open S\<close>] simp: open_contains_ball)
- moreover have "closedin (subtopology euclidean S) (S \<inter> f -` {y})"
+ moreover have "closedin (top_of_set S) (S \<inter> f -` {y})"
by (force intro!: continuous_closedin_preimage [OF contf])
ultimately have "(S \<inter> f -` {y}) = {} \<or> (S \<inter> f -` {y}) = S"
using \<open>connected S\<close> by (simp add: connected_clopen)
--- a/src/HOL/Analysis/Homeomorphism.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Homeomorphism.thy Tue Mar 19 16:14:59 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 \<noteq> {}" "closedin (subtopology euclidean U) T"
+ where "convex U" "U \<noteq> {}" "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 ((\<lambda>x. setdist {x} (- U)) o fst) z" for z :: "'a \<times> 'b"
by (rule continuous_at_compose continuous_intros continuous_at_setdist)+
@@ -1275,9 +1275,9 @@
where
"covering_space c p S \<equiv>
continuous_on c p \<and> p ` c = S \<and>
- (\<forall>x \<in> S. \<exists>T. x \<in> T \<and> openin (subtopology euclidean S) T \<and>
+ (\<forall>x \<in> S. \<exists>T. x \<in> T \<and> openin (top_of_set S) T \<and>
(\<exists>v. \<Union>v = c \<inter> p -` T \<and>
- (\<forall>u \<in> v. openin (subtopology euclidean c) u) \<and>
+ (\<forall>u \<in> v. openin (top_of_set c) u) \<and>
pairwise disjnt v \<and>
(\<forall>u \<in> v. \<exists>q. homeomorphism u T p q)))"
@@ -1295,8 +1295,8 @@
lemma covering_space_local_homeomorphism:
assumes "covering_space c p S" "x \<in> c"
- obtains T u q where "x \<in> T" "openin (subtopology euclidean c) T"
- "p x \<in> u" "openin (subtopology euclidean S) u"
+ obtains T u q where "x \<in> T" "openin (top_of_set c) T"
+ "p x \<in> 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 \<in> S"
obtains x T U q where "p x = y"
- "x \<in> T" "openin (subtopology euclidean c) T"
- "y \<in> U" "openin (subtopology euclidean S) U"
+ "x \<in> T" "openin (top_of_set c) T"
+ "y \<in> U" "openin (top_of_set S) U"
"homeomorphism T U p q"
proof -
obtain x where "p x = y" "x \<in> 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:
"\<And>x. x \<in> S \<Longrightarrow>
- \<exists>X VS. x \<in> X \<and> openin (subtopology euclidean S) X \<and>
+ \<exists>X VS. x \<in> X \<and> openin (top_of_set S) X \<and>
\<Union>VS = c \<inter> p -` X \<and>
- (\<forall>u \<in> VS. openin (subtopology euclidean c) u) \<and>
+ (\<forall>u \<in> VS. openin (top_of_set c) u) \<and>
pairwise disjnt VS \<and>
(\<forall>u \<in> VS. \<exists>q. homeomorphism u X p q)"
using p by (auto simp: covering_space_def)
have "T \<subseteq> c" by (metis openin_euclidean_subtopology_iff T)
- have "\<exists>X. openin (subtopology euclidean S) X \<and> y \<in> X \<and> X \<subseteq> p ` T"
+ have "\<exists>X. openin (top_of_set S) X \<and> y \<in> X \<and> X \<subseteq> p ` T"
if "y \<in> p ` T" for y
proof -
have "y \<in> S" using \<open>T \<subseteq> c\<close> pce that by blast
- obtain U VS where "y \<in> U" and U: "openin (subtopology euclidean S) U"
+ obtain U VS where "y \<in> U" and U: "openin (top_of_set S) U"
and VS: "\<Union>VS = c \<inter> p -` U"
- and openVS: "\<forall>V \<in> VS. openin (subtopology euclidean c) V"
+ and openVS: "\<forall>V \<in> VS. openin (top_of_set c) V"
and homVS: "\<And>V. V \<in> VS \<Longrightarrow> \<exists>q. homeomorphism V U p q"
using covs [OF \<open>y \<in> S\<close>] by auto
obtain x where "x \<in> c" "p x \<in> U" "x \<in> 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 \<inter> V) = U \<inter> q -` (T \<inter> V)"
using VS \<open>V \<in> VS\<close> by (auto simp: homeomorphism_def)
- have ocv: "openin (subtopology euclidean c) V"
+ have ocv: "openin (top_of_set c) V"
by (simp add: \<open>V \<in> VS\<close> openVS)
- have "openin (subtopology euclidean U) (U \<inter> q -` (T \<inter> V))"
+ have "openin (top_of_set U) (U \<inter> q -` (T \<inter> 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 \<inter> q -` (T \<inter> V))"
+ then have os: "openin (top_of_set S) (U \<inter> q -` (T \<inter> V))"
using openin_trans [of U] by (simp add: Collect_conj_eq U)
show ?thesis
apply (rule_tac x = "p ` (T \<inter> 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 \<open>U \<subseteq> T\<close> 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 \<in> U" "g1 z - g2 z = 0"
- obtain v w q where "g1 z \<in> v" and ocv: "openin (subtopology euclidean c) v"
- and "p (g1 z) \<in> w" and osw: "openin (subtopology euclidean S) w"
+ obtain v w q where "g1 z \<in> v" and ocv: "openin (top_of_set c) v"
+ and "p (g1 z) \<in> 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 \<open>U \<subseteq> T\<close> \<open>z \<in> U\<close> g1(2) apply blast+
@@ -1400,15 +1400,15 @@
have "g2 z \<in> v" using \<open>g1 z \<in> v\<close> z by auto
have gg: "U \<inter> g -` v = U \<inter> g -` (v \<inter> g ` U)" for g
by auto
- have "openin (subtopology euclidean (g1 ` U)) (v \<inter> g1 ` U)"
+ have "openin (top_of_set (g1 ` U)) (v \<inter> g1 ` U)"
using ocv \<open>U \<subseteq> T\<close> g1 by (fastforce simp add: openin_open)
- then have 1: "openin (subtopology euclidean U) (U \<inter> g1 -` v)"
+ then have 1: "openin (top_of_set U) (U \<inter> g1 -` v)"
unfolding gg by (blast intro: contu continuous_on_open [THEN iffD1, rule_format])
- have "openin (subtopology euclidean (g2 ` U)) (v \<inter> g2 ` U)"
+ have "openin (top_of_set (g2 ` U)) (v \<inter> g2 ` U)"
using ocv \<open>U \<subseteq> T\<close> g2 by (fastforce simp add: openin_open)
- then have 2: "openin (subtopology euclidean U) (U \<inter> g2 -` v)"
+ then have 2: "openin (top_of_set U) (U \<inter> g2 -` v)"
unfolding gg by (blast intro: contu continuous_on_open [THEN iffD1, rule_format])
- show "\<exists>T. openin (subtopology euclidean U) T \<and> z \<in> T \<and> T \<subseteq> {z \<in> U. g1 z - g2 z = 0}"
+ show "\<exists>T. openin (top_of_set U) T \<and> z \<in> T \<and> T \<subseteq> {z \<in> U. g1 z - g2 z = 0}"
using z
apply (rule_tac x = "(U \<inter> g1 -` v) \<inter> (U \<inter> g2 -` v)" in exI)
apply (intro conjI)
@@ -1417,7 +1417,7 @@
apply (metis \<open>U \<subseteq> T\<close> 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 = {} \<or> G12 = U"
@@ -1468,37 +1468,37 @@
show ?rhs
proof (rule locallyI)
fix V x
- assume V: "openin (subtopology euclidean C) V" and "x \<in> V"
+ assume V: "openin (top_of_set C) V" and "x \<in> V"
have "p x \<in> p ` C"
by (metis IntE V \<open>x \<in> V\<close> imageI openin_open)
then obtain T \<V> where "p x \<in> T"
- and opeT: "openin (subtopology euclidean S) T"
+ and opeT: "openin (top_of_set S) T"
and veq: "\<Union>\<V> = C \<inter> p -` T"
- and ope: "\<forall>U\<in>\<V>. openin (subtopology euclidean C) U"
+ and ope: "\<forall>U\<in>\<V>. openin (top_of_set C) U"
and hom: "\<forall>U\<in>\<V>. \<exists>q. homeomorphism U T p q"
using cov unfolding covering_space_def by (blast intro: that)
have "x \<in> \<Union>\<V>"
using V veq \<open>p x \<in> T\<close> \<open>x \<in> V\<close> openin_imp_subset by fastforce
then obtain U where "x \<in> U" "U \<in> \<V>"
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 \<inter> V)"
+ with V have "openin (top_of_set C) (U \<inter> V)"
by blast
- then have UV: "openin (subtopology euclidean S) (p ` (U \<inter> V))"
+ then have UV: "openin (top_of_set S) (p ` (U \<inter> V))"
using cov covering_space_open_map by blast
- obtain W W' where opeW: "openin (subtopology euclidean S) W" and "\<psi> W'" "p x \<in> W" "W \<subseteq> W'" and W'sub: "W' \<subseteq> p ` (U \<inter> V)"
+ obtain W W' where opeW: "openin (top_of_set S) W" and "\<psi> W'" "p x \<in> W" "W \<subseteq> W'" and W'sub: "W' \<subseteq> p ` (U \<inter> V)"
using locallyE [OF L UV] \<open>x \<in> U\<close> \<open>x \<in> V\<close> by blast
then have "W \<subseteq> T"
by (metis Int_lower1 q homeomorphism_image1 image_Int_subset order_trans)
- show "\<exists>U Z. openin (subtopology euclidean C) U \<and>
+ show "\<exists>U Z. openin (top_of_set C) U \<and>
\<phi> Z \<and> x \<in> U \<and> U \<subseteq> Z \<and> Z \<subseteq> 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 \<open>W \<subseteq> T\<close>)
- 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 "\<phi> (q ` W')"
by (metis (mono_tags, lifting) Int_subset_iff UV W'sub \<open>\<psi> W'\<close> continuous_on_subset dual_order.trans homeomorphism_def image_Int_subset openin_imp_subset q qim)
@@ -1575,18 +1575,18 @@
"\<And>y. y \<in> U \<Longrightarrow> k(0, y) = f y"
"\<And>z. z \<in> {0..1} \<times> U \<Longrightarrow> h z = p(k z)"
proof -
- have "\<exists>V k. openin (subtopology euclidean U) V \<and> y \<in> V \<and>
+ have "\<exists>V k. openin (top_of_set U) V \<and> y \<in> V \<and>
continuous_on ({0..1} \<times> V) k \<and> k ` ({0..1} \<times> V) \<subseteq> C \<and>
(\<forall>z \<in> V. k(0, z) = f z) \<and> (\<forall>z \<in> {0..1} \<times> V. h z = p(k z))"
if "y \<in> U" for y
proof -
- obtain UU where UU: "\<And>s. s \<in> S \<Longrightarrow> s \<in> (UU s) \<and> openin (subtopology euclidean S) (UU s) \<and>
+ obtain UU where UU: "\<And>s. s \<in> S \<Longrightarrow> s \<in> (UU s) \<and> openin (top_of_set S) (UU s) \<and>
(\<exists>\<V>. \<Union>\<V> = C \<inter> p -` UU s \<and>
- (\<forall>U \<in> \<V>. openin (subtopology euclidean C) U) \<and>
+ (\<forall>U \<in> \<V>. openin (top_of_set C) U) \<and>
pairwise disjnt \<V> \<and>
(\<forall>U \<in> \<V>. \<exists>q. homeomorphism U (UU s) p q))"
using cov unfolding covering_space_def by (metis (mono_tags))
- then have ope: "\<And>s. s \<in> S \<Longrightarrow> s \<in> (UU s) \<and> openin (subtopology euclidean S) (UU s)"
+ then have ope: "\<And>s. s \<in> S \<Longrightarrow> s \<in> (UU s) \<and> openin (top_of_set S) (UU s)"
by blast
have "\<exists>k n i. open k \<and> open n \<and>
t \<in> k \<and> y \<in> n \<and> i \<in> S \<and> h ` (({0..1} \<inter> k) \<times> (U \<inter> n)) \<subseteq> UU i" if "t \<in> {0..1}" for t
@@ -1595,7 +1595,7 @@
using \<open>y \<in> U\<close> him that by blast
then have "(t,y) \<in> ({0..1} \<times> U) \<inter> h -` UU(h(t, y))"
using \<open>y \<in> U\<close> \<open>t \<in> {0..1}\<close> by (auto simp: ope)
- moreover have ope_01U: "openin (subtopology euclidean ({0..1} \<times> U)) (({0..1} \<times> U) \<inter> h -` UU(h(t, y)))"
+ moreover have ope_01U: "openin (top_of_set ({0..1} \<times> U)) (({0..1} \<times> U) \<inter> 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 \<in> {0..1} \<inter> V" "t \<in> {0..1} \<inter> V"
and opeW: "open W" and "y \<in> U" "y \<in> W"
@@ -1641,7 +1641,7 @@
using reals_Archimedean2 by blast
then have "N > 0"
using \<open>0 < \<delta>\<close> order.asym by force
- have *: "\<exists>V k. openin (subtopology euclidean U) V \<and> y \<in> V \<and>
+ have *: "\<exists>V k. openin (top_of_set U) V \<and> y \<in> V \<and>
continuous_on ({0..of_nat n / N} \<times> V) k \<and>
k ` ({0..of_nat n / N} \<times> V) \<subseteq> C \<and>
(\<forall>z\<in>V. k (0, z) = f z) \<and>
@@ -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 \<in> V"
and contk: "continuous_on ({0..real n / real N} \<times> V) k"
and kim: "k ` ({0..real n / real N} \<times> V) \<subseteq> C"
@@ -1676,7 +1676,7 @@
have t01: "t \<in> {0..1}"
using \<open>t \<in> tk\<close> \<open>tk \<subseteq> {0..1}\<close> by blast
obtain \<V> where \<V>: "\<Union>\<V> = C \<inter> p -` UU (X t)"
- and opeC: "\<And>U. U \<in> \<V> \<Longrightarrow> openin (subtopology euclidean C) U"
+ and opeC: "\<And>U. U \<in> \<V> \<Longrightarrow> openin (top_of_set C) U"
and "pairwise disjnt \<V>"
and homuu: "\<And>U. U \<in> \<V> \<Longrightarrow> \<exists>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) \<in> W" and "W \<in> \<V>"
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 \<subseteq> C"
using openin_imp_subset by blast
define W' where "W' = UU(X t)"
- have opeVW: "openin (subtopology euclidean V) (V \<inter> (k \<circ> Pair (n / N)) -` W)"
+ have opeVW: "openin (top_of_set V) (V \<inter> (k \<circ> 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: \<open>y \<in> V\<close> W)
done
- obtain N' where opeUN': "openin (subtopology euclidean U) N'"
+ obtain N' where opeUN': "openin (top_of_set U) N'"
and "y \<in> N'" and kimw: "k ` ({(real n / real N)} \<times> N') \<subseteq> W"
apply (rule_tac N' = "(V \<inter> (k \<circ> Pair (n / N)) -` W)" in that)
apply (fastforce simp: \<open>y \<in> V\<close> 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 \<in> Q" "Q \<subseteq> Q'"
and Q': "Q' \<subseteq> (U \<inter> NN(t)) \<inter> N' \<inter> V"
proof -
@@ -1732,9 +1732,9 @@
by (metis Int_iff \<open>N' = U \<inter> VX\<close> \<open>V = U \<inter> VO\<close> \<open>y \<in> N'\<close> \<open>y \<in> V\<close> inUS open_contains_cball t01)
show ?thesis
proof
- show "openin (subtopology euclidean U) (U \<inter> ball y e)"
+ show "openin (top_of_set U) (U \<inter> ball y e)"
by blast
- show "closedin (subtopology euclidean U) (U \<inter> cball y e)"
+ show "closedin (top_of_set U) (U \<inter> cball y e)"
using e by (auto simp: closedin_closed)
qed (use \<open>y \<in> U\<close> \<open>e > 0\<close> VO VX e in auto)
qed
@@ -1749,13 +1749,13 @@
(\<lambda>x. if x \<in> {0..real n / real N} \<times> Q' then k x else (p' \<circ> 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} \<times> Q'))
+ show "closedin (top_of_set ({0..(1 + real n) / real N} \<times> Q'))
({0..real n / real N} \<times> Q')"
apply (simp add: closedin_closed)
apply (rule_tac x="{0 .. real n / real N} \<times> UNIV" in exI)
using n_div_N_in apply (auto simp: closed_Times)
done
- show "closedin (subtopology euclidean ({0..(1 + real n) / real N} \<times> Q'))
+ show "closedin (top_of_set ({0..(1 + real n) / real N} \<times> Q'))
({real n / real N..(1 + real n) / real N} \<times> Q')"
apply (simp add: closedin_closed)
apply (rule_tac x="{real n / real N .. (1 + real n) / real N} \<times> UNIV" in exI)
@@ -1845,7 +1845,7 @@
show ?thesis
using*[OF order_refl] N \<open>0 < \<delta>\<close> by (simp add: split: if_split_asm)
qed
- then obtain V fs where opeV: "\<And>y. y \<in> U \<Longrightarrow> openin (subtopology euclidean U) (V y)"
+ then obtain V fs where opeV: "\<And>y. y \<in> U \<Longrightarrow> openin (top_of_set U) (V y)"
and V: "\<And>y. y \<in> U \<Longrightarrow> y \<in> V y"
and contfs: "\<And>y. y \<in> U \<Longrightarrow> continuous_on ({0..1} \<times> V y) (fs y)"
and *: "\<And>y. y \<in> U \<Longrightarrow> (fs y) ` ({0..1} \<times> V y) \<subseteq> C \<and>
@@ -1857,14 +1857,17 @@
obtain k where contk: "continuous_on ({0..1} \<times> U) k"
and k: "\<And>x i. \<lbrakk>i \<in> U; x \<in> {0..1} \<times> U \<inter> {0..1} \<times> V i\<rbrakk> \<Longrightarrow> k x = fs i x"
proof (rule pasting_lemma_exists)
- show "{0..1} \<times> U \<subseteq> (\<Union>i\<in>U. {0..1} \<times> V i)"
- apply auto
- using V by blast
- show "\<And>i. i \<in> U \<Longrightarrow> openin (subtopology euclidean ({0..1} \<times> U)) ({0..1} \<times> V i)"
+ let ?X = "top_of_set ({0..1::real} \<times> U)"
+ show "topspace ?X \<subseteq> (\<Union>i\<in>U. {0..1} \<times> V i)"
+ using V by force
+ show "\<And>i. i \<in> U \<Longrightarrow> openin (top_of_set ({0..1} \<times> U)) ({0..1} \<times> V i)"
by (simp add: opeV openin_Times)
- show "\<And>i. i \<in> U \<Longrightarrow> continuous_on ({0..1} \<times> V i) (fs i)"
- using contfs by blast
- show "fs i x = fs j x" if "i \<in> U" "j \<in> U" and x: "x \<in> {0..1} \<times> U \<inter> {0..1} \<times> V i \<inter> {0..1} \<times> V j"
+ show "\<And>i. i \<in> U \<Longrightarrow> continuous_map
+ (subtopology (top_of_set ({0..1} \<times> U)) ({0..1} \<times> 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 \<in> U" "j \<in> U" and x: "x \<in> topspace ?X \<inter> {0..1} \<times> V i \<inter> {0..1} \<times> V j"
for i j x
proof -
obtain u y where "x = (u, y)" "y \<in> V i" "y \<in> V j" "0 \<le> u" "u \<le> 1"
@@ -1900,7 +1903,7 @@
using \<open>x = (u, y)\<close> x by blast
qed
qed
- qed blast
+ qed force
show ?thesis
proof
show "k ` ({0..1} \<times> U) \<subseteq> C"
@@ -2282,17 +2285,17 @@
using l [of "linepath z z" z "linepath a a"] by (auto simp: assms)
show LC: "l ` U \<subseteq> C"
by (clarify dest!: *) (metis (full_types) l pathfinish_in_path_image subsetCE)
- have "\<exists>T. openin (subtopology euclidean U) T \<and> y \<in> T \<and> T \<subseteq> U \<inter> l -` X"
- if X: "openin (subtopology euclidean C) X" and "y \<in> U" "l y \<in> X" for X y
+ have "\<exists>T. openin (top_of_set U) T \<and> y \<in> T \<and> T \<subseteq> U \<inter> l -` X"
+ if X: "openin (top_of_set C) X" and "y \<in> U" "l y \<in> X" for X y
proof -
have "X \<subseteq> C"
using X openin_euclidean_subtopology_iff by blast
have "f y \<in> S"
using fim \<open>y \<in> U\<close> by blast
then obtain W \<V>
- where WV: "f y \<in> W \<and> openin (subtopology euclidean S) W \<and>
+ where WV: "f y \<in> W \<and> openin (top_of_set S) W \<and>
(\<Union>\<V> = C \<inter> p -` W \<and>
- (\<forall>U \<in> \<V>. openin (subtopology euclidean C) U) \<and>
+ (\<forall>U \<in> \<V>. openin (top_of_set C) U) \<and>
pairwise disjnt \<V> \<and>
(\<forall>U \<in> \<V>. \<exists>q. homeomorphism U W p q))"
using cov by (force simp: covering_space_def)
@@ -2300,15 +2303,15 @@
using \<open>X \<subseteq> C\<close> pleq that by auto
then obtain W' where "l y \<in> W'" and "W' \<in> \<V>"
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 \<subseteq> W'"
using homUW' homeomorphism_image2 homeomorphism_cont2 by fastforce+
obtain V where "y \<in> V" "y \<in> U" and fimW: "f ` V \<subseteq> W" "V \<subseteq> 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 \<inter> f -` W)"
+ have "openin (top_of_set U) (U \<inter> 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 \<inter> p' -` (W' \<inter> X))"
+ have "openin (top_of_set S) (W \<inter> p' -` (W' \<inter> X))"
proof (rule openin_trans)
- show "openin (subtopology euclidean W) (W \<inter> p' -` (W' \<inter> X))"
+ show "openin (top_of_set W) (W \<inter> p' -` (W' \<inter> X))"
apply (rule continuous_openin_preimage [OF contp' p'im])
using X \<open>W' \<subseteq> C\<close> 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 \<inter> (U \<inter> (f -` (W \<inter> (p' -` (W' \<inter> X))))))"
+ then show "openin (top_of_set U) (V \<inter> (U \<inter> (f -` (W \<inter> (p' -` (W' \<inter> X))))))"
by (blast intro: opeUV openin_subtopology_self continuous_openin_preimage [OF contf fim])
have "p' (f y) \<in> X"
using \<open>l y \<in> W'\<close> homeomorphism_apply1 [OF homUW'] pleq \<open>y \<in> U\<close> \<open>l y \<in> X\<close> by fastforce
--- a/src/HOL/Analysis/Homotopy.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Homotopy.thy Tue Mar 19 16:14:59 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} \<times> X \<union> {1/2..1} \<times> X)) ({0..1/2::real} \<times> X)"
+ have clo1: "closedin (top_of_set ({0..1/2} \<times> X \<union> {1/2..1} \<times> X)) ({0..1/2::real} \<times> X)"
apply (simp add: closedin_closed split_01_prod [symmetric])
apply (rule_tac x="{0..1/2} \<times> UNIV" in exI)
apply (force simp: closed_Times)
done
- have clo2: "closedin (subtopology euclidean ({0..1/2} \<times> X \<union> {1/2..1} \<times> X)) ({1/2..1::real} \<times> X)"
+ have clo2: "closedin (top_of_set ({0..1/2} \<times> X \<union> {1/2..1} \<times> X)) ({1/2..1::real} \<times> X)"
apply (simp add: closedin_closed split_01_prod [symmetric])
apply (rule_tac x="{1/2..1} \<times> UNIV" in exI)
apply (force simp: closed_Times)
@@ -1472,20 +1472,20 @@
definition%important locally :: "('a::topological_space set \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
where
"locally P S \<equiv>
- \<forall>w x. openin (subtopology euclidean S) w \<and> x \<in> w
- \<longrightarrow> (\<exists>u v. openin (subtopology euclidean S) u \<and> P v \<and>
+ \<forall>w x. openin (top_of_set S) w \<and> x \<in> w
+ \<longrightarrow> (\<exists>u v. openin (top_of_set S) u \<and> P v \<and>
x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> w)"
lemma locallyI:
- assumes "\<And>w x. \<lbrakk>openin (subtopology euclidean S) w; x \<in> w\<rbrakk>
- \<Longrightarrow> \<exists>u v. openin (subtopology euclidean S) u \<and> P v \<and>
+ assumes "\<And>w x. \<lbrakk>openin (top_of_set S) w; x \<in> w\<rbrakk>
+ \<Longrightarrow> \<exists>u v. openin (top_of_set S) u \<and> P v \<and>
x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> w"
shows "locally P S"
using assms by (force simp: locally_def)
lemma locallyE:
- assumes "locally P S" "openin (subtopology euclidean S) w" "x \<in> w"
- obtains u v where "openin (subtopology euclidean S) u"
+ assumes "locally P S" "openin (top_of_set S) w" "x \<in> w"
+ obtains u v where "openin (top_of_set S) u"
"P v" "x \<in> u" "u \<subseteq> v" "v \<subseteq> 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:
- "\<lbrakk>locally P S; closedin (subtopology euclidean S) t\<rbrakk> \<Longrightarrow> locally P (S - t)"
+ "\<lbrakk>locally P S; closedin (top_of_set S) t\<rbrakk> \<Longrightarrow> 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 \<times> T)) W" and xy: "(x, y) \<in> W"
- then obtain U V where "openin (subtopology euclidean S) U" "x \<in> U"
- "openin (subtopology euclidean T) V" "y \<in> V" "U \<times> V \<subseteq> W"
+ assume W: "openin (top_of_set (S \<times> T)) W" and xy: "(x, y) \<in> W"
+ then obtain U V where "openin (top_of_set S) U" "x \<in> U"
+ "openin (top_of_set T) V" "y \<in> V" "U \<times> V \<subseteq> W"
using Times_in_interior_subtopology by metis
then obtain U1 U2 V1 V2
- where opeS: "openin (subtopology euclidean S) U1 \<and> P U2 \<and> x \<in> U1 \<and> U1 \<subseteq> U2 \<and> U2 \<subseteq> U"
- and opeT: "openin (subtopology euclidean T) V1 \<and> Q V2 \<and> y \<in> V1 \<and> V1 \<subseteq> V2 \<and> V2 \<subseteq> V"
+ where opeS: "openin (top_of_set S) U1 \<and> P U2 \<and> x \<in> U1 \<and> U1 \<subseteq> U2 \<and> U2 \<subseteq> U"
+ and opeT: "openin (top_of_set T) V1 \<and> Q V2 \<and> y \<in> V1 \<and> V1 \<subseteq> V2 \<and> V2 \<subseteq> V"
by (meson PS QT locallyE)
- with \<open>U \<times> V \<subseteq> W\<close> show "\<exists>u v. openin (subtopology euclidean (S \<times> T)) u \<and> R v \<and> (x,y) \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> W"
+ with \<open>U \<times> V \<subseteq> W\<close> show "\<exists>u v. openin (top_of_set (S \<times> T)) u \<and> R v \<and> (x,y) \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> W"
apply (rule_tac x="U1 \<times> V1" in exI)
apply (rule_tac x="U2 \<times> 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 \<in> W" and "openin (subtopology euclidean t) W"
+ assume "y \<in> W" and "openin (top_of_set t) W"
then obtain T where T: "open T" "W = t \<inter> T"
by (force simp: openin_open)
then have "W \<subseteq> t" by auto
@@ -1586,15 +1586,15 @@
using \<open>g ` t = S\<close> \<open>W \<subseteq> t\<close> apply blast
using g \<open>W \<subseteq> t\<close> apply auto[1]
by (simp add: f rev_image_eqI)
- have \<circ>: "openin (subtopology euclidean S) (g ` W)"
+ have \<circ>: "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 \<open>openin (subtopology euclidean t) W\<close> continuous_on_open f(2))
+ then show "openin (top_of_set S) (g ` W)"
+ by (simp add: gw Collect_conj_eq \<open>openin (top_of_set t) W\<close> continuous_on_open f(2))
qed
then obtain u v
- where osu: "openin (subtopology euclidean S) u" and uv: "P v" "g y \<in> u" "u \<subseteq> v" "v \<subseteq> g ` W"
+ where osu: "openin (top_of_set S) u" and uv: "P v" "g y \<in> u" "u \<subseteq> v" "v \<subseteq> g ` W"
using S [unfolded locally_def, rule_format, of "g ` W" "g y"] \<open>y \<in> W\<close> by force
have "v \<subseteq> S" using uv by (simp add: gw)
have fv: "f ` v = t \<inter> {x. g x \<in> v}"
@@ -1609,7 +1609,7 @@
using \<open>v \<subseteq> S\<close> \<open>W \<subseteq> t\<close> 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 \<inter> g -` u)"
+ have 1: "openin (top_of_set t) (t \<inter> g -` u)"
apply (rule continuous_on_open [THEN iffD1, rule_format])
apply (rule \<open>continuous_on t g\<close>)
using \<open>g ` t = S\<close> apply (simp add: osu)
@@ -1619,7 +1619,7 @@
apply (intro conjI Q [OF \<open>P v\<close> homv])
using \<open>W \<subseteq> t\<close> \<open>y \<in> W\<close> \<open>f ` v \<subseteq> W\<close> uv apply (auto simp: fv)
done
- show "\<exists>U. openin (subtopology euclidean t) U \<and> (\<exists>v. Q v \<and> y \<in> U \<and> U \<subseteq> v \<and> v \<subseteq> W)"
+ show "\<exists>U. openin (top_of_set t) U \<and> (\<exists>v. Q v \<and> y \<in> U \<and> U \<subseteq> v \<and> v \<subseteq> W)"
by (meson 1 2)
qed
@@ -1674,23 +1674,23 @@
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
assumes P: "locally P S"
and f: "continuous_on S f"
- and oo: "\<And>t. openin (subtopology euclidean S) t
- \<Longrightarrow> openin (subtopology euclidean (f ` S)) (f ` t)"
+ and oo: "\<And>t. openin (top_of_set S) t
+ \<Longrightarrow> openin (top_of_set (f ` S)) (f ` t)"
and Q: "\<And>t. \<lbrakk>t \<subseteq> S; P t\<rbrakk> \<Longrightarrow> 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 \<in> W"
+ assume oiw: "openin (top_of_set (f ` S)) W" and "y \<in> W"
then have "W \<subseteq> f ` S" by (simp add: openin_euclidean_subtopology_iff)
- have oivf: "openin (subtopology euclidean S) (S \<inter> f -` W)"
+ have oivf: "openin (top_of_set S) (S \<inter> f -` W)"
by (rule continuous_on_open [THEN iffD1, rule_format, OF f oiw])
then obtain x where "x \<in> S" "f x = y"
using \<open>W \<subseteq> f ` S\<close> \<open>y \<in> W\<close> by blast
then obtain U V
- where "openin (subtopology euclidean S) U" "P V" "x \<in> U" "U \<subseteq> V" "V \<subseteq> S \<inter> f -` W"
+ where "openin (top_of_set S) U" "P V" "x \<in> U" "U \<subseteq> V" "V \<subseteq> S \<inter> f -` W"
using P [unfolded locally_def, rule_format, of "(S \<inter> f -` W)" x] oivf \<open>y \<in> W\<close>
by auto
- then show "\<exists>X. openin (subtopology euclidean (f ` S)) X \<and> (\<exists>Y. Q Y \<and> y \<in> X \<and> X \<subseteq> Y \<and> Y \<subseteq> W)"
+ then show "\<exists>X. openin (top_of_set (f ` S)) X \<and> (\<exists>Y. Q Y \<and> y \<in> X \<and> X \<subseteq> Y \<and> Y \<subseteq> 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: "\<And>T a. \<lbrakk>openin (subtopology euclidean S) T; a \<in> T\<rbrakk> \<Longrightarrow> \<exists>z. z \<in> T \<and> P z"
+ and opD: "\<And>T a. \<lbrakk>openin (top_of_set S) T; a \<in> T\<rbrakk> \<Longrightarrow> \<exists>z. z \<in> T \<and> P z"
and opI: "\<And>a. a \<in> S
- \<Longrightarrow> \<exists>T. openin (subtopology euclidean S) T \<and> a \<in> T \<and>
+ \<Longrightarrow> \<exists>T. openin (top_of_set S) T \<and> a \<in> T \<and>
(\<forall>x \<in> T. \<forall>y \<in> T. P x \<and> P y \<and> Q x \<longrightarrow> Q y)"
and etc: "a \<in> S" "b \<in> S" "P a" "P b" "Q a"
shows "Q b"
proof -
- have 1: "openin (subtopology euclidean S)
- {b. \<exists>T. openin (subtopology euclidean S) T \<and>
+ have 1: "openin (top_of_set S)
+ {b. \<exists>T. openin (top_of_set S) T \<and>
b \<in> T \<and> (\<forall>x\<in>T. P x \<longrightarrow> Q x)}"
apply (subst openin_subopen, clarify)
apply (rule_tac x=T in exI, auto)
done
- have 2: "openin (subtopology euclidean S)
- {b. \<exists>T. openin (subtopology euclidean S) T \<and>
+ have 2: "openin (top_of_set S)
+ {b. \<exists>T. openin (top_of_set S) T \<and>
b \<in> T \<and> (\<forall>x\<in>T. P x \<longrightarrow> \<not> 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 \<in> S" "b \<in> S" "P a" "P b"
and trans: "\<And>x y z. \<lbrakk>R x y; R y z\<rbrakk> \<Longrightarrow> R x z"
- and opD: "\<And>T a. \<lbrakk>openin (subtopology euclidean S) T; a \<in> T\<rbrakk> \<Longrightarrow> \<exists>z. z \<in> T \<and> P z"
+ and opD: "\<And>T a. \<lbrakk>openin (top_of_set S) T; a \<in> T\<rbrakk> \<Longrightarrow> \<exists>z. z \<in> T \<and> P z"
and opI: "\<And>a. a \<in> S
- \<Longrightarrow> \<exists>T. openin (subtopology euclidean S) T \<and> a \<in> T \<and>
+ \<Longrightarrow> \<exists>T. openin (top_of_set S) T \<and> a \<in> T \<and>
(\<forall>x \<in> T. \<forall>y \<in> T. P x \<and> P y \<longrightarrow> R x y)"
shows "R a b"
proof -
@@ -1754,7 +1754,7 @@
assumes "connected S"
and etc: "a \<in> S" "b \<in> S" "P a"
and opI: "\<And>a. a \<in> S
- \<Longrightarrow> \<exists>T. openin (subtopology euclidean S) T \<and> a \<in> T \<and>
+ \<Longrightarrow> \<exists>T. openin (top_of_set S) T \<and> a \<in> T \<and>
(\<forall>x \<in> T. \<forall>y \<in> T. P x \<longrightarrow> P y)"
shows "P b"
apply (rule connected_induction [OF \<open>connected S\<close> _, where P = "\<lambda>x. True"], blast)
@@ -1767,7 +1767,7 @@
and etc: "a \<in> S" "b \<in> S"
and sym: "\<And>x y. \<lbrakk>R x y; x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> R y x"
and trans: "\<And>x y z. \<lbrakk>R x y; R y z; x \<in> S; y \<in> S; z \<in> S\<rbrakk> \<Longrightarrow> R x z"
- and opI: "\<And>a. a \<in> S \<Longrightarrow> \<exists>T. openin (subtopology euclidean S) T \<and> a \<in> T \<and> (\<forall>x \<in> T. R a x)"
+ and opI: "\<And>a. a \<in> S \<Longrightarrow> \<exists>T. openin (top_of_set S) T \<and> a \<in> T \<and> (\<forall>x \<in> T. R a x)"
shows "R a b"
proof -
have "\<And>a b c. \<lbrakk>a \<in> S; b \<in> S; c \<in> S; R a b\<rbrakk> \<Longrightarrow> R a c"
@@ -1779,7 +1779,7 @@
lemma locally_constant_imp_constant:
assumes "connected S"
and opI: "\<And>a. a \<in> S
- \<Longrightarrow> \<exists>T. openin (subtopology euclidean S) T \<and> a \<in> T \<and> (\<forall>x \<in> T. f x = f a)"
+ \<Longrightarrow> \<exists>T. openin (top_of_set S) T \<and> a \<in> T \<and> (\<forall>x \<in> T. f x = f a)"
shows "f constant_on S"
proof -
have "\<And>x y. x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> f x = f y"
@@ -1805,7 +1805,7 @@
shows
"locally compact s \<longleftrightarrow>
(\<forall>x \<in> s. \<exists>u v. x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<and>
- openin (subtopology euclidean s) u \<and> compact v)"
+ openin (top_of_set s) u \<and> compact v)"
(is "?lhs = ?rhs")
proof
assume ?lhs
@@ -1816,11 +1816,11 @@
next
assume r [rule_format]: ?rhs
have *: "\<exists>u v.
- openin (subtopology euclidean s) u \<and>
+ openin (top_of_set s) u \<and>
compact v \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<inter> T"
if "open T" "x \<in> s" "x \<in> T" for x T
proof -
- obtain u v where uv: "x \<in> u" "u \<subseteq> v" "v \<subseteq> s" "compact v" "openin (subtopology euclidean s) u"
+ obtain u v where uv: "x \<in> u" "u \<subseteq> v" "v \<subseteq> s" "compact v" "openin (top_of_set s) u"
using r [OF \<open>x \<in> s\<close>] by auto
obtain e where "e>0" and e: "cball x e \<subseteq> T"
using open_contains_cball \<open>open T\<close> \<open>x \<in> T\<close> by blast
@@ -1842,7 +1842,7 @@
fixes s :: "'a :: metric_space set"
assumes "locally compact s"
obtains u v where "\<And>x. x \<in> s \<Longrightarrow> x \<in> u x \<and> u x \<subseteq> v x \<and> v x \<subseteq> s \<and>
- openin (subtopology euclidean s) (u x) \<and> compact (v x)"
+ openin (top_of_set s) (u x) \<and> compact (v x)"
using assms
unfolding locally_compact by metis
@@ -1850,7 +1850,7 @@
fixes s :: "'a :: heine_borel set"
shows "locally compact s \<longleftrightarrow>
(\<forall>x \<in> s. \<exists>u. x \<in> u \<and>
- openin (subtopology euclidean s) u \<and> compact(closure u) \<and> closure u \<subseteq> s)"
+ openin (top_of_set s) u \<and> compact(closure u) \<and> closure u \<subseteq> 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 \<longleftrightarrow>
(\<forall>k. k \<subseteq> s \<and> compact k
\<longrightarrow> (\<exists>u v. k \<subseteq> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<and>
- openin (subtopology euclidean s) u \<and> compact v))"
+ openin (top_of_set s) u \<and> compact v))"
(is "?lhs = ?rhs")
proof
assume ?lhs
then obtain u v where
uv: "\<And>x. x \<in> s \<Longrightarrow> x \<in> u x \<and> u x \<subseteq> v x \<and> v x \<subseteq> s \<and>
- openin (subtopology euclidean s) (u x) \<and> compact (v x)"
+ openin (top_of_set s) (u x) \<and> compact (v x)"
by (metis locally_compactE)
- have *: "\<exists>u v. k \<subseteq> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<and> openin (subtopology euclidean s) u \<and> compact v"
+ have *: "\<exists>u v. k \<subseteq> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<and> openin (top_of_set s) u \<and> compact v"
if "k \<subseteq> s" "compact k" for k
proof -
- have "\<And>C. (\<forall>c\<in>C. openin (subtopology euclidean k) c) \<and> k \<subseteq> \<Union>C \<Longrightarrow>
+ have "\<And>C. (\<forall>c\<in>C. openin (top_of_set k) c) \<and> k \<subseteq> \<Union>C \<Longrightarrow>
\<exists>D\<subseteq>C. finite D \<and> k \<subseteq> \<Union>D"
using that by (simp add: compact_eq_openin_cover)
- moreover have "\<forall>c \<in> (\<lambda>x. k \<inter> u x) ` k. openin (subtopology euclidean k) c"
+ moreover have "\<forall>c \<in> (\<lambda>x. k \<inter> 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 \<subseteq> \<Union>((\<lambda>x. k \<inter> u x) ` k)"
using that by clarsimp (meson subsetCE uv)
@@ -1931,12 +1931,12 @@
assumes "open s"
shows "locally compact s"
proof -
- have *: "\<exists>u v. x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<and> openin (subtopology euclidean s) u \<and> compact v"
+ have *: "\<exists>u v. x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<and> openin (top_of_set s) u \<and> compact v"
if "x \<in> s" for x
proof -
obtain e where "e>0" and e: "cball x e \<subseteq> s"
using open_contains_cball assms \<open>x \<in> s\<close> 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 *: "\<exists>u v. x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<and>
- openin (subtopology euclidean s) u \<and> compact v"
+ openin (top_of_set s) u \<and> compact v"
if "x \<in> s" for x
proof -
show ?thesis
@@ -1979,7 +1979,7 @@
lemma locally_compact_closedin:
fixes s :: "'a :: heine_borel set"
- shows "\<lbrakk>closedin (subtopology euclidean s) t; locally compact s\<rbrakk>
+ shows "\<lbrakk>closedin (top_of_set s) t; locally compact s\<rbrakk>
\<Longrightarrow> 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 \<union> T)) S"
- and opT: "openin (subtopology euclidean (S \<union> T)) T"
+ and opS: "openin (top_of_set (S \<union> T)) S"
+ and opT: "openin (top_of_set (S \<union> T)) T"
shows "locally compact (S \<union> T)"
proof -
have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if "x \<in> 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 \<union> T)) S"
- and clT: "closedin (subtopology euclidean (S \<union> T)) T"
+ and clS: "closedin (top_of_set (S \<union> T)) S"
+ and clT: "closedin (top_of_set (S \<union> T)) T"
shows "locally compact (S \<union> T)"
proof -
have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if "x \<in> S" "x \<in> T" for x
@@ -2113,7 +2113,7 @@
"locally compact S \<longleftrightarrow>
(\<forall>K T. K \<subseteq> S \<and> compact K \<and> open T \<and> K \<subseteq> T
\<longrightarrow> (\<exists>U V. K \<subseteq> U \<and> U \<subseteq> V \<and> U \<subseteq> T \<and> V \<subseteq> S \<and>
- openin (subtopology euclidean S) U \<and> compact V))"
+ openin (top_of_set S) U \<and> compact V))"
(is "?lhs = ?rhs")
proof
assume L: ?lhs
@@ -2122,10 +2122,10 @@
fix K :: "'a set" and T :: "'a set"
assume "K \<subseteq> S" and "compact K" and "open T" and "K \<subseteq> T"
obtain U V where "K \<subseteq> U" "U \<subseteq> V" "V \<subseteq> 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 \<open>K \<subseteq> S\<close> \<open>compact K\<close>)
show "\<exists>U V. K \<subseteq> U \<and> U \<subseteq> V \<and> U \<subseteq> T \<and> V \<subseteq> S \<and>
- openin (subtopology euclidean S) U \<and> compact V"
+ openin (top_of_set S) U \<and> compact V"
proof (intro exI conjI)
show "K \<subseteq> U \<inter> T"
by (simp add: \<open>K \<subseteq> T\<close> \<open>K \<subseteq> U\<close>)
@@ -2133,7 +2133,7 @@
by (rule closure_subset)
show "closure (U \<inter> T) \<subseteq> S"
by (metis \<open>U \<subseteq> V\<close> \<open>V \<subseteq> S\<close> \<open>compact V\<close> closure_closed closure_mono compact_imp_closed inf.cobounded1 subset_trans)
- show "openin (subtopology euclidean S) (U \<inter> T)"
+ show "openin (top_of_set S) (U \<inter> T)"
by (simp add: \<open>open T\<close> ope openin_Int_open)
show "compact (closure (U \<inter> T))"
by (meson Int_lower1 \<open>U \<subseteq> V\<close> \<open>compact V\<close> 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 \<in> components S"
- shows "C = \<Inter>{T. C \<subseteq> T \<and> openin (subtopology euclidean S) T \<and>
- closedin (subtopology euclidean S) T}"
+ shows "C = \<Inter>{T. C \<subseteq> T \<and> openin (top_of_set S) T \<and>
+ closedin (top_of_set S) T}"
(is "C = \<Inter>?\<T>")
proof
obtain x where x: "C = connected_component_set S x" and "x \<in> S"
@@ -2168,8 +2168,8 @@
have clo: "closed (\<Inter>?\<T>)"
by (simp add: \<open>compact S\<close> closed_Inter closedin_compact_eq compact_imp_closed)
have False
- if K1: "closedin (subtopology euclidean (\<Inter>?\<T>)) K1" and
- K2: "closedin (subtopology euclidean (\<Inter>?\<T>)) K2" and
+ if K1: "closedin (top_of_set (\<Inter>?\<T>)) K1" and
+ K2: "closedin (top_of_set (\<Inter>?\<T>)) K2" and
K12_Int: "K1 \<inter> K2 = {}" and K12_Un: "K1 \<union> K2 = \<Inter>?\<T>" and "K1 \<noteq> {}" "K2 \<noteq> {}"
for K1 K2
proof -
@@ -2187,8 +2187,8 @@
show "(S - (V1 \<union> V2)) \<inter> \<Inter>\<F> \<noteq> {}" if "finite \<F>" and \<F>: "\<F> \<subseteq> ?\<T>" for \<F>
proof
assume djo: "(S - (V1 \<union> V2)) \<inter> \<Inter>\<F> = {}"
- 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 \<subseteq> D" and DV12: "D \<subseteq> V1 \<union> V2"
proof (cases "\<F> = {}")
case True
@@ -2197,9 +2197,9 @@
next
case False show ?thesis
proof
- show ope: "openin (subtopology euclidean S) (\<Inter>\<F>)"
+ show ope: "openin (top_of_set S) (\<Inter>\<F>)"
using openin_Inter \<open>finite \<F>\<close> False \<F> by blast
- then show "closedin (subtopology euclidean S) (\<Inter>\<F>)"
+ then show "closedin (top_of_set S) (\<Inter>\<F>)"
by (meson clo\<T> \<F> closed_Inter closed_subset openin_imp_subset subset_eq)
show "C \<subseteq> \<Inter>\<F>"
using \<F> by auto
@@ -2211,16 +2211,16 @@
by (simp add: x)
have "closed D"
using \<open>compact S\<close> cloD closedin_closed_trans compact_imp_closed by blast
- have cloV1: "closedin (subtopology euclidean D) (D \<inter> closure V1)"
- and cloV2: "closedin (subtopology euclidean D) (D \<inter> closure V2)"
+ have cloV1: "closedin (top_of_set D) (D \<inter> closure V1)"
+ and cloV2: "closedin (top_of_set D) (D \<inter> closure V2)"
by (simp_all add: closedin_closed_Int)
moreover have "D \<inter> closure V1 = D \<inter> V1" "D \<inter> closure V2 = D \<inter> V2"
apply safe
using \<open>D \<subseteq> V1 \<union> V2\<close> \<open>open V1\<close> \<open>open V2\<close> V12
apply (simp_all add: closure_subset [THEN subsetD] closure_iff_nhds_not_empty, blast+)
done
- ultimately have cloDV1: "closedin (subtopology euclidean D) (D \<inter> V1)"
- and cloDV2: "closedin (subtopology euclidean D) (D \<inter> V2)"
+ ultimately have cloDV1: "closedin (top_of_set D) (D \<inter> V1)"
+ and cloDV2: "closedin (top_of_set D) (D \<inter> V2)"
by metis+
then obtain U1 U2 where "closed U1" "closed U2"
and D1: "D \<inter> V1 = D \<inter> U1" and D2: "D \<inter> V2 = D \<inter> U2"
@@ -2274,21 +2274,21 @@
fixes S :: "'a::euclidean_space set"
assumes S: "locally compact S" and C: "C \<in> components S" and "compact C"
and U: "open U" "C \<subseteq> U"
- obtains K where "openin (subtopology euclidean S) K" "compact K" "C \<subseteq> K" "K \<subseteq> U"
+ obtains K where "openin (top_of_set S) K" "compact K" "C \<subseteq> K" "K \<subseteq> U"
proof (rule ccontr)
assume "\<not> thesis"
- with that have neg: "\<nexists>K. openin (subtopology euclidean S) K \<and> compact K \<and> C \<subseteq> K \<and> K \<subseteq> U"
+ with that have neg: "\<nexists>K. openin (top_of_set S) K \<and> compact K \<and> C \<subseteq> K \<and> K \<subseteq> U"
by metis
obtain V K where "C \<subseteq> V" "V \<subseteq> U" "V \<subseteq> K" "K \<subseteq> S" "compact K"
- and opeSV: "openin (subtopology euclidean S) V"
+ and opeSV: "openin (top_of_set S) V"
using S U \<open>compact C\<close>
apply (simp add: locally_compact_compact_subopen)
by (meson C in_components_subset)
- let ?\<T> = "{T. C \<subseteq> T \<and> openin (subtopology euclidean K) T \<and> compact T \<and> T \<subseteq> K}"
+ let ?\<T> = "{T. C \<subseteq> T \<and> openin (top_of_set K) T \<and> compact T \<and> T \<subseteq> K}"
have CK: "C \<in> components K"
by (meson C \<open>C \<subseteq> V\<close> \<open>K \<subseteq> S\<close> \<open>V \<subseteq> K\<close> components_intermediate_subset subset_trans)
with \<open>compact K\<close>
- have "C = \<Inter>{T. C \<subseteq> T \<and> openin (subtopology euclidean K) T \<and> closedin (subtopology euclidean K) T}"
+ have "C = \<Inter>{T. C \<subseteq> T \<and> openin (top_of_set K) T \<and> closedin (top_of_set K) T}"
by (simp add: Sura_Bura_compact)
then have Ceq: "C = \<Inter>?\<T>"
by (simp add: closedin_compact_eq \<open>compact K\<close>)
@@ -2322,11 +2322,11 @@
by (metis (no_types, lifting) compact_Inter False mem_Collect_eq subsetCE \<F>)
moreover have "\<Inter>\<F> \<subseteq> K"
using False that(2) by fastforce
- moreover have opeKF: "openin (subtopology euclidean K) (\<Inter>\<F>)"
+ moreover have opeKF: "openin (top_of_set K) (\<Inter>\<F>)"
using False \<F> \<open>finite \<F>\<close> by blast
- then have opeVF: "openin (subtopology euclidean V) (\<Inter>\<F>)"
+ then have opeVF: "openin (top_of_set V) (\<Inter>\<F>)"
using W \<open>K \<subseteq> S\<close> \<open>V \<subseteq> K\<close> opeKF \<open>\<Inter>\<F> \<subseteq> K\<close> FUW openin_subset_trans by fastforce
- then have "openin (subtopology euclidean S) (\<Inter>\<F>)"
+ then have "openin (top_of_set S) (\<Inter>\<F>)"
by (metis opeSV openin_trans)
moreover have "\<Inter>\<F> \<subseteq> U"
by (meson \<open>V \<subseteq> U\<close> 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 \<in> components S" and "compact C"
- and opeSU: "openin (subtopology euclidean S) U" and "C \<subseteq> U"
- obtains K where "openin (subtopology euclidean S) K" "compact K" "C \<subseteq> K" "K \<subseteq> U"
+ and opeSU: "openin (top_of_set S) U" and "C \<subseteq> U"
+ obtains K where "openin (top_of_set S) K" "compact K" "C \<subseteq> K" "K \<subseteq> U"
proof -
obtain V where "open V" "U = S \<inter> 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 \<in> components S" "compact C"
- shows "C = \<Inter> {K. C \<subseteq> K \<and> compact K \<and> openin (subtopology euclidean S) K}"
+ shows "C = \<Inter> {K. C \<subseteq> K \<and> compact K \<and> openin (top_of_set S) K}"
(is "C = ?rhs")
proof
show "?rhs \<subseteq> C"
proof (clarsimp, rule ccontr)
fix x
- assume *: "\<forall>X. C \<subseteq> X \<and> compact X \<and> openin (subtopology euclidean S) X \<longrightarrow> x \<in> X"
+ assume *: "\<forall>X. C \<subseteq> X \<and> compact X \<and> openin (top_of_set S) X \<longrightarrow> x \<in> X"
and "x \<notin> C"
obtain U V where "open U" "open V" "{x} \<subseteq> U" "C \<subseteq> V" "U \<inter> V = {}"
using separation_normal [of "{x}" C]
@@ -2381,8 +2381,8 @@
lemma locally_connected_1:
assumes
- "\<And>v x. \<lbrakk>openin (subtopology euclidean S) v; x \<in> v\<rbrakk>
- \<Longrightarrow> \<exists>u. openin (subtopology euclidean S) u \<and>
+ "\<And>v x. \<lbrakk>openin (top_of_set S) v; x \<in> v\<rbrakk>
+ \<Longrightarrow> \<exists>u. openin (top_of_set S) u \<and>
connected u \<and> x \<in> u \<and> u \<subseteq> 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 \<in> 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"
"\<forall>w x. openin ?SS w \<and> x \<in> w \<longrightarrow> (\<exists>u. openin ?SS u \<and> (\<exists>v. connected v \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> w))"
and "connected_component t x y"
@@ -2420,29 +2420,29 @@
qed
lemma locally_connected_3:
- assumes "\<And>t x. \<lbrakk>openin (subtopology euclidean S) t; x \<in> t\<rbrakk>
- \<Longrightarrow> openin (subtopology euclidean S)
+ assumes "\<And>t x. \<lbrakk>openin (top_of_set S) t; x \<in> t\<rbrakk>
+ \<Longrightarrow> openin (top_of_set S)
(connected_component_set t x)"
- "openin (subtopology euclidean S) v" "x \<in> v"
- shows "\<exists>u. openin (subtopology euclidean S) u \<and> connected u \<and> x \<in> u \<and> u \<subseteq> v"
+ "openin (top_of_set S) v" "x \<in> v"
+ shows "\<exists>u. openin (top_of_set S) u \<and> connected u \<and> x \<in> u \<and> u \<subseteq> v"
using assms connected_component_subset by fastforce
lemma locally_connected:
"locally connected S \<longleftrightarrow>
- (\<forall>v x. openin (subtopology euclidean S) v \<and> x \<in> v
- \<longrightarrow> (\<exists>u. openin (subtopology euclidean S) u \<and> connected u \<and> x \<in> u \<and> u \<subseteq> v))"
+ (\<forall>v x. openin (top_of_set S) v \<and> x \<in> v
+ \<longrightarrow> (\<exists>u. openin (top_of_set S) u \<and> connected u \<and> x \<in> u \<and> u \<subseteq> v))"
by (metis locally_connected_1 locally_connected_2 locally_connected_3)
lemma locally_connected_open_connected_component:
"locally connected S \<longleftrightarrow>
- (\<forall>t x. openin (subtopology euclidean S) t \<and> x \<in> t
- \<longrightarrow> openin (subtopology euclidean S) (connected_component_set t x))"
+ (\<forall>t x. openin (top_of_set S) t \<and> x \<in> t
+ \<longrightarrow> 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
- "\<And>v x. \<lbrakk>openin (subtopology euclidean S) v; x \<in> v\<rbrakk>
- \<Longrightarrow> \<exists>u. openin (subtopology euclidean S) u \<and> path_connected u \<and> x \<in> u \<and> u \<subseteq> v"
+ "\<And>v x. \<lbrakk>openin (top_of_set S) v; x \<in> v\<rbrakk>
+ \<Longrightarrow> \<exists>u. openin (top_of_set S) u \<and> path_connected u \<and> x \<in> u \<and> u \<subseteq> 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 \<in> 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"
"\<forall>w x. openin ?SS w \<and> x \<in> w \<longrightarrow> (\<exists>u. openin ?SS u \<and> (\<exists>v. path_connected v \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> w))"
and "path_component t x y"
@@ -2479,10 +2479,10 @@
qed
lemma locally_path_connected_3:
- assumes "\<And>t x. \<lbrakk>openin (subtopology euclidean S) t; x \<in> t\<rbrakk>
- \<Longrightarrow> openin (subtopology euclidean S) (path_component_set t x)"
- "openin (subtopology euclidean S) v" "x \<in> v"
- shows "\<exists>u. openin (subtopology euclidean S) u \<and> path_connected u \<and> x \<in> u \<and> u \<subseteq> v"
+ assumes "\<And>t x. \<lbrakk>openin (top_of_set S) t; x \<in> t\<rbrakk>
+ \<Longrightarrow> openin (top_of_set S) (path_component_set t x)"
+ "openin (top_of_set S) v" "x \<in> v"
+ shows "\<exists>u. openin (top_of_set S) u \<and> path_connected u \<and> x \<in> u \<and> u \<subseteq> 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 \<longleftrightarrow>
- (\<forall>v x. openin (subtopology euclidean S) v \<and> x \<in> v
- \<longrightarrow> (\<exists>u. openin (subtopology euclidean S) u \<and> path_connected u \<and> x \<in> u \<and> u \<subseteq> v))"
+ (\<forall>v x. openin (top_of_set S) v \<and> x \<in> v
+ \<longrightarrow> (\<exists>u. openin (top_of_set S) u \<and> path_connected u \<and> x \<in> u \<and> u \<subseteq> 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 \<longleftrightarrow>
- (\<forall>t x. openin (subtopology euclidean S) t \<and> x \<in> t
- \<longrightarrow> openin (subtopology euclidean S) (path_component_set t x))"
+ (\<forall>t x. openin (top_of_set S) t \<and> x \<in> t
+ \<longrightarrow> 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 \<longleftrightarrow>
- (\<forall>t c. openin (subtopology euclidean S) t \<and> c \<in> components t
- \<longrightarrow> openin (subtopology euclidean S) c)"
+ (\<forall>t c. openin (top_of_set S) t \<and> c \<in> components t
+ \<longrightarrow> openin (top_of_set S) c)"
by (metis components_iff locally_connected_open_connected_component)
proposition locally_connected_im_kleinen:
"locally connected S \<longleftrightarrow>
- (\<forall>v x. openin (subtopology euclidean S) v \<and> x \<in> v
- \<longrightarrow> (\<exists>u. openin (subtopology euclidean S) u \<and>
+ (\<forall>v x. openin (top_of_set S) v \<and> x \<in> v
+ \<longrightarrow> (\<exists>u. openin (top_of_set S) u \<and>
x \<in> u \<and> u \<subseteq> v \<and>
(\<forall>y. y \<in> u \<longrightarrow> (\<exists>c. connected c \<and> c \<subseteq> v \<and> x \<in> c \<and> y \<in> c))))"
(is "?lhs = ?rhs")
@@ -2521,12 +2521,12 @@
by (fastforce simp add: locally_connected)
next
assume ?rhs
- have *: "\<exists>T. openin (subtopology euclidean S) T \<and> x \<in> T \<and> T \<subseteq> c"
- if "openin (subtopology euclidean S) t" and c: "c \<in> components t" and "x \<in> c" for t c x
+ have *: "\<exists>T. openin (top_of_set S) T \<and> x \<in> T \<and> T \<subseteq> c"
+ if "openin (top_of_set S) t" and c: "c \<in> components t" and "x \<in> c" for t c x
proof -
from that \<open>?rhs\<close> [rule_format, of t x]
obtain u where u:
- "openin (subtopology euclidean S) u \<and> x \<in> u \<and> u \<subseteq> t \<and>
+ "openin (top_of_set S) u \<and> x \<in> u \<and> u \<subseteq> t \<and>
(\<forall>y. y \<in> u \<longrightarrow> (\<exists>c. connected c \<and> c \<subseteq> t \<and> x \<in> c \<and> y \<in> c))"
using in_components_subset by auto
obtain F :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a" where
@@ -2551,8 +2551,8 @@
proposition locally_path_connected_im_kleinen:
"locally path_connected S \<longleftrightarrow>
- (\<forall>v x. openin (subtopology euclidean S) v \<and> x \<in> v
- \<longrightarrow> (\<exists>u. openin (subtopology euclidean S) u \<and>
+ (\<forall>v x. openin (top_of_set S) v \<and> x \<in> v
+ \<longrightarrow> (\<exists>u. openin (top_of_set S) u \<and>
x \<in> u \<and> u \<subseteq> v \<and>
(\<forall>y. y \<in> u \<longrightarrow> (\<exists>p. path p \<and> path_image p \<subseteq> v \<and>
pathstart p = x \<and> pathfinish p = y))))"
@@ -2565,15 +2565,15 @@
by (meson dual_order.trans)
next
assume ?rhs
- have *: "\<exists>T. openin (subtopology euclidean S) T \<and>
+ have *: "\<exists>T. openin (top_of_set S) T \<and>
x \<in> T \<and> T \<subseteq> path_component_set u z"
- if "openin (subtopology euclidean S) u" and "z \<in> u" and c: "path_component u z x" for u z x
+ if "openin (top_of_set S) u" and "z \<in> u" and c: "path_component u z x" for u z x
proof -
have "x \<in> u"
by (meson c path_component_mem(2))
with that \<open>?rhs\<close> [rule_format, of u x]
obtain U where U:
- "openin (subtopology euclidean S) U \<and> x \<in> U \<and> U \<subseteq> u \<and>
+ "openin (top_of_set S) U \<and> x \<in> U \<and> U \<subseteq> u \<and>
(\<forall>y. y \<in> U \<longrightarrow> (\<exists>p. path p \<and> path_image p \<subseteq> u \<and> pathstart p = x \<and> pathfinish p = y))"
by blast
show ?thesis
@@ -2626,22 +2626,22 @@
lemma openin_connected_component_locally_connected:
"locally connected S
- \<Longrightarrow> openin (subtopology euclidean S) (connected_component_set S x)"
+ \<Longrightarrow> 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:
- "\<lbrakk>locally connected S; c \<in> components S\<rbrakk> \<Longrightarrow> openin (subtopology euclidean S) c"
+ "\<lbrakk>locally connected S; c \<in> components S\<rbrakk> \<Longrightarrow> 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
- \<Longrightarrow> openin (subtopology euclidean S) (path_component_set S x)"
+ \<Longrightarrow> 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
- \<Longrightarrow> closedin (subtopology euclidean S) (path_component_set S x)"
+ \<Longrightarrow> 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 \<in> 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: "\<And>T. T \<subseteq> f ` S
- \<Longrightarrow> openin (subtopology euclidean S) (S \<inter> f -` T) \<longleftrightarrow>
- openin (subtopology euclidean (f ` S)) T"
+ \<Longrightarrow> openin (top_of_set S) (S \<inter> f -` T) \<longleftrightarrow>
+ 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 \<in> components U"
+ assume opefSU: "openin (top_of_set (f ` S)) U" and "C \<in> components U"
then have "C \<subseteq> U" "U \<subseteq> f ` S"
by (meson in_components_subset openin_imp_subset)+
- then have "openin (subtopology euclidean (f ` S)) C \<longleftrightarrow>
- openin (subtopology euclidean S) (S \<inter> f -` C)"
+ then have "openin (top_of_set (f ` S)) C \<longleftrightarrow>
+ openin (top_of_set S) (S \<inter> f -` C)"
by (auto simp: oo)
- moreover have "openin (subtopology euclidean S) (S \<inter> f -` C)"
+ moreover have "openin (top_of_set S) (S \<inter> f -` C)"
proof (subst openin_subopen, clarify)
fix x
assume "x \<in> S" "f x \<in> C"
- show "\<exists>T. openin (subtopology euclidean S) T \<and> x \<in> T \<and> T \<subseteq> (S \<inter> f -` C)"
+ show "\<exists>T. openin (top_of_set S) T \<and> x \<in> T \<and> T \<subseteq> (S \<inter> f -` C)"
proof (intro conjI exI)
- show "openin (subtopology euclidean S) (connected_component_set (S \<inter> f -` U) x)"
+ show "openin (top_of_set S) (connected_component_set (S \<inter> f -` U) x)"
proof (rule ccontr)
- assume **: "\<not> openin (subtopology euclidean S) (connected_component_set (S \<inter> f -` U) x)"
+ assume **: "\<not> openin (top_of_set S) (connected_component_set (S \<inter> f -` U) x)"
then have "x \<notin> (S \<inter> f -` U)"
using \<open>U \<subseteq> f ` S\<close> opefSU lcS locally_connected_2 oo by blast
with ** show False
@@ -2768,7 +2768,7 @@
finally show "connected_component_set (S \<inter> f -` U) x \<subseteq> (S \<inter> 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: "\<And>T. T \<subseteq> f ` S
- \<Longrightarrow> openin (subtopology euclidean S) (S \<inter> f -` T) \<longleftrightarrow> openin (subtopology euclidean (f ` S)) T"
+ \<Longrightarrow> openin (top_of_set S) (S \<inter> f -` T) \<longleftrightarrow> 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 \<in> U"
+ assume opefSU: "openin (top_of_set (f ` S)) U" and "y \<in> U"
then have "path_component_set U y \<subseteq> U" "U \<subseteq> f ` S"
by (meson path_component_subset openin_imp_subset)+
- then have "openin (subtopology euclidean (f ` S)) (path_component_set U y) \<longleftrightarrow>
- openin (subtopology euclidean S) (S \<inter> f -` path_component_set U y)"
+ then have "openin (top_of_set (f ` S)) (path_component_set U y) \<longleftrightarrow>
+ openin (top_of_set S) (S \<inter> f -` path_component_set U y)"
proof -
have "path_component_set U y \<subseteq> f ` S"
using \<open>U \<subseteq> f ` S\<close> \<open>path_component_set U y \<subseteq> U\<close> by blast
then show ?thesis
using oo by blast
qed
- moreover have "openin (subtopology euclidean S) (S \<inter> f -` path_component_set U y)"
+ moreover have "openin (top_of_set S) (S \<inter> f -` path_component_set U y)"
proof (subst openin_subopen, clarify)
fix x
assume "x \<in> S" and Uyfx: "path_component U y (f x)"
then have "f x \<in> U"
using path_component_mem by blast
- show "\<exists>T. openin (subtopology euclidean S) T \<and> x \<in> T \<and> T \<subseteq> (S \<inter> f -` path_component_set U y)"
+ show "\<exists>T. openin (top_of_set S) T \<and> x \<in> T \<and> T \<subseteq> (S \<inter> f -` path_component_set U y)"
proof (intro conjI exI)
- show "openin (subtopology euclidean S) (path_component_set (S \<inter> f -` U) x)"
+ show "openin (top_of_set S) (path_component_set (S \<inter> f -` U) x)"
proof (rule ccontr)
- assume **: "\<not> openin (subtopology euclidean S) (path_component_set (S \<inter> f -` U) x)"
+ assume **: "\<not> openin (top_of_set S) (path_component_set (S \<inter> f -` U) x)"
then have "x \<notin> (S \<inter> f -` U)"
by (metis (no_types, lifting) \<open>U \<subseteq> f ` S\<close> opefSU lcS oo locally_path_connected_open_path_component)
then show False
@@ -2842,7 +2842,7 @@
finally show "path_component_set (S \<inter> f -` U) x \<subseteq> (S \<inter> 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 \<Rightarrow> 'b::topological_space"
assumes "\<And>c. c \<in> components S \<Longrightarrow>
- openin (subtopology euclidean S) c \<and> continuous_on c f"
+ openin (top_of_set S) c \<and> 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 \<inter> f -` t = (\<Union>c \<in> components S. c \<inter> f -` t)"
by auto
- show "openin (subtopology euclidean S) (S \<inter> f -` t)"
+ show "openin (top_of_set S) (S \<inter> f -` t)"
unfolding * using \<open>open t\<close> 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 \<subseteq> components(u - S)"
- shows "closedin (subtopology euclidean u) (S \<union> \<Union>c)"
+ shows "closedin (top_of_set u) (S \<union> \<Union>c)"
proof -
have di: "(\<And>S t. S \<in> c \<and> t \<in> c' \<Longrightarrow> disjnt S t) \<Longrightarrow> disjnt (\<Union> c) (\<Union> 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 \<union> \<Union>c = u - (\<Union>(components(u - S) - c))"
by (auto simp: disjnt_def)
- have *: "openin (subtopology euclidean u) (\<Union>(components (u - S) - c))"
+ have *: "openin (top_of_set u) (\<Union>(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 - \<Union>(components (u - S) - c)))"
+ have "openin (top_of_set u) (u - (u - \<Union>(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 \<subseteq> components(- S)"
shows "closed(S \<union> \<Union> c)"
proof -
- have "closedin (subtopology euclidean UNIV) (S \<union> \<Union>c)"
+ have "closedin (top_of_set UNIV) (S \<union> \<Union>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 \<in> components(u - S)"
- shows "closedin (subtopology euclidean u) (S \<union> c)"
+ shows "closedin (top_of_set u) (S \<union> c)"
proof -
- have "closedin (subtopology euclidean u) (S \<union> \<Union>{c})"
+ have "closedin (top_of_set u) (S \<union> \<Union>{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 "\<not> collinear S" "countable T"
shows "path_connected(S - T)"
proof (clarsimp simp add: path_connected_component)
@@ -3995,9 +3995,9 @@
assume xy: "x \<in> S" "x \<notin> T" "y \<in> S" "y \<notin> T"
show "path_component (S - T) x y"
proof (rule connected_equivalence_relation_gen [OF \<open>connected S\<close>, where P = "\<lambda>x. x \<notin> T"])
- show "\<exists>z. z \<in> U \<and> z \<notin> T" if opeU: "openin (subtopology euclidean S) U" and "x \<in> U" for U x
+ show "\<exists>z. z \<in> U \<and> z \<notin> T" if opeU: "openin (top_of_set S) U" and "x \<in> 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 \<open>x \<in> U\<close> obtain r where Usub: "U \<subseteq> affine hull S" and "r > 0"
and subU: "ball x r \<inter> affine hull S \<subseteq> U"
@@ -4024,7 +4024,7 @@
using \<open>countable T\<close> countable_subset by blast
then show ?thesis by blast
qed
- show "\<exists>U. openin (subtopology euclidean S) U \<and> x \<in> U \<and>
+ show "\<exists>U. openin (top_of_set S) U \<and> x \<in> U \<and>
(\<forall>x\<in>U. \<forall>y\<in>U. x \<notin> T \<and> y \<notin> T \<longrightarrow> path_component (S - T) x y)"
if "x \<in> S" for x
proof -
@@ -4046,9 +4046,9 @@
proof (intro exI conjI)
show "x \<in> ball x r \<inter> affine hull S"
using \<open>x \<in> S\<close> \<open>r > 0\<close> by (simp add: hull_inc)
- have "openin (subtopology euclidean (affine hull S)) (ball x r \<inter> affine hull S)"
+ have "openin (top_of_set (affine hull S)) (ball x r \<inter> affine hull S)"
by (subst inf.commute) (simp add: openin_Int_open)
- then show "openin (subtopology euclidean S) (ball x r \<inter> affine hull S)"
+ then show "openin (top_of_set S) (ball x r \<inter> affine hull S)"
by (rule openin_subset_trans [OF _ subS Ssub])
qed (use * path_component_trans in \<open>auto simp: path_connected_component path_component_of_subset [OF ST]\<close>)
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 "\<not> 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 "\<not> 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 \<subseteq> T"
and TS: "T \<subseteq> affine hull S"
and S: "connected S" "a \<in> S" "b \<in> S"
@@ -4371,7 +4371,7 @@
then show "bounded {x. \<not> ((f2 \<circ> f1) x = x \<and> (g1 \<circ> g2) x = x)}"
by (rule bounded_subset) auto
qed
- have 3: "\<exists>U. openin (subtopology euclidean S) U \<and>
+ have 3: "\<exists>U. openin (top_of_set S) U \<and>
d \<in> U \<and>
(\<forall>x\<in>U.
\<exists>f g. homeomorphism T T f g \<and> f d = x \<and>
@@ -4410,7 +4410,7 @@
assumes K: "finite K" "\<And>i. i \<in> K \<Longrightarrow> x i \<in> S \<and> y i \<in> S"
"pairwise (\<lambda>i j. (x i \<noteq> x j) \<and> (y i \<noteq> y j)) K"
and "2 \<le> aff_dim S"
- and ope: "openin (subtopology euclidean (affine hull S)) S"
+ and ope: "openin (top_of_set (affine hull S)) S"
and "S \<subseteq> T" "T \<subseteq> affine hull S" "connected S"
shows "\<exists>f g. homeomorphism T T f g \<and> (\<forall>i \<in> K. f(x i) = y i) \<and>
{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> g x = x)}"
@@ -4449,7 +4449,7 @@
and hk_sub: "{x. \<not> (h x = x \<and> k x = x)} \<subseteq> S - y ` K"
and bo_hk: "bounded {x. \<not> (h x = x \<and> 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 \<subseteq> T"
using \<open>S \<subseteq> T\<close> by auto
@@ -4499,7 +4499,7 @@
case False
then have affS: "affine hull S = UNIV"
by (simp add: affine_hull_open \<open>open S\<close>)
- then have ope: "openin (subtopology euclidean (affine hull S)) S"
+ then have ope: "openin (top_of_set (affine hull S)) S"
using \<open>open S\<close> open_openin by auto
have "2 \<le> DIM('a)" by (rule 2)
also have "\<dots> = aff_dim (UNIV :: 'a set)"
@@ -4714,7 +4714,7 @@
using \<open>cbox w z \<subseteq> S\<close> \<open>S \<subseteq> T\<close> by auto
show "homeomorphism T T f' g'"
proof
- have clo: "closedin (subtopology euclidean (cbox w z \<union> (T - box w z))) (T - box w z)"
+ have clo: "closedin (top_of_set (cbox w z \<union> (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 \<union> (T - box w z)) f'" "continuous_on (cbox w z \<union> (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 \<noteq> {}" "finite K" "K \<subseteq> S" and S: "S \<subseteq> T" "T \<subseteq> affine hull S" "connected S"
obtains f g where "homeomorphism T T f g" "{x. (\<not> (f x = x \<and> g x = x))} \<subseteq> S"
"bounded {x. (\<not> (f x = x \<and> g x = x))}" "\<And>x. x \<in> K \<Longrightarrow> f x \<in> U"
proof (cases "2 \<le> 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 \<in> U" "u \<in> S"
using \<open>U \<noteq> {}\<close> opeU openin_imp_subset by fastforce+
@@ -4892,7 +4892,7 @@
by (meson Topological_Spaces.connected_continuous_image \<open>connected S\<close> homeomorphism_cont1 homeomorphism_of_subsets homhj hull_subset top_greatest)
have hUS: "h ` U \<subseteq> 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 \<Longrightarrow> open (h ` U)" for U
+ have opn: "openin (top_of_set (affine hull S)) U \<Longrightarrow> 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)
--- a/src/HOL/Analysis/Jordan_Curve.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Jordan_Curve.thy Tue Mar 19 16:14:59 2019 +0000
@@ -18,7 +18,7 @@
proof -
have [simp]: "a \<notin> S" "a \<notin> T" "b \<notin> S" "b \<notin> T"
by (meson ComplD ccS ccT connected_component_in)+
- have clo: "closedin (subtopology euclidean (S \<union> T)) S" "closedin (subtopology euclidean (S \<union> T)) T"
+ have clo: "closedin (top_of_set (S \<union> T)) S" "closedin (top_of_set (S \<union> T)) T"
by (simp_all add: assms closed_subset compact_imp_closed)
obtain g where contg: "continuous_on S g"
and g: "\<And>x. x \<in> S \<Longrightarrow> exp (\<i>* 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} \<subseteq> middle"
proof (rule dense_accessible_frontier_point_pairs [OF \<open>open middle\<close> \<open>connected middle\<close>, of "path_image c \<inter> ball a0 (dist a0 b0)" "path_image c \<inter> ball b0 (dist a0 b0)"])
- show "openin (subtopology euclidean (frontier middle)) (path_image c \<inter> ball a0 (dist a0 b0))"
- "openin (subtopology euclidean (frontier middle)) (path_image c \<inter> ball b0 (dist a0 b0))"
+ show "openin (top_of_set (frontier middle)) (path_image c \<inter> ball a0 (dist a0 b0))"
+ "openin (top_of_set (frontier middle)) (path_image c \<inter> ball b0 (dist a0 b0))"
by (simp_all add: \<open>frontier middle = path_image c\<close> openin_open_Int)
show "path_image c \<inter> ball a0 (dist a0 b0) \<noteq> path_image c \<inter> ball b0 (dist a0 b0)"
using \<open>a0 \<noteq> b0\<close> \<open>b0 \<in> path_image c\<close> by auto
--- a/src/HOL/Analysis/Lebesgue_Measure.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Lebesgue_Measure.thy Tue Mar 19 16:14:59 2019 +0000
@@ -1341,11 +1341,11 @@
qed
lemma lebesgue_openin:
- "\<lbrakk>openin (subtopology euclidean S) T; S \<in> sets lebesgue\<rbrakk> \<Longrightarrow> T \<in> sets lebesgue"
+ "\<lbrakk>openin (top_of_set S) T; S \<in> sets lebesgue\<rbrakk> \<Longrightarrow> T \<in> sets lebesgue"
by (metis borel_open openin_open sets.Int sets_completionI_sets sets_lborel)
lemma lebesgue_closedin:
- "\<lbrakk>closedin (subtopology euclidean S) T; S \<in> sets lebesgue\<rbrakk> \<Longrightarrow> T \<in> sets lebesgue"
+ "\<lbrakk>closedin (top_of_set S) T; S \<in> sets lebesgue\<rbrakk> \<Longrightarrow> T \<in> sets lebesgue"
by (metis borel_closed closedin_closed sets.Int sets_completionI_sets sets_lborel)
end
--- a/src/HOL/Analysis/Polytope.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Polytope.thy Tue Mar 19 16:14:59 2019 +0000
@@ -3921,9 +3921,9 @@
proof (rule Baire [OF \<open>closed U\<close>])
show "countable {U - C |C. C \<in> \<S> \<and> aff_dim C < aff_dim U}"
using \<open>finite \<S>\<close> uncountable_infinite by fastforce
- have "\<And>C. C \<in> \<S> \<Longrightarrow> openin (subtopology euclidean U) (U-C)"
+ have "\<And>C. C \<in> \<S> \<Longrightarrow> 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 \<and> U \<subseteq> closure T"
+ then show "openin (top_of_set U) T \<and> U \<subseteq> closure T"
if "T \<in> {U - C |C. C \<in> \<S> \<and> aff_dim C < aff_dim U}" for T
using that dense_complement_convex_closed \<open>closed U\<close> \<open>convex U\<close> by auto
qed
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Analysis/Product_Topology.thy Tue Mar 19 16:14:59 2019 +0000
@@ -0,0 +1,459 @@
+theory Product_Topology
+imports Function_Topology
+begin
+
+lemma subset_UnE: (*FIXME MOVE*)
+ assumes "C \<subseteq> A \<union> B"
+ obtains A' B' where "A' \<subseteq> A" "B' \<subseteq> B" "C = A' \<union> B'"
+ by (metis assms Int_Un_distrib inf.order_iff inf_le2)
+
+section \<open>Product Topology\<close>
+
+subsection\<open>A binary product topology where the two types can be different.\<close>
+
+definition prod_topology :: "'a topology \<Rightarrow> 'b topology \<Rightarrow> ('a \<times> 'b) topology" where
+ "prod_topology X Y \<equiv> topology (arbitrary union_of (\<lambda>U. U \<in> {S \<times> T |S T. openin X S \<and> openin Y T}))"
+
+lemma open_product_open:
+ assumes "open A"
+ shows "\<exists>\<U>. \<U> \<subseteq> {S \<times> T |S T. open S \<and> open T} \<and> \<Union> \<U> = A"
+proof -
+ obtain f g where *: "\<And>u. u \<in> A \<Longrightarrow> open (f u) \<and> open (g u) \<and> u \<in> (f u) \<times> (g u) \<and> (f u) \<times> (g u) \<subseteq> A"
+ using open_prod_def [of A] assms by metis
+ let ?\<U> = "(\<lambda>u. f u \<times> g u) ` A"
+ show ?thesis
+ by (rule_tac x="?\<U>" in exI) (auto simp: dest: *)
+qed
+
+lemma open_product_open_eq: "(arbitrary union_of (\<lambda>U. \<exists>S T. U = S \<times> T \<and> open S \<and> 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 (\<lambda>U. U \<in> {S \<times> T |S T. openin X S \<and> openin Y T})"
+ unfolding prod_topology_def
+proof (rule topology_inverse')
+ show "istopology (arbitrary union_of (\<lambda>U. U \<in> {S \<times> T |S T. openin X S \<and> 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 \<times> topspace Y"
+proof -
+ have "topspace(prod_topology X Y) = \<Union> (Collect (openin (prod_topology X Y)))" (is "_ = ?Z")
+ unfolding topspace_def ..
+ also have "\<dots> = topspace X \<times> topspace Y"
+ proof
+ show "?Z \<subseteq> topspace X \<times> topspace Y"
+ apply (auto simp: openin_prod_topology union_of_def arbitrary_def)
+ using openin_subset by force+
+ next
+ have *: "\<exists>A B. topspace X \<times> topspace Y = A \<times> B \<and> openin X A \<and> openin Y B"
+ by blast
+ show "topspace X \<times> topspace Y \<subseteq> ?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 \<times> T) = prod_topology (subtopology X S) (subtopology Y T)"
+proof -
+ have "((\<lambda>U. \<exists>S T. U = S \<times> T \<and> openin X S \<and> openin Y T) relative_to S \<times> T) =
+ (\<lambda>U. \<exists>S' T'. U = S' \<times> T' \<and> (openin X relative_to S) S' \<and> (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 \<times> topspace Y)"
+ "prod_topology X (subtopology Y T) = subtopology (prod_topology X Y) (topspace X \<times> T)"
+by (auto simp: subtopology_Times)
+
+lemma prod_topology_discrete_topology:
+ "discrete_topology (S \<times> 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 \<times> 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 \<longleftrightarrow> continuous_on S h \<and> h ` S \<subseteq> 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 \<longleftrightarrow>
+ (\<forall>x y. (x,y) \<in> S \<longrightarrow> (\<exists>U V. openin X U \<and> openin Y V \<and> x \<in> U \<and> y \<in> V \<and> U \<times> V \<subseteq> 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 \<times> T) \<longleftrightarrow> S = {} \<or> T = {} \<or> openin X S \<and> openin Y T"
+proof (cases "S = {} \<or> 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 \<times> T) = (X closure_of S) \<times> (Y closure_of T)" (is "?lhs = ?rhs")
+proof
+ show "?lhs \<subseteq> ?rhs"
+ by (clarsimp simp: closure_of_def openin_prod_topology_alt) blast
+ show "?rhs \<subseteq> ?lhs"
+ by (clarsimp simp: closure_of_def openin_prod_topology_alt) (meson SigmaI subsetD)
+qed
+
+lemma closedin_Times:
+ "closedin (prod_topology X Y) (S \<times> T) \<longleftrightarrow> S = {} \<or> T = {} \<or> closedin X S \<and> 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 \<longleftrightarrow> continuous_map Z X (fst \<circ> f) \<and> continuous_map Z Y (snd \<circ> f)"
+ (is "?lhs = ?rhs")
+proof -
+ let ?g = "fst \<circ> f" and ?h = "snd \<circ> f"
+ have f: "f x = (?g x, ?h x)" for x
+ by auto
+ show ?thesis
+ proof (cases "(\<forall>x \<in> topspace Z. ?g x \<in> topspace X) \<and> (\<forall>x \<in> topspace Z. ?h x \<in> topspace Y)")
+ case True
+ show ?thesis
+ proof safe
+ assume "continuous_map Z (prod_topology X Y) f"
+ then have "openin Z {x \<in> topspace Z. fst (f x) \<in> U}" if "openin X U" for U
+ unfolding continuous_map_def using True that
+ apply clarify
+ apply (drule_tac x="U \<times> topspace Y" in spec)
+ by (simp add: openin_Times mem_Times_iff cong: conj_cong)
+ with True show "continuous_map Z X (fst \<circ> f)"
+ by (auto simp: continuous_map_def)
+ next
+ assume "continuous_map Z (prod_topology X Y) f"
+ then have "openin Z {x \<in> topspace Z. snd (f x) \<in> V}" if "openin Y V" for V
+ unfolding continuous_map_def using True that
+ apply clarify
+ apply (drule_tac x="topspace X \<times> V" in spec)
+ by (simp add: openin_Times mem_Times_iff cong: conj_cong)
+ with True show "continuous_map Z Y (snd \<circ> f)"
+ by (auto simp: continuous_map_def)
+ next
+ assume Z: "continuous_map Z X (fst \<circ> f)" "continuous_map Z Y (snd \<circ> f)"
+ have *: "openin Z {x \<in> topspace Z. f x \<in> W}"
+ if "\<And>w. w \<in> W \<Longrightarrow> \<exists>U V. openin X U \<and> openin Y V \<and> w \<in> U \<times> V \<and> U \<times> V \<subseteq> W" for W
+ proof (subst openin_subopen, clarify)
+ fix x :: "'a"
+ assume "x \<in> topspace Z" and "f x \<in> W"
+ with that [OF \<open>f x \<in> W\<close>]
+ obtain U V where UV: "openin X U" "openin Y V" "f x \<in> U \<times> V" "U \<times> V \<subseteq> W"
+ by auto
+ with Z UV show "\<exists>T. openin Z T \<and> x \<in> T \<and> T \<subseteq> {x \<in> topspace Z. f x \<in> W}"
+ apply (rule_tac x="{x \<in> topspace Z. ?g x \<in> U} \<inter> {x \<in> topspace Z. ?h x \<in> V}" in exI)
+ apply (auto simp: \<open>x \<in> topspace Z\<close> 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) (\<lambda>x. (f x,g x)) \<longleftrightarrow> continuous_map Z X f \<and> continuous_map Z Y g"
+ by (simp add: continuous_map_pairwise o_def)
+
+lemma continuous_map_pairedI [continuous_intros]:
+ "\<lbrakk>continuous_map Z X f; continuous_map Z Y g\<rbrakk> \<Longrightarrow> continuous_map Z (prod_topology X Y) (\<lambda>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 \<Longrightarrow> continuous_map Z X (fst \<circ> f)"
+ by (simp add: continuous_map_pairwise)
+
+lemma continuous_map_snd_of [continuous_intros]:
+ "continuous_map Z (prod_topology X Y) f \<Longrightarrow> continuous_map Z Y (snd \<circ> f)"
+ by (simp add: continuous_map_pairwise)
+
+lemma continuous_map_if_iff [simp]: "continuous_map X Y (\<lambda>x. if P then f x else g x) \<longleftrightarrow> continuous_map X Y (if P then f else g)"
+ by simp
+
+lemma continuous_map_if [continuous_intros]: "\<lbrakk>P \<Longrightarrow> continuous_map X Y f; ~P \<Longrightarrow> continuous_map X Y g\<rbrakk>
+ \<Longrightarrow> continuous_map X Y (\<lambda>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 \<longleftrightarrow> (topspace Y = {} \<longrightarrow> 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 \<longleftrightarrow> (topspace X = {} \<longrightarrow> 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 \<longleftrightarrow> (topspace Y = {} \<longrightarrow> 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 "\<exists>g. continuous_map X (prod_topology X Y) g \<and> (\<forall>x\<in>topspace X. fst (g x) = x)"
+ if y: "y \<in> topspace Y" for y
+ by (rule_tac x="\<lambda>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 \<longleftrightarrow> (topspace X = {} \<longrightarrow> 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 "\<exists>g. continuous_map Y (prod_topology X Y) g \<and> (\<forall>y\<in>topspace Y. snd (g y) = y)"
+ if x: "x \<in> topspace X" for x
+ by (rule_tac x="\<lambda>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 \<circ> fst) \<longleftrightarrow> topspace Y = {} \<or> 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 \<circ> snd) \<longleftrightarrow> topspace X = {} \<or> 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') (\<lambda>(x,y). (f x, g y)) \<longleftrightarrow>
+ topspace (prod_topology X Y) = {} \<or> continuous_map X X' f \<and> 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') (\<lambda>(x,y). (f x, g y)) (\<lambda>(x,y). (f' x, g' y)) \<longleftrightarrow>
+ topspace(prod_topology X Y) = {} \<and>
+ topspace(prod_topology X' Y') = {} \<or>
+ homeomorphic_maps X X' f f' \<and>
+ 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) (\<lambda>x. (x, f x)) \<longleftrightarrow> continuous_map X Y f"
+ (is "?lhs = ?rhs")
+proof
+ assume L: ?lhs
+ have "snd \<circ> (\<lambda>x. (x, f x)) = f"
+ by force
+ moreover have "continuous_map X Y (snd \<circ> (\<lambda>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 \<in> (prod_topology X Y) closure_of S"
+ shows "fst z \<in> X closure_of (fst ` S)" "snd z \<in> 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) \<longleftrightarrow> topspace(prod_topology X Y) = {} \<or> compact_space X \<and> 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 \<noteq> {}" "topspace Y \<noteq> {}"
+ by auto
+ have "compact_space X" "compact_space Y" if "compact_space(prod_topology X Y)"
+ proof -
+ have "compactin X (fst ` (topspace X \<times> topspace Y))"
+ by (metis compact_space_def continuous_map_fst image_compactin that topspace_prod_topology)
+ moreover
+ have "compactin Y (snd ` (topspace X \<times> 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 \<X> where "\<X> \<equiv> (\<lambda>V. topspace X \<times> V) ` Collect (openin Y)"
+ define \<Y> where "\<Y> \<equiv> (\<lambda>U. U \<times> 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 \<times> topspace Y \<subseteq> \<Union>(\<X> \<union> \<Y>)"
+ unfolding \<X>_def \<Y>_def by auto
+ fix \<C> :: "('a \<times> 'b) set set"
+ assume \<C>: "\<C> \<subseteq> \<X> \<union> \<Y>" "topspace X \<times> topspace Y \<subseteq> \<Union>\<C>"
+ then obtain \<X>' \<Y>' where XY: "\<X>' \<subseteq> \<X>" "\<Y>' \<subseteq> \<Y>" and \<C>eq: "\<C> = \<X>' \<union> \<Y>'"
+ using subset_UnE by metis
+ then have sub: "topspace X \<times> topspace Y \<subseteq> \<Union>(\<X>' \<union> \<Y>')"
+ using \<C> by simp
+ obtain \<U> \<V> where \<U>: "\<And>U. U \<in> \<U> \<Longrightarrow> openin X U" "\<Y>' = (\<lambda>U. U \<times> topspace Y) ` \<U>"
+ and \<V>: "\<And>V. V \<in> \<V> \<Longrightarrow> openin Y V" "\<X>' = (\<lambda>V. topspace X \<times> V) ` \<V>"
+ using XY by (clarsimp simp add: \<X>_def \<Y>_def subset_image_iff) (force simp add: subset_iff)
+ have "\<exists>\<D>. finite \<D> \<and> \<D> \<subseteq> \<X>' \<union> \<Y>' \<and> topspace X \<times> topspace Y \<subseteq> \<Union> \<D>"
+ proof -
+ have "topspace X \<subseteq> \<Union>\<U> \<or> topspace Y \<subseteq> \<Union>\<V>"
+ using \<U> \<V> \<C> \<C>eq by auto
+ then have *: "\<exists>\<D>. finite \<D> \<and>
+ (\<forall>x \<in> \<D>. x \<in> (\<lambda>V. topspace X \<times> V) ` \<V> \<or> x \<in> (\<lambda>U. U \<times> topspace Y) ` \<U>) \<and>
+ (topspace X \<times> topspace Y \<subseteq> \<Union>\<D>)"
+ proof
+ assume "topspace X \<subseteq> \<Union>\<U>"
+ with \<open>compact_space X\<close> \<U> obtain \<F> where "finite \<F>" "\<F> \<subseteq> \<U>" "topspace X \<subseteq> \<Union>\<F>"
+ by (meson compact_space_alt)
+ with that show ?thesis
+ by (rule_tac x="(\<lambda>D. D \<times> topspace Y) ` \<F>" in exI) auto
+ next
+ assume "topspace Y \<subseteq> \<Union>\<V>"
+ with \<open>compact_space Y\<close> \<V> obtain \<F> where "finite \<F>" "\<F> \<subseteq> \<V>" "topspace Y \<subseteq> \<Union>\<F>"
+ by (meson compact_space_alt)
+ with that show ?thesis
+ by (rule_tac x="(\<lambda>C. topspace X \<times> C) ` \<F>" in exI) auto
+ qed
+ then show ?thesis
+ using that \<U> \<V> by blast
+ qed
+ then show "\<exists>\<D>. finite \<D> \<and> \<D> \<subseteq> \<C> \<and> topspace X \<times> topspace Y \<subseteq> \<Union> \<D>"
+ using \<C> \<C>eq by blast
+ next
+ have "(finite intersection_of (\<lambda>x. x \<in> \<X> \<or> x \<in> \<Y>) relative_to topspace X \<times> topspace Y)
+ = (\<lambda>U. \<exists>S T. U = S \<times> T \<and> openin X S \<and> openin Y T)"
+ (is "?lhs = ?rhs")
+ proof -
+ have "?rhs U" if "?lhs U" for U
+ proof -
+ have "topspace X \<times> topspace Y \<inter> \<Inter> T \<in> {A \<times> B |A B. A \<in> Collect (openin X) \<and> B \<in> Collect (openin Y)}"
+ if "finite T" "T \<subseteq> \<X> \<union> \<Y>" for T
+ using that
+ proof induction
+ case (insert B \<B>)
+ then show ?case
+ unfolding \<X>_def \<Y>_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 \<times> V" "openin X U" "openin Y V"
+ using Z by blast
+ then have UV: "U \<times> V = (topspace X \<times> topspace Y) \<inter> (U \<times> V)"
+ by (simp add: Sigma_mono inf_absorb2 openin_subset)
+ moreover
+ have "?lhs ((topspace X \<times> topspace Y) \<inter> (U \<times> V))"
+ proof (rule relative_to_inc)
+ show "(finite intersection_of (\<lambda>x. x \<in> \<X> \<or> x \<in> \<Y>)) (U \<times> V)"
+ apply (simp add: intersection_of_def \<X>_def \<Y>_def)
+ apply (rule_tac x="{(U \<times> topspace Y),(topspace X \<times> V)}" in exI)
+ using \<open>openin X U\<close> \<open>openin Y V\<close> openin_subset UV apply (fastforce simp add:)
+ done
+ qed
+ ultimately show ?thesis
+ using \<open>Z = U \<times> V\<close> by auto
+ qed
+ ultimately show ?thesis
+ by meson
+ qed
+ then show "topology (arbitrary union_of (finite intersection_of (\<lambda>x. x \<in> \<X> \<union> \<Y>)
+ relative_to (topspace X \<times> 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 \<times> T) \<longleftrightarrow> S = {} \<or> T = {} \<or> compactin X S \<and> compactin Y T"
+ by (auto simp: compactin_subspace subtopology_Times compact_space_prod_topology)
+
+end
+
--- a/src/HOL/Analysis/Riemann_Mapping.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Riemann_Mapping.thy Tue Mar 19 16:14:59 2019 +0000
@@ -1029,14 +1029,14 @@
by (metis frontierS C)
obtain K where "C \<subseteq> K" "compact K"
and Ksub: "K \<subseteq> \<Inter>(range X)" and clo: "closed(\<Inter>(range X) - K)"
- proof (cases "{k. C \<subseteq> k \<and> compact k \<and> openin (subtopology euclidean (\<Inter>(range X))) k} = {}")
+ proof (cases "{k. C \<subseteq> k \<and> compact k \<and> openin (top_of_set (\<Inter>(range X))) k} = {}")
case True
then show ?thesis
using Sura_Bura [OF lcX Cco \<open>compact C\<close>] boC
by (simp add: True)
next
case False
- then obtain L where "compact L" "C \<subseteq> L" and K: "openin (subtopology euclidean (\<Inter>x. X x)) L"
+ then obtain L where "compact L" "C \<subseteq> L" and K: "openin (top_of_set (\<Inter>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 \<inter> closure U)"
+ then have "openin (top_of_set (X j)) (X j \<inter> closure U)"
by (simp add: openin_open_Int \<open>open U\<close>)
- moreover have "closedin (subtopology euclidean (X j)) (X j \<inter> closure U)"
+ moreover have "closedin (top_of_set (X j)) (X j \<inter> closure U)"
by (simp add: closedin_closed_Int)
moreover have "X j \<inter> closure U \<noteq> X j"
by (metis unboundedX \<open>compact (closure U)\<close> bounded_subset compact_eq_bounded_closed inf.order_iff)
--- a/src/HOL/Analysis/Starlike.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Starlike.thy Tue Mar 19 16:14:59 2019 +0000
@@ -8,7 +8,7 @@
chapter \<open>Unsorted\<close>
theory Starlike
-imports Convex_Euclidean_Space
+imports Convex_Euclidean_Space Abstract_Limits
begin
section \<open>Line Segments\<close>
@@ -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 \<inter> 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 \<subseteq> S'" "T \<subseteq> T'" "S' \<inter> T' = {}"
proof (cases "S = {} \<or> T = {}")
case True with that show ?thesis
@@ -4810,10 +4810,10 @@
proof (rule_tac S' = "(U \<inter> f -` {0<..})" and T' = "(U \<inter> f -` {..<0})" in that)
show "(U \<inter> f -` {0<..}) \<inter> (U \<inter> f -` {..<0}) = {}"
by auto
- show "openin (subtopology euclidean U) (U \<inter> f -` {0<..})"
+ show "openin (top_of_set U) (U \<inter> f -` {0<..})"
by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf)
next
- show "openin (subtopology euclidean U) (U \<inter> f -` {..<0})"
+ show "openin (top_of_set U) (U \<inter> f -` {..<0})"
by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf)
next
have "S \<subseteq> U" "T \<subseteq> U"
@@ -4993,10 +4993,10 @@
proposition proper_map:
fixes f :: "'a::heine_borel \<Rightarrow> 'b::heine_borel"
- assumes "closedin (subtopology euclidean S) K"
+ assumes "closedin (top_of_set S) K"
and com: "\<And>U. \<lbrakk>U \<subseteq> T; compact U\<rbrakk> \<Longrightarrow> compact (S \<inter> f -` U)"
and "f ` S \<subseteq> T"
- shows "closedin (subtopology euclidean T) (f ` K)"
+ shows "closedin (top_of_set T) (f ` K)"
proof -
have "K \<subseteq> S"
using assms closedin_imp_subset by metis
@@ -5221,7 +5221,7 @@
lemma proper_map_from_compact:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes contf: "continuous_on S f" and imf: "f ` S \<subseteq> T" and "compact S"
- "closedin (subtopology euclidean T) K"
+ "closedin (top_of_set T) K"
shows "compact (S \<inter> f -` K)"
by (rule closedin_compact [OF \<open>compact S\<close>] 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 \<times> T)) c"
- shows "closedin (subtopology euclidean S) (fst ` c)"
+ assumes "compact T" "closedin (top_of_set (S \<times> T)) c"
+ shows "closedin (top_of_set S) (fst ` c)"
proof -
have *: "fst ` (S \<times> T) \<subseteq> 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 \<times> T)) c"
- shows "closedin (subtopology euclidean T) (snd ` c)"
+ assumes "compact S" "closedin (top_of_set (S \<times> T)) c"
+ shows "closedin (top_of_set T) (snd ` c)"
proof -
have *: "snd ` (S \<times> T) \<subseteq> 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 \<times> T)) U"
- shows "closedin (subtopology euclidean T) {y. \<exists>x. x \<in> S \<and> (x, y) \<in> U}"
+ assumes "compact S" and clo: "closedin (top_of_set (S \<times> T)) U"
+ shows "closedin (top_of_set T) {y. \<exists>x. x \<in> S \<and> (x, y) \<in> U}"
proof -
have "U \<subseteq> S \<times> T"
by (metis clo closedin_imp_subset)
then have "{y. \<exists>x. x \<in> S \<and> (x, y) \<in> 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 \<inter> T \<noteq> {}"
+ assumes "convex S" "openin (top_of_set (affine hull S)) T" "S \<inter> T \<noteq> {}"
shows "affine hull (S \<inter> 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 \<noteq> {}"
+ assumes "openin (top_of_set (affine hull T)) S" "S \<noteq> {}"
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 \<subset> S"
+ assumes ope: "openin (top_of_set (affine hull S)) S" and "finite F" "F \<subset> 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 \<noteq> {}"
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 \<noteq> {}"
+ assumes ope: "openin (top_of_set T) S" and "affine T" "S \<noteq> {}"
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 \<noteq> {}"
+ assumes ope: "openin (top_of_set T) S" and "subspace T" "S \<noteq> {}"
shows "dim S = dim T"
proof (rule order_antisym)
show "dim S \<le> 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 \<subseteq> 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: "\<And>T. T \<in> \<C> \<Longrightarrow> openin (subtopology euclidean U) T"
+ assumes cin: "closedin (top_of_set U) S"
+ and oin: "\<And>T. T \<in> \<C> \<Longrightarrow> openin (top_of_set U) T"
and "S \<subseteq> \<Union>\<C>"
obtains \<C>' where "S \<subseteq> \<Union> \<C>'"
- and "\<And>V. V \<in> \<C>' \<Longrightarrow> openin (subtopology euclidean U) V \<and> (\<exists>T. T \<in> \<C> \<and> V \<subseteq> T)"
+ and "\<And>V. V \<in> \<C>' \<Longrightarrow> openin (top_of_set U) V \<and> (\<exists>T. T \<in> \<C> \<and> V \<subseteq> T)"
and "\<And>x. x \<in> U
- \<Longrightarrow> \<exists>V. openin (subtopology euclidean U) V \<and> x \<in> V \<and>
+ \<Longrightarrow> \<exists>V. openin (top_of_set U) V \<and> x \<in> V \<and>
finite {X. X \<in> \<C>' \<and> (X \<inter> V \<noteq> {})}"
proof -
have "\<exists>Z. open Z \<and> (T = U \<inter> Z)" if "T \<in> \<C>" for T
@@ -6830,12 +6830,12 @@
proof (rule_tac \<C>' = "{U \<inter> V |V. V \<in> \<D> \<and> (V \<inter> K \<noteq> {})}" in that)
show "S \<subseteq> \<Union>?C"
using \<open>U \<inter> K = S\<close> \<open>U \<subseteq> \<Union>\<D>\<close> K by (blast dest!: subsetD)
- show "\<And>V. V \<in> ?C \<Longrightarrow> openin (subtopology euclidean U) V \<and> (\<exists>T. T \<in> \<C> \<and> V \<subseteq> T)"
+ show "\<And>V. V \<in> ?C \<Longrightarrow> openin (top_of_set U) V \<and> (\<exists>T. T \<in> \<C> \<and> V \<subseteq> T)"
using D1 intF by fastforce
have *: "{X. (\<exists>V. X = U \<inter> V \<and> V \<in> \<D> \<and> V \<inter> K \<noteq> {}) \<and> X \<inter> (U \<inter> V) \<noteq> {}} \<subseteq>
(\<lambda>x. U \<inter> x) ` {U \<in> \<D>. U \<inter> V \<noteq> {}}" for V
by blast
- show "\<exists>V. openin (subtopology euclidean U) V \<and> x \<in> V \<and> finite {X \<in> ?C. X \<inter> V \<noteq> {}}"
+ show "\<exists>V. openin (top_of_set U) V \<and> x \<in> V \<and> finite {X \<in> ?C. X \<inter> V \<noteq> {}}"
if "x \<in> 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 \<subseteq> T"
- shows "closedin (subtopology euclidean (S \<times> T)) ((\<lambda>x. Pair x (f x)) ` S)"
+ shows "closedin (top_of_set (S \<times> T)) ((\<lambda>x. Pair x (f x)) ` S)"
proof -
have eq: "((\<lambda>x. Pair x (f x)) ` S) =(S \<times> T \<inter> (\<lambda>z. (f \<circ> fst)z - snd z) -` {0})"
using fim by auto
@@ -6876,16 +6876,16 @@
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "compact T" and fim: "f ` S \<subseteq> T"
shows "continuous_on S f \<longleftrightarrow>
- closedin (subtopology euclidean (S \<times> T)) ((\<lambda>x. Pair x (f x)) ` S)"
+ closedin (top_of_set (S \<times> T)) ((\<lambda>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 \<inter> f -` U) = fst ` (((\<lambda>x. Pair x (f x)) ` S) \<inter> (S \<times> U))"
by (force simp: image_iff)
- show "closedin (subtopology euclidean S) (S \<inter> f -` U)"
+ show "closedin (top_of_set S) (S \<inter> f -` U)"
by (simp add: U closedin_Int closedin_Times closed_map_fst [OF \<open>compact T\<close>] 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 \<open>compact T\<close> fim])
lemma continuous_on_Un_local_open:
- assumes opS: "openin (subtopology euclidean (S \<union> T)) S"
- and opT: "openin (subtopology euclidean (S \<union> T)) T"
+ assumes opS: "openin (top_of_set (S \<union> T)) S"
+ and opT: "openin (top_of_set (S \<union> T)) T"
and contf: "continuous_on S f" and contg: "continuous_on T f"
shows "continuous_on (S \<union> T) f"
-using pasting_lemma [of "{S,T}" "S \<union> T" "\<lambda>i. i" "\<lambda>i. f" f] contf contg opS opT by blast
+ using pasting_lemma [of "{S,T}" "top_of_set (S \<union> T)" id euclidean "\<lambda>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 \<union> T)) S"
- and opT: "openin (subtopology euclidean (S \<union> T)) T"
+ assumes opS: "openin (top_of_set (S \<union> T)) S"
+ and opT: "openin (top_of_set (S \<union> T)) T"
and contf: "continuous_on S f" and contg: "continuous_on T g"
and fg: "\<And>x. x \<in> S \<and> \<not>P x \<or> x \<in> T \<and> P x \<Longrightarrow> f x = g x"
shows "continuous_on (S \<union> T) (\<lambda>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 \<in> topspace X \<and> p x \<le> q x}) Y f"
+ and contg: "continuous_map (subtopology X {x. x \<in> topspace X \<and> q x \<le> p x}) Y g"
+ and fg: "\<And>x. \<lbrakk>x \<in> topspace X; p x = q x\<rbrakk> \<Longrightarrow> f x = g x"
+ shows "continuous_map X Y (\<lambda>x. if p x \<le> q x then f x else g x)"
+proof -
+ have "continuous_map X Y (\<lambda>x. if q x - p x \<in> {0..} then f x else g x)"
+ proof (rule continuous_map_cases_function)
+ show "continuous_map X euclideanreal (\<lambda>x. q x - p x)"
+ by (intro contp contq continuous_intros)
+ show "continuous_map (subtopology X {x \<in> topspace X. q x - p x \<in> euclideanreal closure_of {0..}}) Y f"
+ by (simp add: contf)
+ show "continuous_map (subtopology X {x \<in> topspace X. q x - p x \<in> 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 \<in> topspace X \<and> p x \<le> q x}) Y f"
+ and contg: "continuous_map (subtopology X {x. x \<in> topspace X \<and> q x \<le> p x}) Y g"
+ and fg: "\<And>x. \<lbrakk>x \<in> topspace X; p x = q x\<rbrakk> \<Longrightarrow> f x = g x"
+ shows "continuous_map X Y (\<lambda>x. if p x < q x then f x else g x)"
+proof -
+ have "continuous_map X Y (\<lambda>x. if q x - p x \<in> {0<..} then f x else g x)"
+ proof (rule continuous_map_cases_function)
+ show "continuous_map X euclideanreal (\<lambda>x. q x - p x)"
+ by (intro contp contq continuous_intros)
+ show "continuous_map (subtopology X {x \<in> topspace X. q x - p x \<in> euclideanreal closure_of {0<..}}) Y f"
+ by (simp add: contf)
+ show "continuous_map (subtopology X {x \<in> topspace X. q x - p x \<in> 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\<open>The union of two collinear segments is another segment\<close>
proposition%unimportant in_convex_hull_exchange:
--- a/src/HOL/Analysis/Topology_Euclidean_Space.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy Tue Mar 19 16:14:59 2019 +0000
@@ -1874,8 +1874,8 @@
obtain \<B> :: "'a set set"
where "countable \<B>"
and "{} \<notin> \<B>"
- and ope: "\<And>C. C \<in> \<B> \<Longrightarrow> openin(subtopology euclidean S) C"
- and if_ope: "\<And>T. openin(subtopology euclidean S) T \<Longrightarrow> \<exists>\<U>. \<U> \<subseteq> \<B> \<and> T = \<Union>\<U>"
+ and ope: "\<And>C. C \<in> \<B> \<Longrightarrow> openin(top_of_set S) C"
+ and if_ope: "\<And>T. openin(top_of_set S) T \<Longrightarrow> \<exists>\<U>. \<U> \<subseteq> \<B> \<and> T = \<Union>\<U>"
by (meson subset_second_countable)
then obtain f where f: "\<And>C. C \<in> \<B> \<Longrightarrow> f C \<in> C"
by (metis equals0I)
@@ -1889,7 +1889,7 @@
proof (clarsimp simp: closure_approachable)
fix x and e::real
assume "x \<in> S" "0 < e"
- have "openin (subtopology euclidean S) (S \<inter> ball x e)"
+ have "openin (top_of_set S) (S \<inter> ball x e)"
by (simp add: openin_Int_open)
with if_ope obtain \<U> where \<U>: "\<U> \<subseteq> \<B>" "S \<inter> ball x e = \<Union>\<U>"
by meson
@@ -2324,12 +2324,12 @@
by (simp add: setdist_eq_0_sing_1)
lemma setdist_eq_0_closedin:
- shows "\<lbrakk>closedin (subtopology euclidean U) S; x \<in> U\<rbrakk>
+ shows "\<lbrakk>closedin (top_of_set U) S; x \<in> U\<rbrakk>
\<Longrightarrow> (setdist {x} S = 0 \<longleftrightarrow> S = {} \<or> x \<in> S)"
by (auto simp: closedin_limpt setdist_eq_0_sing_1 closure_def)
lemma setdist_gt_0_closedin:
- shows "\<lbrakk>closedin (subtopology euclidean U) S; x \<in> U; S \<noteq> {}; x \<notin> S\<rbrakk>
+ shows "\<lbrakk>closedin (top_of_set U) S; x \<in> U; S \<noteq> {}; x \<notin> S\<rbrakk>
\<Longrightarrow> setdist {x} S > 0"
using less_eq_real_def setdist_eq_0_closedin by fastforce
--- a/src/HOL/Analysis/Vitali_Covering_Theorem.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Analysis/Vitali_Covering_Theorem.thy Tue Mar 19 16:14:59 2019 +0000
@@ -542,10 +542,10 @@
next
assume R [rule_format]: "\<forall>x \<in> S. \<forall>e > 0. ?Q x e"
let ?\<mu> = "measure lebesgue"
- have "\<exists>U. openin (subtopology euclidean S) U \<and> z \<in> U \<and> negligible U"
+ have "\<exists>U. openin (top_of_set S) U \<and> z \<in> U \<and> negligible U"
if "z \<in> S" for z
proof (intro exI conjI)
- show "openin (subtopology euclidean S) (S \<inter> ball z 1)"
+ show "openin (top_of_set S) (S \<inter> ball z 1)"
by (simp add: openin_open_Int)
show "z \<in> S \<inter> ball z 1"
using \<open>z \<in> S\<close> by auto
--- a/src/HOL/Product_Type.thy Mon Mar 18 21:50:51 2019 +0100
+++ b/src/HOL/Product_Type.thy Tue Mar 19 16:14:59 2019 +0000
@@ -1124,7 +1124,7 @@
lemma vimage_snd: "snd -` A = UNIV \<times> A"
by auto
-lemma insert_times_insert [simp]:
+lemma insert_Times_insert [simp]:
"insert a A \<times> insert b B = insert (a,b) (A \<times> insert b B \<union> insert a A \<times> B)"
by blast