| author | Andreas Lochbihler <mail@andreas-lochbihler.de> | 
| Sun, 31 Jan 2021 12:10:20 +0100 | |
| changeset 73213 | bb35f7f60d6c | 
| parent 71633 | 07bec530f02e | 
| child 77935 | 7f240b0dabd9 | 
| permissions | -rw-r--r-- | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 1 | section\<open>The binary product topology\<close> | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 2 | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3 | theory Product_Topology | 
| 71200 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
70178diff
changeset | 4 | imports Function_Topology | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5 | begin | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 6 | |
| 71200 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
70178diff
changeset | 7 | section \<open>Product Topology\<close> | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 8 | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 9 | subsection\<open>Definition\<close> | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 10 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 11 | definition prod_topology :: "'a topology \<Rightarrow> 'b topology \<Rightarrow> ('a \<times> 'b) topology" where
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 12 |  "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}))"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 13 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 14 | lemma open_product_open: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 15 | assumes "open A" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 16 |   shows "\<exists>\<U>. \<U> \<subseteq> {S \<times> T |S T. open S \<and> open T} \<and> \<Union> \<U> = A"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 17 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 18 | 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" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 19 | using open_prod_def [of A] assms by metis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 20 | let ?\<U> = "(\<lambda>u. f u \<times> g u) ` A" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 21 | show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 22 | by (rule_tac x="?\<U>" in exI) (auto simp: dest: *) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 23 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 24 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 25 | lemma open_product_open_eq: "(arbitrary union_of (\<lambda>U. \<exists>S T. U = S \<times> T \<and> open S \<and> open T)) = open" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 26 | by (force simp: union_of_def arbitrary_def intro: open_product_open open_Times) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 27 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 28 | lemma openin_prod_topology: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 29 |    "openin (prod_topology X Y) = arbitrary union_of (\<lambda>U. U \<in> {S \<times> T |S T. openin X S \<and> openin Y T})"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 30 | unfolding prod_topology_def | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 31 | proof (rule topology_inverse') | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 32 |   show "istopology (arbitrary union_of (\<lambda>U. U \<in> {S \<times> T |S T. openin X S \<and> openin Y T}))"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 33 | apply (rule istopology_base, simp) | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 34 | by (metis openin_Int Times_Int_Times) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 35 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 36 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 37 | lemma topspace_prod_topology [simp]: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 38 | "topspace (prod_topology X Y) = topspace X \<times> topspace Y" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 39 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 40 | have "topspace(prod_topology X Y) = \<Union> (Collect (openin (prod_topology X Y)))" (is "_ = ?Z") | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 41 | unfolding topspace_def .. | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 42 | also have "\<dots> = topspace X \<times> topspace Y" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 43 | proof | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 44 | show "?Z \<subseteq> topspace X \<times> topspace Y" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 45 | apply (auto simp: openin_prod_topology union_of_def arbitrary_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 46 | using openin_subset by force+ | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 47 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 48 | have *: "\<exists>A B. topspace X \<times> topspace Y = A \<times> B \<and> openin X A \<and> openin Y B" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 49 | by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 50 | show "topspace X \<times> topspace Y \<subseteq> ?Z" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 51 | apply (rule Union_upper) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 52 | using * by (simp add: openin_prod_topology arbitrary_union_of_inc) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 53 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 54 | finally show ?thesis . | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 55 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 56 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 57 | lemma subtopology_Times: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 58 | shows "subtopology (prod_topology X Y) (S \<times> T) = prod_topology (subtopology X S) (subtopology Y T)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 59 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 60 | have "((\<lambda>U. \<exists>S T. U = S \<times> T \<and> openin X S \<and> openin Y T) relative_to S \<times> T) = | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 61 | (\<lambda>U. \<exists>S' T'. U = S' \<times> T' \<and> (openin X relative_to S) S' \<and> (openin Y relative_to T) T')" | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 62 | by (auto simp: relative_to_def Times_Int_Times fun_eq_iff) metis | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 63 | then show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 64 | by (simp add: topology_eq openin_prod_topology arbitrary_union_of_relative_to flip: openin_relative_to) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 65 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 66 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 67 | lemma prod_topology_subtopology: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 68 | "prod_topology (subtopology X S) Y = subtopology (prod_topology X Y) (S \<times> topspace Y)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 69 | "prod_topology X (subtopology Y T) = subtopology (prod_topology X Y) (topspace X \<times> T)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 70 | by (auto simp: subtopology_Times) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 71 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 72 | lemma prod_topology_discrete_topology: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 73 | "discrete_topology (S \<times> T) = prod_topology (discrete_topology S) (discrete_topology T)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 74 | by (auto simp: discrete_topology_unique openin_prod_topology intro: arbitrary_union_of_inc) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 75 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 76 | lemma prod_topology_euclidean [simp]: "prod_topology euclidean euclidean = euclidean" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 77 | by (simp add: prod_topology_def open_product_open_eq) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 78 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 79 | lemma prod_topology_subtopology_eu [simp]: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 80 | "prod_topology (subtopology euclidean S) (subtopology euclidean T) = subtopology euclidean (S \<times> T)" | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 81 | by (simp add: prod_topology_subtopology subtopology_subtopology Times_Int_Times) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 82 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 83 | lemma openin_prod_topology_alt: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 84 | "openin (prod_topology X Y) S \<longleftrightarrow> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 85 | (\<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))" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 86 | apply (auto simp: openin_prod_topology arbitrary_union_of_alt, fastforce) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 87 | by (metis mem_Sigma_iff) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 88 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 89 | lemma open_map_fst: "open_map (prod_topology X Y) X fst" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 90 | unfolding open_map_def openin_prod_topology_alt | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 91 | by (force simp: openin_subopen [of X "fst ` _"] intro: subset_fst_imageI) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 92 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 93 | lemma open_map_snd: "open_map (prod_topology X Y) Y snd" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 94 | unfolding open_map_def openin_prod_topology_alt | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 95 | by (force simp: openin_subopen [of Y "snd ` _"] intro: subset_snd_imageI) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 96 | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 97 | lemma openin_prod_Times_iff: | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 98 |      "openin (prod_topology X Y) (S \<times> T) \<longleftrightarrow> S = {} \<or> T = {} \<or> openin X S \<and> openin Y T"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 99 | proof (cases "S = {} \<or> T = {}")
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 100 | case False | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 101 | then show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 102 | apply (simp add: openin_prod_topology_alt openin_subopen [of X S] openin_subopen [of Y T] times_subset_iff, safe) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 103 | apply (meson|force)+ | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 104 | done | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 105 | qed force | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 106 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 107 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 108 | lemma closure_of_Times: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 109 | "(prod_topology X Y) closure_of (S \<times> T) = (X closure_of S) \<times> (Y closure_of T)" (is "?lhs = ?rhs") | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 110 | proof | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 111 | show "?lhs \<subseteq> ?rhs" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 112 | by (clarsimp simp: closure_of_def openin_prod_topology_alt) blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 113 | show "?rhs \<subseteq> ?lhs" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 114 | by (clarsimp simp: closure_of_def openin_prod_topology_alt) (meson SigmaI subsetD) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 115 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 116 | |
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 117 | lemma closedin_prod_Times_iff: | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 118 |    "closedin (prod_topology X Y) (S \<times> T) \<longleftrightarrow> S = {} \<or> T = {} \<or> closedin X S \<and> closedin Y T"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 119 | by (auto simp: closure_of_Times times_eq_iff simp flip: closure_of_eq) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 120 | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 121 | lemma interior_of_Times: "(prod_topology X Y) interior_of (S \<times> T) = (X interior_of S) \<times> (Y interior_of T)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 122 | proof (rule interior_of_unique) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 123 | show "(X interior_of S) \<times> Y interior_of T \<subseteq> S \<times> T" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 124 | by (simp add: Sigma_mono interior_of_subset) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 125 | show "openin (prod_topology X Y) ((X interior_of S) \<times> Y interior_of T)" | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 126 | by (simp add: openin_prod_Times_iff) | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 127 | next | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 128 | show "T' \<subseteq> (X interior_of S) \<times> Y interior_of T" if "T' \<subseteq> S \<times> T" "openin (prod_topology X Y) T'" for T' | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 129 | proof (clarsimp; intro conjI) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 130 | fix a :: "'a" and b :: "'b" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 131 | assume "(a, b) \<in> T'" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 132 | with that obtain U V where UV: "openin X U" "openin Y V" "a \<in> U" "b \<in> V" "U \<times> V \<subseteq> T'" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 133 | by (metis openin_prod_topology_alt) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 134 | then show "a \<in> X interior_of S" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 135 | using interior_of_maximal_eq that(1) by fastforce | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 136 | show "b \<in> Y interior_of T" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 137 | using UV interior_of_maximal_eq that(1) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 138 | by (metis SigmaI mem_Sigma_iff subset_eq) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 139 | qed | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 140 | qed | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 141 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 142 | subsection \<open>Continuity\<close> | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 143 | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 144 | lemma continuous_map_pairwise: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 145 | "continuous_map Z (prod_topology X Y) f \<longleftrightarrow> continuous_map Z X (fst \<circ> f) \<and> continuous_map Z Y (snd \<circ> f)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 146 | (is "?lhs = ?rhs") | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 147 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 148 | let ?g = "fst \<circ> f" and ?h = "snd \<circ> f" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 149 | have f: "f x = (?g x, ?h x)" for x | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 150 | by auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 151 | show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 152 | proof (cases "(\<forall>x \<in> topspace Z. ?g x \<in> topspace X) \<and> (\<forall>x \<in> topspace Z. ?h x \<in> topspace Y)") | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 153 | case True | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 154 | show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 155 | proof safe | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 156 | assume "continuous_map Z (prod_topology X Y) f" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 157 |       then have "openin Z {x \<in> topspace Z. fst (f x) \<in> U}" if "openin X U" for U
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 158 | unfolding continuous_map_def using True that | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 159 | apply clarify | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 160 | apply (drule_tac x="U \<times> topspace Y" in spec) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 161 | by (simp add: openin_prod_Times_iff mem_Times_iff cong: conj_cong) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 162 | with True show "continuous_map Z X (fst \<circ> f)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 163 | by (auto simp: continuous_map_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 164 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 165 | assume "continuous_map Z (prod_topology X Y) f" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 166 |       then have "openin Z {x \<in> topspace Z. snd (f x) \<in> V}" if "openin Y V" for V
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 167 | unfolding continuous_map_def using True that | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 168 | apply clarify | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 169 | apply (drule_tac x="topspace X \<times> V" in spec) | 
| 69986 
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
 paulson <lp15@cam.ac.uk> parents: 
69945diff
changeset | 170 | by (simp add: openin_prod_Times_iff mem_Times_iff cong: conj_cong) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 171 | with True show "continuous_map Z Y (snd \<circ> f)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 172 | by (auto simp: continuous_map_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 173 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 174 | assume Z: "continuous_map Z X (fst \<circ> f)" "continuous_map Z Y (snd \<circ> f)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 175 |       have *: "openin Z {x \<in> topspace Z. f x \<in> W}"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 176 | 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 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 177 | proof (subst openin_subopen, clarify) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 178 | fix x :: "'a" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 179 | assume "x \<in> topspace Z" and "f x \<in> W" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 180 | with that [OF \<open>f x \<in> W\<close>] | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 181 | obtain U V where UV: "openin X U" "openin Y V" "f x \<in> U \<times> V" "U \<times> V \<subseteq> W" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 182 | by auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 183 |         with Z  UV show "\<exists>T. openin Z T \<and> x \<in> T \<and> T \<subseteq> {x \<in> topspace Z. f x \<in> W}"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 184 |           apply (rule_tac x="{x \<in> topspace Z. ?g x \<in> U} \<inter> {x \<in> topspace Z. ?h x \<in> V}" in exI)
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 185 | apply (auto simp: \<open>x \<in> topspace Z\<close> continuous_map_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 186 | done | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 187 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 188 | show "continuous_map Z (prod_topology X Y) f" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 189 | using True by (simp add: continuous_map_def openin_prod_topology_alt mem_Times_iff *) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 190 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 191 | qed (auto simp: continuous_map_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 192 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 193 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 194 | lemma continuous_map_paired: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 195 | "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" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 196 | by (simp add: continuous_map_pairwise o_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 197 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 198 | lemma continuous_map_pairedI [continuous_intros]: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 199 | "\<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))" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 200 | by (simp add: continuous_map_pairwise o_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 201 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 202 | lemma continuous_map_fst [continuous_intros]: "continuous_map (prod_topology X Y) X fst" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 203 | using continuous_map_pairwise [of "prod_topology X Y" X Y id] | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 204 | by (simp add: continuous_map_pairwise) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 205 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 206 | lemma continuous_map_snd [continuous_intros]: "continuous_map (prod_topology X Y) Y snd" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 207 | using continuous_map_pairwise [of "prod_topology X Y" X Y id] | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 208 | by (simp add: continuous_map_pairwise) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 209 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 210 | lemma continuous_map_fst_of [continuous_intros]: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 211 | "continuous_map Z (prod_topology X Y) f \<Longrightarrow> continuous_map Z X (fst \<circ> f)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 212 | by (simp add: continuous_map_pairwise) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 213 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 214 | lemma continuous_map_snd_of [continuous_intros]: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 215 | "continuous_map Z (prod_topology X Y) f \<Longrightarrow> continuous_map Z Y (snd \<circ> f)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 216 | by (simp add: continuous_map_pairwise) | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 217 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 218 | lemma continuous_map_prod_fst: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 219 | "i \<in> I \<Longrightarrow> continuous_map (prod_topology (product_topology (\<lambda>i. Y) I) X) Y (\<lambda>x. fst x i)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 220 | using continuous_map_componentwise_UNIV continuous_map_fst by fastforce | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 221 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 222 | lemma continuous_map_prod_snd: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 223 | "i \<in> I \<Longrightarrow> continuous_map (prod_topology X (product_topology (\<lambda>i. Y) I)) Y (\<lambda>x. snd x i)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 224 | using continuous_map_componentwise_UNIV continuous_map_snd by fastforce | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 225 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 226 | 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)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 227 | by simp | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 228 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 229 | lemma continuous_map_if [continuous_intros]: "\<lbrakk>P \<Longrightarrow> continuous_map X Y f; ~P \<Longrightarrow> continuous_map X Y g\<rbrakk> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 230 | \<Longrightarrow> continuous_map X Y (\<lambda>x. if P then f x else g x)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 231 | by simp | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 232 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 233 | lemma continuous_map_subtopology_fst [continuous_intros]: "continuous_map (subtopology (prod_topology X Y) Z) X fst" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 234 | using continuous_map_from_subtopology continuous_map_fst by force | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 235 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 236 | lemma continuous_map_subtopology_snd [continuous_intros]: "continuous_map (subtopology (prod_topology X Y) Z) Y snd" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 237 | using continuous_map_from_subtopology continuous_map_snd by force | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 238 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 239 | lemma quotient_map_fst [simp]: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 240 |    "quotient_map(prod_topology X Y) X fst \<longleftrightarrow> (topspace Y = {} \<longrightarrow> topspace X = {})"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 241 | by (auto simp: continuous_open_quotient_map open_map_fst continuous_map_fst) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 242 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 243 | lemma quotient_map_snd [simp]: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 244 |    "quotient_map(prod_topology X Y) Y snd \<longleftrightarrow> (topspace X = {} \<longrightarrow> topspace Y = {})"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 245 | by (auto simp: continuous_open_quotient_map open_map_snd continuous_map_snd) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 246 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 247 | lemma retraction_map_fst: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 248 |    "retraction_map (prod_topology X Y) X fst \<longleftrightarrow> (topspace Y = {} \<longrightarrow> topspace X = {})"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 249 | proof (cases "topspace Y = {}")
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 250 | case True | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 251 | then show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 252 | using continuous_map_image_subset_topspace | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 253 | by (fastforce simp: retraction_map_def retraction_maps_def continuous_map_fst continuous_map_on_empty) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 254 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 255 | case False | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 256 | have "\<exists>g. continuous_map X (prod_topology X Y) g \<and> (\<forall>x\<in>topspace X. fst (g x) = x)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 257 | if y: "y \<in> topspace Y" for y | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 258 | by (rule_tac x="\<lambda>x. (x,y)" in exI) (auto simp: y continuous_map_paired) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 259 | with False have "retraction_map (prod_topology X Y) X fst" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 260 | by (fastforce simp: retraction_map_def retraction_maps_def continuous_map_fst) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 261 | with False show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 262 | by simp | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 263 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 264 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 265 | lemma retraction_map_snd: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 266 |    "retraction_map (prod_topology X Y) Y snd \<longleftrightarrow> (topspace X = {} \<longrightarrow> topspace Y = {})"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 267 | proof (cases "topspace X = {}")
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 268 | case True | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 269 | then show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 270 | using continuous_map_image_subset_topspace | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 271 | by (fastforce simp: retraction_map_def retraction_maps_def continuous_map_fst continuous_map_on_empty) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 272 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 273 | case False | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 274 | have "\<exists>g. continuous_map Y (prod_topology X Y) g \<and> (\<forall>y\<in>topspace Y. snd (g y) = y)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 275 | if x: "x \<in> topspace X" for x | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 276 | by (rule_tac x="\<lambda>y. (x,y)" in exI) (auto simp: x continuous_map_paired) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 277 | with False have "retraction_map (prod_topology X Y) Y snd" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 278 | by (fastforce simp: retraction_map_def retraction_maps_def continuous_map_snd) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 279 | with False show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 280 | by simp | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 281 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 282 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 283 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 284 | lemma continuous_map_of_fst: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 285 |    "continuous_map (prod_topology X Y) Z (f \<circ> fst) \<longleftrightarrow> topspace Y = {} \<or> continuous_map X Z f"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 286 | proof (cases "topspace Y = {}")
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 287 | case True | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 288 | then show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 289 | by (simp add: continuous_map_on_empty) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 290 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 291 | case False | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 292 | then show ?thesis | 
| 71633 | 293 | by (simp add: continuous_compose_quotient_map_eq) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 294 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 295 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 296 | lemma continuous_map_of_snd: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 297 |    "continuous_map (prod_topology X Y) Z (f \<circ> snd) \<longleftrightarrow> topspace X = {} \<or> continuous_map Y Z f"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 298 | proof (cases "topspace X = {}")
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 299 | case True | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 300 | then show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 301 | by (simp add: continuous_map_on_empty) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 302 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 303 | case False | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 304 | then show ?thesis | 
| 71633 | 305 | by (simp add: continuous_compose_quotient_map_eq) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 306 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 307 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 308 | lemma continuous_map_prod_top: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 309 | "continuous_map (prod_topology X Y) (prod_topology X' Y') (\<lambda>(x,y). (f x, g y)) \<longleftrightarrow> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 310 |     topspace (prod_topology X Y) = {} \<or> continuous_map X X' f \<and> continuous_map Y Y' g"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 311 | proof (cases "topspace (prod_topology X Y) = {}")
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 312 | case True | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 313 | then show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 314 | by (simp add: continuous_map_on_empty) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 315 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 316 | case False | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 317 | then show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 318 | by (simp add: continuous_map_paired case_prod_unfold continuous_map_of_fst [unfolded o_def] continuous_map_of_snd [unfolded o_def]) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 319 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 320 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 321 | lemma in_prod_topology_closure_of: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 322 | assumes "z \<in> (prod_topology X Y) closure_of S" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 323 | shows "fst z \<in> X closure_of (fst ` S)" "snd z \<in> Y closure_of (snd ` S)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 324 | using assms continuous_map_eq_image_closure_subset continuous_map_fst apply fastforce | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 325 | using assms continuous_map_eq_image_closure_subset continuous_map_snd apply fastforce | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 326 | done | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 327 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 328 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 329 | proposition compact_space_prod_topology: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 330 |    "compact_space(prod_topology X Y) \<longleftrightarrow> topspace(prod_topology X Y) = {} \<or> compact_space X \<and> compact_space Y"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 331 | proof (cases "topspace(prod_topology X Y) = {}")
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 332 | case True | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 333 | then show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 334 | using compact_space_topspace_empty by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 335 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 336 | case False | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 337 |   then have non_mt: "topspace X \<noteq> {}" "topspace Y \<noteq> {}"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 338 | by auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 339 | have "compact_space X" "compact_space Y" if "compact_space(prod_topology X Y)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 340 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 341 | have "compactin X (fst ` (topspace X \<times> topspace Y))" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 342 | by (metis compact_space_def continuous_map_fst image_compactin that topspace_prod_topology) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 343 | moreover | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 344 | have "compactin Y (snd ` (topspace X \<times> topspace Y))" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 345 | by (metis compact_space_def continuous_map_snd image_compactin that topspace_prod_topology) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 346 | ultimately show "compact_space X" "compact_space Y" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 347 | by (simp_all add: non_mt compact_space_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 348 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 349 | moreover | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 350 | define \<X> where "\<X> \<equiv> (\<lambda>V. topspace X \<times> V) ` Collect (openin Y)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 351 | define \<Y> where "\<Y> \<equiv> (\<lambda>U. U \<times> topspace Y) ` Collect (openin X)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 352 | have "compact_space(prod_topology X Y)" if "compact_space X" "compact_space Y" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 353 | proof (rule Alexander_subbase_alt) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 354 | show toptop: "topspace X \<times> topspace Y \<subseteq> \<Union>(\<X> \<union> \<Y>)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 355 | unfolding \<X>_def \<Y>_def by auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 356 |     fix \<C> :: "('a \<times> 'b) set set"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 357 | assume \<C>: "\<C> \<subseteq> \<X> \<union> \<Y>" "topspace X \<times> topspace Y \<subseteq> \<Union>\<C>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 358 | then obtain \<X>' \<Y>' where XY: "\<X>' \<subseteq> \<X>" "\<Y>' \<subseteq> \<Y>" and \<C>eq: "\<C> = \<X>' \<union> \<Y>'" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 359 | using subset_UnE by metis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 360 | then have sub: "topspace X \<times> topspace Y \<subseteq> \<Union>(\<X>' \<union> \<Y>')" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 361 | using \<C> by simp | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 362 | obtain \<U> \<V> where \<U>: "\<And>U. U \<in> \<U> \<Longrightarrow> openin X U" "\<Y>' = (\<lambda>U. U \<times> topspace Y) ` \<U>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 363 | and \<V>: "\<And>V. V \<in> \<V> \<Longrightarrow> openin Y V" "\<X>' = (\<lambda>V. topspace X \<times> V) ` \<V>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 364 | using XY by (clarsimp simp add: \<X>_def \<Y>_def subset_image_iff) (force simp add: subset_iff) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 365 | have "\<exists>\<D>. finite \<D> \<and> \<D> \<subseteq> \<X>' \<union> \<Y>' \<and> topspace X \<times> topspace Y \<subseteq> \<Union> \<D>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 366 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 367 | have "topspace X \<subseteq> \<Union>\<U> \<or> topspace Y \<subseteq> \<Union>\<V>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 368 | using \<U> \<V> \<C> \<C>eq by auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 369 | then have *: "\<exists>\<D>. finite \<D> \<and> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 370 | (\<forall>x \<in> \<D>. x \<in> (\<lambda>V. topspace X \<times> V) ` \<V> \<or> x \<in> (\<lambda>U. U \<times> topspace Y) ` \<U>) \<and> | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 371 | (topspace X \<times> topspace Y \<subseteq> \<Union>\<D>)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 372 | proof | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 373 | assume "topspace X \<subseteq> \<Union>\<U>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 374 | with \<open>compact_space X\<close> \<U> obtain \<F> where "finite \<F>" "\<F> \<subseteq> \<U>" "topspace X \<subseteq> \<Union>\<F>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 375 | by (meson compact_space_alt) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 376 | with that show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 377 | by (rule_tac x="(\<lambda>D. D \<times> topspace Y) ` \<F>" in exI) auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 378 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 379 | assume "topspace Y \<subseteq> \<Union>\<V>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 380 | with \<open>compact_space Y\<close> \<V> obtain \<F> where "finite \<F>" "\<F> \<subseteq> \<V>" "topspace Y \<subseteq> \<Union>\<F>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 381 | by (meson compact_space_alt) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 382 | with that show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 383 | by (rule_tac x="(\<lambda>C. topspace X \<times> C) ` \<F>" in exI) auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 384 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 385 | then show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 386 | using that \<U> \<V> by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 387 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 388 | then show "\<exists>\<D>. finite \<D> \<and> \<D> \<subseteq> \<C> \<and> topspace X \<times> topspace Y \<subseteq> \<Union> \<D>" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 389 | using \<C> \<C>eq by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 390 | next | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 391 | have "(finite intersection_of (\<lambda>x. x \<in> \<X> \<or> x \<in> \<Y>) relative_to topspace X \<times> topspace Y) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 392 | = (\<lambda>U. \<exists>S T. U = S \<times> T \<and> openin X S \<and> openin Y T)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 393 | (is "?lhs = ?rhs") | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 394 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 395 | have "?rhs U" if "?lhs U" for U | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 396 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 397 |         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)}"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 398 | if "finite T" "T \<subseteq> \<X> \<union> \<Y>" for T | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 399 | using that | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 400 | proof induction | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 401 | case (insert B \<B>) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 402 | then show ?case | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 403 | unfolding \<X>_def \<Y>_def | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 404 | apply (simp add: Int_ac subset_eq image_def) | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 405 | apply (metis (no_types) openin_Int openin_topspace Times_Int_Times) | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 406 | done | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 407 | qed auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 408 | then show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 409 | using that | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 410 | by (auto simp: subset_eq elim!: relative_toE intersection_ofE) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 411 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 412 | moreover | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 413 | have "?lhs Z" if Z: "?rhs Z" for Z | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 414 | proof - | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 415 | obtain U V where "Z = U \<times> V" "openin X U" "openin Y V" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 416 | using Z by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 417 | then have UV: "U \<times> V = (topspace X \<times> topspace Y) \<inter> (U \<times> V)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 418 | by (simp add: Sigma_mono inf_absorb2 openin_subset) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 419 | moreover | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 420 | have "?lhs ((topspace X \<times> topspace Y) \<inter> (U \<times> V))" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 421 | proof (rule relative_to_inc) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 422 | show "(finite intersection_of (\<lambda>x. x \<in> \<X> \<or> x \<in> \<Y>)) (U \<times> V)" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 423 | apply (simp add: intersection_of_def \<X>_def \<Y>_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 424 |             apply (rule_tac x="{(U \<times> topspace Y),(topspace X \<times> V)}" in exI)
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 425 | using \<open>openin X U\<close> \<open>openin Y V\<close> openin_subset UV apply (fastforce simp add:) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 426 | done | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 427 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 428 | ultimately show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 429 | using \<open>Z = U \<times> V\<close> by auto | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 430 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 431 | ultimately show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 432 | by meson | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 433 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 434 | then show "topology (arbitrary union_of (finite intersection_of (\<lambda>x. x \<in> \<X> \<union> \<Y>) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 435 | relative_to (topspace X \<times> topspace Y))) = | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 436 | prod_topology X Y" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 437 | by (simp add: prod_topology_def) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 438 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 439 | ultimately show ?thesis | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 440 | using False by blast | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 441 | qed | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 442 | |
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 443 | lemma compactin_Times: | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 444 |    "compactin (prod_topology X Y) (S \<times> T) \<longleftrightarrow> S = {} \<or> T = {} \<or> compactin X S \<and> compactin Y T"
 | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 445 | by (auto simp: compactin_subspace subtopology_Times compact_space_prod_topology) | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 446 | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 447 | subsection\<open>Homeomorphic maps\<close> | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 448 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 449 | lemma homeomorphic_maps_prod: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 450 | "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> | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 451 |         topspace(prod_topology X Y) = {} \<and>
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 452 |         topspace(prod_topology X' Y') = {} \<or>
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 453 | homeomorphic_maps X X' f f' \<and> | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 454 | homeomorphic_maps Y Y' g g'" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 455 | unfolding homeomorphic_maps_def continuous_map_prod_top | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 456 | by (auto simp: continuous_map_def homeomorphic_maps_def continuous_map_prod_top) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 457 | |
| 70178 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69994diff
changeset | 458 | lemma homeomorphic_maps_swap: | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69994diff
changeset | 459 | "homeomorphic_maps (prod_topology X Y) (prod_topology Y X) | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69994diff
changeset | 460 | (\<lambda>(x,y). (y,x)) (\<lambda>(y,x). (x,y))" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69994diff
changeset | 461 | by (auto simp: homeomorphic_maps_def case_prod_unfold continuous_map_fst continuous_map_pairedI continuous_map_snd) | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69994diff
changeset | 462 | |
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69994diff
changeset | 463 | lemma homeomorphic_map_swap: | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69994diff
changeset | 464 | "homeomorphic_map (prod_topology X Y) (prod_topology Y X) (\<lambda>(x,y). (y,x))" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69994diff
changeset | 465 | using homeomorphic_map_maps homeomorphic_maps_swap by metis | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69994diff
changeset | 466 | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 467 | lemma embedding_map_graph: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 468 | "embedding_map X (prod_topology X Y) (\<lambda>x. (x, f x)) \<longleftrightarrow> continuous_map X Y f" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 469 | (is "?lhs = ?rhs") | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 470 | proof | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 471 | assume L: ?lhs | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 472 | have "snd \<circ> (\<lambda>x. (x, f x)) = f" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 473 | by force | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 474 | moreover have "continuous_map X Y (snd \<circ> (\<lambda>x. (x, f x)))" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 475 | using L | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 476 | unfolding embedding_map_def | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 477 | by (meson continuous_map_in_subtopology continuous_map_snd_of homeomorphic_imp_continuous_map) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 478 | ultimately show ?rhs | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 479 | by simp | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 480 | next | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 481 | assume R: ?rhs | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 482 | then show ?lhs | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 483 | unfolding homeomorphic_map_maps embedding_map_def homeomorphic_maps_def | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 484 | by (rule_tac x=fst in exI) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 485 | (auto simp: continuous_map_in_subtopology continuous_map_paired continuous_map_from_subtopology | 
| 70178 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
69994diff
changeset | 486 | continuous_map_fst) | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 487 | qed | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 488 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 489 | lemma homeomorphic_space_prod_topology: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 490 | "\<lbrakk>X homeomorphic_space X''; Y homeomorphic_space Y'\<rbrakk> | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 491 | \<Longrightarrow> prod_topology X Y homeomorphic_space prod_topology X'' Y'" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 492 | using homeomorphic_maps_prod unfolding homeomorphic_space_def by blast | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 493 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 494 | lemma prod_topology_homeomorphic_space_left: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 495 |    "topspace Y = {b} \<Longrightarrow> prod_topology X Y homeomorphic_space X"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 496 | unfolding homeomorphic_space_def | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 497 | by (rule_tac x=fst in exI) (simp add: homeomorphic_map_def inj_on_def flip: homeomorphic_map_maps) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 498 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 499 | lemma prod_topology_homeomorphic_space_right: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 500 |    "topspace X = {a} \<Longrightarrow> prod_topology X Y homeomorphic_space Y"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 501 | unfolding homeomorphic_space_def | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 502 | by (rule_tac x=snd in exI) (simp add: homeomorphic_map_def inj_on_def flip: homeomorphic_map_maps) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 503 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 504 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 505 | lemma homeomorphic_space_prod_topology_sing1: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 506 |      "b \<in> topspace Y \<Longrightarrow> X homeomorphic_space (prod_topology X (subtopology Y {b}))"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 507 | by (metis empty_subsetI homeomorphic_space_sym inf.absorb_iff2 insert_subset prod_topology_homeomorphic_space_left topspace_subtopology) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 508 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 509 | lemma homeomorphic_space_prod_topology_sing2: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 510 |      "a \<in> topspace X \<Longrightarrow> Y homeomorphic_space (prod_topology (subtopology X {a}) Y)"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 511 | by (metis empty_subsetI homeomorphic_space_sym inf.absorb_iff2 insert_subset prod_topology_homeomorphic_space_right topspace_subtopology) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 512 | |
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 513 | lemma topological_property_of_prod_component: | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 514 | assumes major: "P(prod_topology X Y)" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 515 |     and X: "\<And>x. \<lbrakk>x \<in> topspace X; P(prod_topology X Y)\<rbrakk> \<Longrightarrow> P(subtopology (prod_topology X Y) ({x} \<times> topspace Y))"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 516 |     and Y: "\<And>y. \<lbrakk>y \<in> topspace Y; P(prod_topology X Y)\<rbrakk> \<Longrightarrow> P(subtopology (prod_topology X Y) (topspace X \<times> {y}))"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 517 | and PQ: "\<And>X X'. X homeomorphic_space X' \<Longrightarrow> (P X \<longleftrightarrow> Q X')" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 518 | and PR: "\<And>X X'. X homeomorphic_space X' \<Longrightarrow> (P X \<longleftrightarrow> R X')" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 519 |   shows "topspace(prod_topology X Y) = {} \<or> Q X \<and> R Y"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 520 | proof - | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 521 |   have "Q X \<and> R Y" if "topspace(prod_topology X Y) \<noteq> {}"
 | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 522 | proof - | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 523 | from that obtain a b where a: "a \<in> topspace X" and b: "b \<in> topspace Y" | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 524 | by force | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 525 | show ?thesis | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 526 | using X [OF a major] and Y [OF b major] homeomorphic_space_prod_topology_sing1 [OF b, of X] homeomorphic_space_prod_topology_sing2 [OF a, of Y] | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 527 | by (simp add: subtopology_Times) (meson PQ PR homeomorphic_space_prod_topology_sing2 homeomorphic_space_sym) | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 528 | qed | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 529 | then show ?thesis by metis | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 530 | qed | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 531 | |
| 69945 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 532 | lemma limitin_pairwise: | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 533 | "limitin (prod_topology X Y) f l F \<longleftrightarrow> limitin X (fst \<circ> f) (fst l) F \<and> limitin Y (snd \<circ> f) (snd l) F" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 534 | (is "?lhs = ?rhs") | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 535 | proof | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 536 | assume ?lhs | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 537 | then obtain a b where ev: "\<And>U. \<lbrakk>(a,b) \<in> U; openin (prod_topology X Y) U\<rbrakk> \<Longrightarrow> \<forall>\<^sub>F x in F. f x \<in> U" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 538 | and a: "a \<in> topspace X" and b: "b \<in> topspace Y" and l: "l = (a,b)" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 539 | by (auto simp: limitin_def) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 540 | moreover have "\<forall>\<^sub>F x in F. fst (f x) \<in> U" if "openin X U" "a \<in> U" for U | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 541 | proof - | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 542 | have "\<forall>\<^sub>F c in F. f c \<in> U \<times> topspace Y" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 543 | using b that ev [of "U \<times> topspace Y"] by (auto simp: openin_prod_topology_alt) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 544 | then show ?thesis | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 545 | by (rule eventually_mono) (metis (mono_tags, lifting) SigmaE2 prod.collapse) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 546 | qed | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 547 | moreover have "\<forall>\<^sub>F x in F. snd (f x) \<in> U" if "openin Y U" "b \<in> U" for U | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 548 | proof - | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 549 | have "\<forall>\<^sub>F c in F. f c \<in> topspace X \<times> U" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 550 | using a that ev [of "topspace X \<times> U"] by (auto simp: openin_prod_topology_alt) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 551 | then show ?thesis | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 552 | by (rule eventually_mono) (metis (mono_tags, lifting) SigmaE2 prod.collapse) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 553 | qed | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 554 | ultimately show ?rhs | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 555 | by (simp add: limitin_def) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 556 | next | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 557 | have "limitin (prod_topology X Y) f (a,b) F" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 558 | if "limitin X (fst \<circ> f) a F" "limitin Y (snd \<circ> f) b F" for a b | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 559 | using that | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 560 | proof (clarsimp simp: limitin_def) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 561 |     fix Z :: "('a \<times> 'b) set"
 | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 562 | assume a: "a \<in> topspace X" "\<forall>U. openin X U \<and> a \<in> U \<longrightarrow> (\<forall>\<^sub>F x in F. fst (f x) \<in> U)" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 563 | and b: "b \<in> topspace Y" "\<forall>U. openin Y U \<and> b \<in> U \<longrightarrow> (\<forall>\<^sub>F x in F. snd (f x) \<in> U)" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 564 | and Z: "openin (prod_topology X Y) Z" "(a, b) \<in> Z" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 565 | then obtain U V where "openin X U" "openin Y V" "a \<in> U" "b \<in> V" "U \<times> V \<subseteq> Z" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 566 | using Z by (force simp: openin_prod_topology_alt) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 567 | then have "\<forall>\<^sub>F x in F. fst (f x) \<in> U" "\<forall>\<^sub>F x in F. snd (f x) \<in> V" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 568 | by (simp_all add: a b) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 569 | then show "\<forall>\<^sub>F x in F. f x \<in> Z" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 570 | by (rule eventually_elim2) (use \<open>U \<times> V \<subseteq> Z\<close> subsetD in auto) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 571 | qed | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 572 | then show "?rhs \<Longrightarrow> ?lhs" | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 573 | by (metis prod.collapse) | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 574 | qed | 
| 
35ba13ac6e5c
New abstract topological material
 paulson <lp15@cam.ac.uk> parents: 
69939diff
changeset | 575 | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 576 | end | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 577 |