| author | wenzelm | 
| Sun, 20 Oct 2024 21:48:38 +0200 | |
| changeset 81214 | 7c2745efec69 | 
| parent 78037 | 37894dff0111 | 
| child 82520 | 1b17f0a41aa3 | 
| permissions | -rw-r--r-- | 
| 77939 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1 | section\<open>Disjoint sum of arbitarily many spaces\<close> | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2 | |
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3 | theory Sum_Topology | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4 | imports Abstract_Topology | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5 | begin | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 6 | |
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 7 | |
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 8 | definition sum_topology :: "('a \<Rightarrow> 'b topology) \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'b) topology" where
 | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 9 | "sum_topology X I \<equiv> | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 10 |     topology (\<lambda>U. U \<subseteq> Sigma I (topspace \<circ> X) \<and> (\<forall>i \<in> I. openin (X i) {x. (i,x) \<in> U}))"
 | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 11 | |
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 12 | lemma is_sum_topology: "istopology (\<lambda>U. U \<subseteq> Sigma I (topspace \<circ> X) \<and> (\<forall>i\<in>I. openin (X i) {x. (i, x) \<in> U}))"
 | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 13 | proof - | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 14 |   have 1: "{x. (i, x) \<in> S \<inter> T} = {x. (i, x) \<in> S} \<inter> {x::'b. (i, x) \<in> T}" for S T and i::'a
 | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 15 | by auto | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 16 |   have 2: "{x. (i, x) \<in> \<Union> \<K>} = (\<Union>K\<in>\<K>. {x::'b. (i, x) \<in> K})" for \<K> and i::'a
 | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 17 | by auto | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 18 | show ?thesis | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 19 | unfolding istopology_def 1 2 by blast | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 20 | qed | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 21 | |
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 22 | lemma openin_sum_topology: | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 23 | "openin (sum_topology X I) U \<longleftrightarrow> | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 24 |         U \<subseteq> Sigma I (topspace \<circ> X) \<and> (\<forall>i \<in> I. openin (X i) {x. (i,x) \<in> U})"
 | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 25 | by (auto simp: sum_topology_def is_sum_topology) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 26 | |
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 27 | lemma openin_disjoint_union: | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 28 | "openin (sum_topology X I) (Sigma I U) \<longleftrightarrow> (\<forall>i \<in> I. openin (X i) (U i))" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 29 | using openin_subset by (force simp: openin_sum_topology) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 30 | |
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 31 | lemma topspace_sum_topology [simp]: | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 32 | "topspace(sum_topology X I) = Sigma I (topspace \<circ> X)" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 33 | by (metis comp_apply openin_disjoint_union openin_subset openin_sum_topology openin_topspace subset_antisym) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 34 | |
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 35 | lemma openin_sum_topology_alt: | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 36 | "openin (sum_topology X I) U \<longleftrightarrow> (\<exists>T. U = Sigma I T \<and> (\<forall>i \<in> I. openin (X i) (T i)))" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 37 | by (bestsimp simp: openin_sum_topology dest: openin_subset) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 38 | |
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 39 | lemma forall_openin_sum_topology: | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 40 | "(\<forall>U. openin (sum_topology X I) U \<longrightarrow> P U) \<longleftrightarrow> (\<forall>T. (\<forall>i \<in> I. openin (X i) (T i)) \<longrightarrow> P(Sigma I T))" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 41 | by (auto simp: openin_sum_topology_alt) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 42 | |
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 43 | lemma exists_openin_sum_topology: | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 44 | "(\<exists>U. openin (sum_topology X I) U \<and> P U) \<longleftrightarrow> | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 45 | (\<exists>T. (\<forall>i \<in> I. openin (X i) (T i)) \<and> P(Sigma I T))" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 46 | by (auto simp: openin_sum_topology_alt) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 47 | |
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 48 | lemma closedin_sum_topology: | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 49 |    "closedin (sum_topology X I) U \<longleftrightarrow> U \<subseteq> Sigma I (topspace \<circ> X) \<and> (\<forall>i \<in> I. closedin (X i) {x. (i,x) \<in> U})"
 | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 50 | (is "?lhs = ?rhs") | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 51 | proof | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 52 | assume L: ?lhs | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 53 | then have U: "U \<subseteq> Sigma I (topspace \<circ> X)" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 54 | using closedin_subset by fastforce | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 55 |   then have "\<forall>i\<in>I. {x. (i, x) \<in> U} \<subseteq> topspace (X i)"
 | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 56 | by fastforce | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 57 |   moreover have "openin (X i) (topspace (X i) - {x. (i, x) \<in> U})" if "i\<in>I" for i
 | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 58 | apply (subst openin_subopen) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 59 | using L that unfolding closedin_def openin_sum_topology | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 60 | by (bestsimp simp: o_def subset_iff) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 61 | ultimately show ?rhs | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 62 | by (simp add: U closedin_def) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 63 | next | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 64 | assume R: ?rhs | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 65 |   then have "openin (X i) {x. (i, x) \<in> topspace (sum_topology X I) - U}" if "i\<in>I" for i
 | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 66 | apply (subst openin_subopen) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 67 | using that unfolding closedin_def openin_sum_topology | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 68 | by (bestsimp simp: o_def subset_iff) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 69 | then show ?lhs | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 70 | by (simp add: R closedin_def openin_sum_topology) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 71 | qed | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 72 | |
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 73 | lemma closedin_disjoint_union: | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 74 | "closedin (sum_topology X I) (Sigma I U) \<longleftrightarrow> (\<forall>i \<in> I. closedin (X i) (U i))" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 75 | using closedin_subset by (force simp: closedin_sum_topology) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 76 | |
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 77 | lemma closedin_sum_topology_alt: | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 78 | "closedin (sum_topology X I) U \<longleftrightarrow> (\<exists>T. U = Sigma I T \<and> (\<forall>i \<in> I. closedin (X i) (T i)))" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 79 | using closedin_subset unfolding closedin_sum_topology by auto blast | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 80 | |
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 81 | lemma forall_closedin_sum_topology: | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 82 | "(\<forall>U. closedin (sum_topology X I) U \<longrightarrow> P U) \<longleftrightarrow> | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 83 | (\<forall>t. (\<forall>i \<in> I. closedin (X i) (t i)) \<longrightarrow> P(Sigma I t))" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 84 | by (metis closedin_sum_topology_alt) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 85 | |
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 86 | lemma exists_closedin_sum_topology: | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 87 | "(\<exists>U. closedin (sum_topology X I) U \<and> P U) \<longleftrightarrow> | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 88 | (\<exists>T. (\<forall>i \<in> I. closedin (X i) (T i)) \<and> P(Sigma I T))" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 89 | by (simp add: closedin_sum_topology_alt) blast | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 90 | |
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 91 | lemma open_map_component_injection: | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 92 | "i \<in> I \<Longrightarrow> open_map (X i) (sum_topology X I) (\<lambda>x. (i,x))" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 93 | unfolding open_map_def openin_sum_topology | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 94 | using openin_subset openin_subopen by (force simp: image_iff) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 95 | |
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 96 | lemma closed_map_component_injection: | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 97 | assumes "i \<in> I" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 98 | shows "closed_map(X i) (sum_topology X I) (\<lambda>x. (i,x))" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 99 | proof - | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 100 |   have "closedin (X j) {x. j = i \<and> x \<in> U}"
 | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 101 | if "\<And>U S. closedin U S \<Longrightarrow> S \<subseteq> topspace U" and "closedin (X i) U" and "j \<in> I" for U j | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 102 | using that by (cases "j=i") auto | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 103 | then show ?thesis | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 104 | using closedin_subset assms by (force simp: closed_map_def closedin_sum_topology image_iff) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 105 | qed | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 106 | |
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 107 | lemma continuous_map_component_injection: | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 108 | "i \<in> I \<Longrightarrow> continuous_map(X i) (sum_topology X I) (\<lambda>x. (i,x))" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 109 | apply (clarsimp simp: continuous_map_def openin_sum_topology) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 110 | by (smt (verit, best) Collect_cong mem_Collect_eq openin_subset subsetD) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 111 | |
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 112 | lemma subtopology_sum_topology: | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 113 | "subtopology (sum_topology X I) (Sigma I S) = | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 114 | sum_topology (\<lambda>i. subtopology (X i) (S i)) I" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 115 | unfolding topology_eq | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 116 | proof (intro strip iffI) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 117 | fix T | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 118 | assume *: "openin (subtopology (sum_topology X I) (Sigma I S)) T" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 119 | then obtain U where U: "\<forall>i\<in>I. openin (X i) (U i) \<and> T = Sigma I U \<inter> Sigma I S" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 120 | by (auto simp: openin_subtopology openin_sum_topology_alt) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 121 | have "T = (SIGMA i:I. U i \<inter> S i)" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 122 | proof | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 123 | show "T \<subseteq> (SIGMA i:I. U i \<inter> S i)" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 124 | by (metis "*" SigmaE Sigma_Int_distrib2 U openin_imp_subset subset_iff) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 125 | show "(SIGMA i:I. U i \<inter> S i) \<subseteq> T" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 126 | using U by fastforce | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 127 | qed | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 128 | then show "openin (sum_topology (\<lambda>i. subtopology (X i) (S i)) I) T" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 129 | by (simp add: U openin_disjoint_union openin_subtopology_Int) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 130 | next | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 131 | fix T | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 132 | assume *: "openin (sum_topology (\<lambda>i. subtopology (X i) (S i)) I) T" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 133 | then obtain U where "\<forall>i\<in>I. \<exists>T. openin (X i) T \<and> U i = T \<inter> S i" and Teq: "T = Sigma I U" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 134 | by (auto simp: openin_subtopology openin_sum_topology_alt) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 135 | then obtain B where "\<And>i. i\<in>I \<Longrightarrow> openin (X i) (B i) \<and> U i = B i \<inter> S i" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 136 | by metis | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 137 | then show "openin (subtopology (sum_topology X I) (Sigma I S)) T" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 138 | by (auto simp: openin_subtopology openin_sum_topology_alt Teq) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 139 | qed | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 140 | |
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 141 | lemma embedding_map_component_injection: | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 142 | "i \<in> I \<Longrightarrow> embedding_map (X i) (sum_topology X I) (\<lambda>x. (i,x))" | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 143 | by (metis injective_open_imp_embedding_map continuous_map_component_injection | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 144 | open_map_component_injection inj_onI prod.inject) | 
| 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 145 | |
| 78037 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77939diff
changeset | 146 | lemma topological_property_of_sum_component: | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77939diff
changeset | 147 | assumes major: "P (sum_topology X I)" | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77939diff
changeset | 148 | and minor: "\<And>X S. \<lbrakk>P X; closedin X S; openin X S\<rbrakk> \<Longrightarrow> P(subtopology X S)" | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77939diff
changeset | 149 | and PQ: "\<And>X Y. X homeomorphic_space Y \<Longrightarrow> (P X \<longleftrightarrow> Q Y)" | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77939diff
changeset | 150 | shows "(\<forall>i \<in> I. Q(X i))" | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77939diff
changeset | 151 | proof - | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77939diff
changeset | 152 | have "Q(X i)" if "i \<in> I" for i | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77939diff
changeset | 153 | proof - | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77939diff
changeset | 154 | have "P(subtopology (sum_topology X I) (Pair i ` topspace (X i)))" | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77939diff
changeset | 155 | by (meson closed_map_component_injection closed_map_def closedin_topspace major minor open_map_component_injection open_map_def openin_topspace that) | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77939diff
changeset | 156 | then show ?thesis | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77939diff
changeset | 157 | by (metis PQ embedding_map_component_injection embedding_map_imp_homeomorphic_space homeomorphic_space_sym that) | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77939diff
changeset | 158 | qed | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77939diff
changeset | 159 | then show ?thesis by metis | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77939diff
changeset | 160 | qed | 
| 
37894dff0111
More material from the HOL Light metric space library
 paulson <lp15@cam.ac.uk> parents: 
77939diff
changeset | 161 | |
| 77939 
98879407d33c
Two new theories containing material ported from HOL Light about abstract topology
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 162 | end |