| author | wenzelm | 
| Sat, 24 Feb 2024 11:27:04 +0100 | |
| changeset 79718 | fba02e281b44 | 
| parent 77226 | 69956724ad4f | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1 | theory Isolated | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2 | imports "HOL-Analysis.Elementary_Metric_Spaces" | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3 | |
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4 | begin | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5 | |
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 6 | subsection \<open>Isolate and discrete\<close> | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 7 | |
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 8 | definition (in topological_space) isolated_in:: "'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixr "isolated'_in" 60) | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 9 |   where "x isolated_in S \<longleftrightarrow> (x\<in>S \<and> (\<exists>T. open T \<and> T \<inter> S = {x}))"
 | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 10 | |
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 11 | definition (in topological_space) discrete:: "'a set \<Rightarrow> bool" | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 12 | where "discrete S \<longleftrightarrow> (\<forall>x\<in>S. x isolated_in S)" | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 13 | |
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 14 | definition (in metric_space) uniform_discrete :: "'a set \<Rightarrow> bool" where | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 15 | "uniform_discrete S \<longleftrightarrow> (\<exists>e>0. \<forall>x\<in>S. \<forall>y\<in>S. dist x y < e \<longrightarrow> x = y)" | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 16 | |
| 77226 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 17 | lemma discreteI: "(\<And>x. x \<in> X \<Longrightarrow> x isolated_in X ) \<Longrightarrow> discrete X" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 18 | unfolding discrete_def by auto | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 19 | |
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 20 | lemma discreteD: "discrete X \<Longrightarrow> x \<in> X \<Longrightarrow> x isolated_in X " | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 21 | unfolding discrete_def by auto | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 22 | |
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 23 | lemma uniformI1: | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 24 | assumes "e>0" "\<And>x y. \<lbrakk>x\<in>S;y\<in>S;dist x y<e\<rbrakk> \<Longrightarrow> x =y " | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 25 | shows "uniform_discrete S" | 
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 26 | unfolding uniform_discrete_def using assms by auto | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 27 | |
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 28 | lemma uniformI2: | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 29 | assumes "e>0" "\<And>x y. \<lbrakk>x\<in>S;y\<in>S;x\<noteq>y\<rbrakk> \<Longrightarrow> dist x y\<ge>e " | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 30 | shows "uniform_discrete S" | 
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 31 | unfolding uniform_discrete_def using assms not_less by blast | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 32 | |
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 33 | lemma isolated_in_islimpt_iff:"(x isolated_in S) \<longleftrightarrow> (\<not> (x islimpt S) \<and> x\<in>S)" | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 34 | unfolding isolated_in_def islimpt_def by auto | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 35 | |
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 36 | lemma isolated_in_dist_Ex_iff: | 
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 37 | fixes x::"'a::metric_space" | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 38 | shows "x isolated_in S \<longleftrightarrow> (x\<in>S \<and> (\<exists>e>0. \<forall>y\<in>S. dist x y < e \<longrightarrow> y=x))" | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 39 | unfolding isolated_in_islimpt_iff islimpt_approachable by (metis dist_commute) | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 40 | |
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 41 | lemma discrete_empty[simp]: "discrete {}"
 | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 42 | unfolding discrete_def by auto | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 43 | |
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 44 | lemma uniform_discrete_empty[simp]: "uniform_discrete {}"
 | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 45 | unfolding uniform_discrete_def by (simp add: gt_ex) | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 46 | |
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 47 | lemma isolated_in_insert: | 
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 48 | fixes x :: "'a::t1_space" | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 49 | shows "x isolated_in (insert a S) \<longleftrightarrow> x isolated_in S \<or> (x=a \<and> \<not> (x islimpt S))" | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 50 | by (meson insert_iff islimpt_insert isolated_in_islimpt_iff) | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 51 | |
| 77226 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 52 | lemma isolated_inI: | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 53 |   assumes "x\<in>S" "open T" "T \<inter> S = {x}"
 | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 54 | shows "x isolated_in S" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 55 | using assms unfolding isolated_in_def by auto | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 56 | |
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 57 | lemma isolated_inE: | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 58 | assumes "x isolated_in S" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 59 |   obtains T where "x \<in> S" "open T" "T \<inter> S = {x}"
 | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 60 | using assms that unfolding isolated_in_def by force | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 61 | |
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 62 | lemma isolated_inE_dist: | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 63 | assumes "x isolated_in S" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 64 | obtains d where "d > 0" "\<And>y. y \<in> S \<Longrightarrow> dist x y < d \<Longrightarrow> y = x" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 65 | by (meson assms isolated_in_dist_Ex_iff) | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 66 | |
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 67 | lemma isolated_in_altdef: | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 68 | "x isolated_in S \<longleftrightarrow> (x\<in>S \<and> eventually (\<lambda>y. y \<notin> S) (at x))" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 69 | proof | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 70 | assume "x isolated_in S" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 71 | from isolated_inE[OF this] | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 72 |   obtain T where "x \<in> S" and T:"open T" "T \<inter> S = {x}"
 | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 73 | by metis | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 74 | have "\<forall>\<^sub>F y in nhds x. y \<in> T" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 75 | apply (rule eventually_nhds_in_open) | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 76 | using T by auto | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 77 |   then have  "eventually (\<lambda>y. y \<in> T - {x}) (at x)"
 | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 78 | unfolding eventually_at_filter by eventually_elim auto | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 79 | then have "eventually (\<lambda>y. y \<notin> S) (at x)" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 80 | by eventually_elim (use T in auto) | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 81 | then show " x \<in> S \<and> (\<forall>\<^sub>F y in at x. y \<notin> S)" using \<open>x \<in> S\<close> by auto | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 82 | next | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 83 | assume "x \<in> S \<and> (\<forall>\<^sub>F y in at x. y \<notin> S)" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 84 | then have "\<forall>\<^sub>F y in at x. y \<notin> S" "x\<in>S" by auto | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 85 | from this(1) have "eventually (\<lambda>y. y \<notin> S \<or> y = x) (nhds x)" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 86 | unfolding eventually_at_filter by eventually_elim auto | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 87 | then obtain T where T:"open T" "x \<in> T" "(\<forall>y\<in>T. y \<notin> S \<or> y = x)" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 88 | unfolding eventually_nhds by auto | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 89 |   with \<open>x \<in> S\<close> have "T \<inter> S = {x}"  
 | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 90 | by fastforce | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 91 | with \<open>x\<in>S\<close> \<open>open T\<close> | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 92 | show "x isolated_in S" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 93 | unfolding isolated_in_def by auto | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 94 | qed | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 95 | |
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 96 | lemma discrete_altdef: | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 97 | "discrete S \<longleftrightarrow> (\<forall>x\<in>S. \<forall>\<^sub>F y in at x. y \<notin> S)" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 98 | unfolding discrete_def isolated_in_altdef by auto | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 99 | |
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 100 | (* | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 101 | TODO. | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 102 | Other than | 
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 103 | |
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 104 | uniform_discrete S \<longrightarrow> discrete S | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 105 | uniform_discrete S \<longrightarrow> closed S | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 106 | |
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 107 | , we should be able to prove | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 108 | |
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 109 | discrete S \<and> closed S \<longrightarrow> uniform_discrete S | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 110 | |
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 111 | but the proof (based on Tietze Extension Theorem) seems not very trivial to me. Informal proofs can be found in | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 112 | |
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 113 | http://topology.auburn.edu/tp/reprints/v30/tp30120.pdf | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 114 | http://msp.org/pjm/1959/9-2/pjm-v9-n2-p19-s.pdf | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 115 | *) | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 116 | |
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 117 | lemma uniform_discrete_imp_closed: | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 118 | "uniform_discrete S \<Longrightarrow> closed S" | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 119 | by (meson discrete_imp_closed uniform_discrete_def) | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 120 | |
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 121 | lemma uniform_discrete_imp_discrete: | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 122 | "uniform_discrete S \<Longrightarrow> discrete S" | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 123 | by (metis discrete_def isolated_in_dist_Ex_iff uniform_discrete_def) | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 124 | |
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 125 | lemma isolated_in_subset:"x isolated_in S \<Longrightarrow> T \<subseteq> S \<Longrightarrow> x\<in>T \<Longrightarrow> x isolated_in T" | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 126 | unfolding isolated_in_def by fastforce | 
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 127 | |
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 128 | lemma discrete_subset[elim]: "discrete S \<Longrightarrow> T \<subseteq> S \<Longrightarrow> discrete T" | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 129 | unfolding discrete_def using islimpt_subset isolated_in_islimpt_iff by blast | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 130 | |
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 131 | lemma uniform_discrete_subset[elim]: "uniform_discrete S \<Longrightarrow> T \<subseteq> S \<Longrightarrow> uniform_discrete T" | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 132 | by (meson subsetD uniform_discrete_def) | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 133 | |
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 134 | lemma continuous_on_discrete: "discrete S \<Longrightarrow> continuous_on S f" | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 135 | unfolding continuous_on_topological by (metis discrete_def islimptI isolated_in_islimpt_iff) | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 136 | |
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 137 | lemma uniform_discrete_insert: "uniform_discrete (insert a S) \<longleftrightarrow> uniform_discrete S" | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 138 | proof | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 139 | assume asm:"uniform_discrete S" | 
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 140 | let ?thesis = "uniform_discrete (insert a S)" | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 141 | have ?thesis when "a\<in>S" using that asm by (simp add: insert_absorb) | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 142 |   moreover have ?thesis when "S={}" using that asm by (simp add: uniform_discrete_def)
 | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 143 |   moreover have ?thesis when "a\<notin>S" "S\<noteq>{}"
 | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 144 | proof - | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 145 | obtain e1 where "e1>0" and e1_dist:"\<forall>x\<in>S. \<forall>y\<in>S. dist y x < e1 \<longrightarrow> y = x" | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 146 | using asm unfolding uniform_discrete_def by auto | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 147 |     define e2 where "e2 \<equiv> min (setdist {a} S) e1"
 | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 148 | have "closed S" using asm uniform_discrete_imp_closed by auto | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 149 | then have "e2>0" | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 150 | by (smt (verit) \<open>0 < e1\<close> e2_def infdist_eq_setdist infdist_pos_not_in_closed that) | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 151 | moreover have "x = y" if "x\<in>insert a S" "y\<in>insert a S" "dist x y < e2" for x y | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 152 | proof (cases "x=a \<or> y=a") | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 153 | case True then show ?thesis | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 154 | by (smt (verit, best) dist_commute e2_def infdist_eq_setdist infdist_le insertE that) | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 155 | next | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 156 | case False then show ?thesis | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 157 | using e1_dist e2_def that by force | 
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 158 | qed | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 159 | ultimately show ?thesis unfolding uniform_discrete_def by meson | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 160 | qed | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 161 | ultimately show ?thesis by auto | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 162 | qed (simp add: subset_insertI uniform_discrete_subset) | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 163 | |
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 164 | lemma discrete_compact_finite_iff: | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 165 | fixes S :: "'a::t1_space set" | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 166 | shows "discrete S \<and> compact S \<longleftrightarrow> finite S" | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 167 | proof | 
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 168 | assume "finite S" | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 169 | then have "compact S" using finite_imp_compact by auto | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 170 | moreover have "discrete S" | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 171 | unfolding discrete_def using isolated_in_islimpt_iff islimpt_finite[OF \<open>finite S\<close>] by auto | 
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 172 | ultimately show "discrete S \<and> compact S" by auto | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 173 | next | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 174 | assume "discrete S \<and> compact S" | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 175 | then show "finite S" | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 176 | by (meson discrete_def Heine_Borel_imp_Bolzano_Weierstrass isolated_in_islimpt_iff order_refl) | 
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 177 | qed | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 178 | |
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 179 | lemma uniform_discrete_finite_iff: | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 180 | fixes S :: "'a::heine_borel set" | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 181 | shows "uniform_discrete S \<and> bounded S \<longleftrightarrow> finite S" | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 182 | proof | 
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 183 | assume "uniform_discrete S \<and> bounded S" | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 184 | then have "discrete S" "compact S" | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 185 | using uniform_discrete_imp_discrete uniform_discrete_imp_closed compact_eq_bounded_closed | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 186 | by auto | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 187 | then show "finite S" using discrete_compact_finite_iff by auto | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 188 | next | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 189 | assume asm:"finite S" | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 190 | let ?thesis = "uniform_discrete S \<and> bounded S" | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 191 |   have ?thesis when "S={}" using that by auto
 | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 192 |   moreover have ?thesis when "S\<noteq>{}"
 | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 193 | proof - | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 194 | have "\<forall>x. \<exists>d>0. \<forall>y\<in>S. y \<noteq> x \<longrightarrow> d \<le> dist x y" | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 195 | using finite_set_avoid[OF \<open>finite S\<close>] by auto | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 196 | then obtain f where f_pos:"f x>0" | 
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 197 | and f_dist: "\<forall>y\<in>S. y \<noteq> x \<longrightarrow> f x \<le> dist x y" | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 198 | if "x\<in>S" for x | 
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 199 | by metis | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 200 | define f_min where "f_min \<equiv> Min (f ` S)" | 
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 201 | have "f_min > 0" | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 202 | unfolding f_min_def | 
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 203 | by (simp add: asm f_pos that) | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 204 | moreover have "\<forall>x\<in>S. \<forall>y\<in>S. f_min > dist x y \<longrightarrow> x=y" | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 205 | using f_dist unfolding f_min_def | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 206 | by (metis Min_le asm finite_imageI imageI le_less_trans linorder_not_less) | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 207 | ultimately have "uniform_discrete S" | 
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 208 | unfolding uniform_discrete_def by auto | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 209 | moreover have "bounded S" using \<open>finite S\<close> by auto | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 210 | ultimately show ?thesis by auto | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 211 | qed | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 212 | ultimately show ?thesis by blast | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 213 | qed | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 214 | |
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 215 | lemma uniform_discrete_image_scale: | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 216 | assumes "uniform_discrete S" and dist:"\<forall>x\<in>S. \<forall>y\<in>S. dist x y = c * dist (f x) (f y)" | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 217 | shows "uniform_discrete (f ` S)" | 
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 218 | proof - | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 219 |   have ?thesis when "S={}" using that by auto
 | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 220 |   moreover have ?thesis when "S\<noteq>{}" "c\<le>0"
 | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 221 | proof - | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 222 |     obtain x1 where "x1\<in>S" using \<open>S\<noteq>{}\<close> by auto
 | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 223 |     have ?thesis when "S-{x1} = {}"
 | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 224 | using \<open>x1 \<in> S\<close> subset_antisym that uniform_discrete_insert by fastforce | 
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 225 |     moreover have ?thesis when "S-{x1} \<noteq> {}"
 | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 226 | proof - | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 227 |       obtain x2 where "x2\<in>S-{x1}" using \<open>S-{x1} \<noteq> {}\<close> by auto
 | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 228 | then have "x2\<in>S" "x1\<noteq>x2" by auto | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 229 | then have "dist x1 x2 > 0" by auto | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 230 | moreover have "dist x1 x2 = c * dist (f x1) (f x2)" | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 231 | by (simp add: \<open>x1 \<in> S\<close> \<open>x2 \<in> S\<close> dist) | 
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 232 | moreover have "dist (f x2) (f x2) \<ge> 0" by auto | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 233 | ultimately have False using \<open>c\<le>0\<close> by (simp add: zero_less_mult_iff) | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 234 | then show ?thesis by auto | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 235 | qed | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 236 | ultimately show ?thesis by auto | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 237 | qed | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 238 |   moreover have ?thesis when "S\<noteq>{}" "c>0"
 | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 239 | proof - | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 240 | obtain e1 where "e1>0" and e1_dist:"\<forall>x\<in>S. \<forall>y\<in>S. dist y x < e1 \<longrightarrow> y = x" | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 241 | using \<open>uniform_discrete S\<close> unfolding uniform_discrete_def by auto | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 242 | define e where "e \<equiv> e1/c" | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 243 | have "x1 = x2" when "x1 \<in> f ` S" "x2 \<in> f ` S" and d: "dist x1 x2 < e" for x1 x2 | 
| 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 244 | by (smt (verit) \<open>0 < c\<close> d dist divide_right_mono e1_dist e_def imageE nonzero_mult_div_cancel_left that) | 
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 245 | moreover have "e>0" using \<open>e1>0\<close> \<open>c>0\<close> unfolding e_def by auto | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 246 | ultimately show ?thesis unfolding uniform_discrete_def by meson | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 247 | qed | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 248 | ultimately show ?thesis by fastforce | 
| 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 249 | qed | 
| 77200 
8f2e6186408f
Some more new material and some tidying of existing proofs
 paulson <lp15@cam.ac.uk> parents: 
77102diff
changeset | 250 | |
| 77226 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 251 | definition sparse :: "real \<Rightarrow> 'a :: metric_space set \<Rightarrow> bool" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 252 |   where "sparse \<epsilon> X \<longleftrightarrow> (\<forall>x\<in>X. \<forall>y\<in>X-{x}. dist x y > \<epsilon>)"
 | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 253 | |
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 254 | lemma sparse_empty [simp, intro]: "sparse \<epsilon> {}"
 | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 255 | by (auto simp: sparse_def) | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 256 | |
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 257 | lemma sparseI [intro?]: | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 258 | "(\<And>x y. x \<in> X \<Longrightarrow> y \<in> X \<Longrightarrow> x \<noteq> y \<Longrightarrow> dist x y > \<epsilon>) \<Longrightarrow> sparse \<epsilon> X" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 259 | unfolding sparse_def by auto | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 260 | |
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 261 | lemma sparseD: | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 262 | "sparse \<epsilon> X \<Longrightarrow> x \<in> X \<Longrightarrow> y \<in> X \<Longrightarrow> x \<noteq> y \<Longrightarrow> dist x y > \<epsilon>" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 263 | unfolding sparse_def by auto | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 264 | |
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 265 | lemma sparseD': | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 266 | "sparse \<epsilon> X \<Longrightarrow> x \<in> X \<Longrightarrow> y \<in> X \<Longrightarrow> dist x y \<le> \<epsilon> \<Longrightarrow> x = y" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 267 | unfolding sparse_def by force | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 268 | |
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 269 | lemma sparse_singleton [simp, intro]: "sparse \<epsilon> {x}"
 | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 270 | by (auto simp: sparse_def) | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 271 | |
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 272 | definition setdist_gt where "setdist_gt \<epsilon> X Y \<longleftrightarrow> (\<forall>x\<in>X. \<forall>y\<in>Y. dist x y > \<epsilon>)" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 273 | |
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 274 | lemma setdist_gt_empty [simp]: "setdist_gt \<epsilon> {} Y" "setdist_gt \<epsilon> X {}"
 | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 275 | by (auto simp: setdist_gt_def) | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 276 | |
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 277 | lemma setdist_gtI: "(\<And>x y. x \<in> X \<Longrightarrow> y \<in> Y \<Longrightarrow> dist x y > \<epsilon>) \<Longrightarrow> setdist_gt \<epsilon> X Y" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 278 | unfolding setdist_gt_def by auto | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 279 | |
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 280 | lemma setdist_gtD: "setdist_gt \<epsilon> X Y \<Longrightarrow> x \<in> X \<Longrightarrow> y \<in> Y \<Longrightarrow> dist x y > \<epsilon>" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 281 | unfolding setdist_gt_def by auto | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 282 | |
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 283 | lemma setdist_gt_setdist: "\<epsilon> < setdist A B \<Longrightarrow> setdist_gt \<epsilon> A B" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 284 | unfolding setdist_gt_def using setdist_le_dist by fastforce | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 285 | |
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 286 | lemma setdist_gt_mono: "setdist_gt \<epsilon>' A B \<Longrightarrow> \<epsilon> \<le> \<epsilon>' \<Longrightarrow> A' \<subseteq> A \<Longrightarrow> B' \<subseteq> B \<Longrightarrow> setdist_gt \<epsilon> A' B'" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 287 | by (force simp: setdist_gt_def) | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 288 | |
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 289 | lemma setdist_gt_Un_left: "setdist_gt \<epsilon> (A \<union> B) C \<longleftrightarrow> setdist_gt \<epsilon> A C \<and> setdist_gt \<epsilon> B C" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 290 | by (auto simp: setdist_gt_def) | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 291 | |
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 292 | lemma setdist_gt_Un_right: "setdist_gt \<epsilon> C (A \<union> B) \<longleftrightarrow> setdist_gt \<epsilon> C A \<and> setdist_gt \<epsilon> C B" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 293 | by (auto simp: setdist_gt_def) | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 294 | |
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 295 | lemma compact_closed_imp_eventually_setdist_gt_at_right_0: | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 296 |   assumes "compact A" "closed B" "A \<inter> B = {}"
 | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 297 | shows "eventually (\<lambda>\<epsilon>. setdist_gt \<epsilon> A B) (at_right 0)" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 298 | proof (cases "A = {} \<or> B = {}")
 | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 299 | case False | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 300 | hence "setdist A B > 0" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 301 | by (metis IntI assms empty_iff in_closed_iff_infdist_zero order_less_le setdist_attains_inf setdist_pos_le setdist_sym) | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 302 | hence "eventually (\<lambda>\<epsilon>. \<epsilon> < setdist A B) (at_right 0)" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 303 | using eventually_at_right_field by blast | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 304 | thus ?thesis | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 305 | by eventually_elim (auto intro: setdist_gt_setdist) | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 306 | qed auto | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 307 | |
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 308 | lemma setdist_gt_symI: "setdist_gt \<epsilon> A B \<Longrightarrow> setdist_gt \<epsilon> B A" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 309 | by (force simp: setdist_gt_def dist_commute) | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 310 | |
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 311 | lemma setdist_gt_sym: "setdist_gt \<epsilon> A B \<longleftrightarrow> setdist_gt \<epsilon> B A" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 312 | by (force simp: setdist_gt_def dist_commute) | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 313 | |
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 314 | lemma eventually_setdist_gt_at_right_0_mult_iff: | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 315 | assumes "c > 0" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 316 | shows "eventually (\<lambda>\<epsilon>. setdist_gt (c * \<epsilon>) A B) (at_right 0) \<longleftrightarrow> | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 317 | eventually (\<lambda>\<epsilon>. setdist_gt \<epsilon> A B) (at_right 0)" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 318 | proof - | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 319 | have "eventually (\<lambda>\<epsilon>. setdist_gt (c * \<epsilon>) A B) (at_right 0) \<longleftrightarrow> | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 320 | eventually (\<lambda>\<epsilon>. setdist_gt \<epsilon> A B) (filtermap ((*) c) (at_right 0))" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 321 | by (simp add: eventually_filtermap) | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 322 | also have "filtermap ((*) c) (at_right 0) = at_right 0" | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 323 | by (subst filtermap_times_pos_at_right) (use assms in auto) | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 324 | finally show ?thesis . | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 325 | qed | 
| 
69956724ad4f
More material for Analysis and Complex_Analysis
 paulson <lp15@cam.ac.uk> parents: 
77200diff
changeset | 326 | |
| 77102 
780161d4b55c
Moved in some material from the AFP entry Winding_number_eval
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 327 | end |