| author | immler | 
| Sat, 29 Sep 2018 16:30:44 -0400 | |
| changeset 69096 | 62a0d10386c1 | 
| parent 68077 | ee8c13ae81e9 | 
| child 69423 | 3922aa1df44e | 
| permissions | -rw-r--r-- | 
| 65039 | 1  | 
section \<open>Winding Numbers\<close>  | 
2  | 
||
3  | 
text\<open>By John Harrison et al. Ported from HOL Light by L C Paulson (2017)\<close>  | 
|
4  | 
||
5  | 
theory Winding_Numbers  | 
|
| 
66826
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
6  | 
imports Polytope Jordan_Curve Riemann_Mapping  | 
| 65039 | 7  | 
begin  | 
8  | 
||
| 
66826
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
9  | 
lemma simply_connected_inside_simple_path:  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
10  | 
fixes p :: "real \<Rightarrow> complex"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
11  | 
shows "simple_path p \<Longrightarrow> simply_connected(inside(path_image p))"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
12  | 
using Jordan_inside_outside connected_simple_path_image inside_simple_curve_imp_closed simply_connected_eq_frontier_properties  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
13  | 
by fastforce  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
14  | 
|
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
15  | 
lemma simply_connected_Int:  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
16  | 
fixes S :: "complex set"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
17  | 
assumes "open S" "open T" "simply_connected S" "simply_connected T" "connected (S \<inter> T)"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
18  | 
shows "simply_connected (S \<inter> T)"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
19  | 
using assms by (force simp: simply_connected_eq_winding_number_zero open_Int)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
20  | 
|
| 65039 | 21  | 
subsection\<open>Winding number for a triangle\<close>  | 
22  | 
||
23  | 
lemma wn_triangle1:  | 
|
24  | 
  assumes "0 \<in> interior(convex hull {a,b,c})"
 | 
|
25  | 
shows "~ (Im(a/b) \<le> 0 \<and> 0 \<le> Im(b/c))"  | 
|
26  | 
proof -  | 
|
27  | 
  { assume 0: "Im(a/b) \<le> 0" "0 \<le> Im(b/c)"
 | 
|
28  | 
    have "0 \<notin> interior (convex hull {a,b,c})"
 | 
|
29  | 
proof (cases "a=0 \<or> b=0 \<or> c=0")  | 
|
30  | 
case True then show ?thesis  | 
|
31  | 
by (auto simp: not_in_interior_convex_hull_3)  | 
|
32  | 
next  | 
|
33  | 
case False  | 
|
34  | 
then have "b \<noteq> 0" by blast  | 
|
35  | 
      { fix x y::complex and u::real
 | 
|
36  | 
assume eq_f': "Im x * Re b \<le> Im b * Re x" "Im y * Re b \<le> Im b * Re y" "0 \<le> u" "u \<le> 1"  | 
|
37  | 
then have "((1 - u) * Im x) * Re b \<le> Im b * ((1 - u) * Re x)"  | 
|
38  | 
by (simp add: mult_left_mono mult.assoc mult.left_commute [of "Im b"])  | 
|
39  | 
then have "((1 - u) * Im x + u * Im y) * Re b \<le> Im b * ((1 - u) * Re x + u * Re y)"  | 
|
40  | 
using eq_f' ordered_comm_semiring_class.comm_mult_left_mono  | 
|
41  | 
by (fastforce simp add: algebra_simps)  | 
|
42  | 
}  | 
|
43  | 
      with False 0 have "convex hull {a,b,c} \<le> {z. Im z * Re b \<le> Im b * Re z}"
 | 
|
44  | 
apply (simp add: Complex.Im_divide divide_simps complex_neq_0 [symmetric])  | 
|
45  | 
apply (simp add: algebra_simps)  | 
|
46  | 
apply (rule hull_minimal)  | 
|
47  | 
apply (auto simp: algebra_simps convex_alt)  | 
|
48  | 
done  | 
|
49  | 
      moreover have "0 \<notin> interior({z. Im z * Re b \<le> Im b * Re z})"
 | 
|
50  | 
proof  | 
|
51  | 
        assume "0 \<in> interior {z. Im z * Re b \<le> Im b * Re z}"
 | 
|
52  | 
        then obtain e where "e>0" and e: "ball 0 e \<subseteq> {z. Im z * Re b \<le> Im b * Re z}"
 | 
|
53  | 
by (meson mem_interior)  | 
|
| 
65064
 
a4abec71279a
Renamed ii to imaginary_unit in order to free up ii as a variable name.  Also replaced some legacy def commands
 
paulson <lp15@cam.ac.uk> 
parents: 
65039 
diff
changeset
 | 
54  | 
define z where "z \<equiv> - sgn (Im b) * (e/3) + sgn (Re b) * (e/3) * \<i>"  | 
| 65039 | 55  | 
have "z \<in> ball 0 e"  | 
| 66304 | 56  | 
using \<open>e>0\<close>  | 
| 65039 | 57  | 
apply (simp add: z_def dist_norm)  | 
58  | 
apply (rule le_less_trans [OF norm_triangle_ineq4])  | 
|
59  | 
apply (simp add: norm_mult abs_sgn_eq)  | 
|
60  | 
done  | 
|
61  | 
        then have "z \<in> {z. Im z * Re b \<le> Im b * Re z}"
 | 
|
62  | 
using e by blast  | 
|
63  | 
then show False  | 
|
| 66304 | 64  | 
using \<open>e>0\<close> \<open>b \<noteq> 0\<close>  | 
| 65039 | 65  | 
apply (simp add: z_def dist_norm sgn_if less_eq_real_def mult_less_0_iff complex.expand split: if_split_asm)  | 
66  | 
apply (auto simp: algebra_simps)  | 
|
67  | 
apply (meson less_asym less_trans mult_pos_pos neg_less_0_iff_less)  | 
|
68  | 
by (metis less_asym mult_pos_pos neg_less_0_iff_less)  | 
|
69  | 
qed  | 
|
70  | 
ultimately show ?thesis  | 
|
71  | 
using interior_mono by blast  | 
|
72  | 
qed  | 
|
73  | 
} with assms show ?thesis by blast  | 
|
74  | 
qed  | 
|
75  | 
||
76  | 
lemma wn_triangle2_0:  | 
|
77  | 
  assumes "0 \<in> interior(convex hull {a,b,c})"
 | 
|
78  | 
shows  | 
|
79  | 
"0 < Im((b - a) * cnj (b)) \<and>  | 
|
80  | 
0 < Im((c - b) * cnj (c)) \<and>  | 
|
81  | 
0 < Im((a - c) * cnj (a))  | 
|
82  | 
\<or>  | 
|
83  | 
Im((b - a) * cnj (b)) < 0 \<and>  | 
|
84  | 
0 < Im((b - c) * cnj (b)) \<and>  | 
|
85  | 
0 < Im((a - b) * cnj (a)) \<and>  | 
|
86  | 
0 < Im((c - a) * cnj (c))"  | 
|
87  | 
proof -  | 
|
88  | 
  have [simp]: "{b,c,a} = {a,b,c}" "{c,a,b} = {a,b,c}" by auto
 | 
|
89  | 
show ?thesis  | 
|
90  | 
using wn_triangle1 [OF assms] wn_triangle1 [of b c a] wn_triangle1 [of c a b] assms  | 
|
91  | 
by (auto simp: algebra_simps Im_complex_div_gt_0 Im_complex_div_lt_0 not_le not_less)  | 
|
92  | 
qed  | 
|
93  | 
||
94  | 
lemma wn_triangle2:  | 
|
95  | 
  assumes "z \<in> interior(convex hull {a,b,c})"
 | 
|
96  | 
shows "0 < Im((b - a) * cnj (b - z)) \<and>  | 
|
97  | 
0 < Im((c - b) * cnj (c - z)) \<and>  | 
|
98  | 
0 < Im((a - c) * cnj (a - z))  | 
|
99  | 
\<or>  | 
|
100  | 
Im((b - a) * cnj (b - z)) < 0 \<and>  | 
|
101  | 
0 < Im((b - c) * cnj (b - z)) \<and>  | 
|
102  | 
0 < Im((a - b) * cnj (a - z)) \<and>  | 
|
103  | 
0 < Im((c - a) * cnj (c - z))"  | 
|
104  | 
proof -  | 
|
105  | 
  have 0: "0 \<in> interior(convex hull {a-z, b-z, c-z})"
 | 
|
106  | 
    using assms convex_hull_translation [of "-z" "{a,b,c}"]
 | 
|
107  | 
interior_translation [of "-z"]  | 
|
108  | 
by simp  | 
|
109  | 
show ?thesis using wn_triangle2_0 [OF 0]  | 
|
110  | 
by simp  | 
|
111  | 
qed  | 
|
112  | 
||
113  | 
lemma wn_triangle3:  | 
|
114  | 
  assumes z: "z \<in> interior(convex hull {a,b,c})"
 | 
|
115  | 
and "0 < Im((b-a) * cnj (b-z))"  | 
|
116  | 
"0 < Im((c-b) * cnj (c-z))"  | 
|
117  | 
"0 < Im((a-c) * cnj (a-z))"  | 
|
118  | 
shows "winding_number (linepath a b +++ linepath b c +++ linepath c a) z = 1"  | 
|
119  | 
proof -  | 
|
120  | 
have znot[simp]: "z \<notin> closed_segment a b" "z \<notin> closed_segment b c" "z \<notin> closed_segment c a"  | 
|
121  | 
using z interior_of_triangle [of a b c]  | 
|
122  | 
by (auto simp: closed_segment_def)  | 
|
123  | 
have gt0: "0 < Re (winding_number (linepath a b +++ linepath b c +++ linepath c a) z)"  | 
|
124  | 
using assms  | 
|
125  | 
by (simp add: winding_number_linepath_pos_lt path_image_join winding_number_join_pos_combined)  | 
|
126  | 
have lt2: "Re (winding_number (linepath a b +++ linepath b c +++ linepath c a) z) < 2"  | 
|
127  | 
using winding_number_lt_half_linepath [of _ a b]  | 
|
128  | 
using winding_number_lt_half_linepath [of _ b c]  | 
|
129  | 
using winding_number_lt_half_linepath [of _ c a] znot  | 
|
130  | 
apply (fastforce simp add: winding_number_join path_image_join)  | 
|
131  | 
done  | 
|
132  | 
show ?thesis  | 
|
133  | 
by (rule winding_number_eq_1) (simp_all add: path_image_join gt0 lt2)  | 
|
134  | 
qed  | 
|
135  | 
||
136  | 
proposition winding_number_triangle:  | 
|
137  | 
  assumes z: "z \<in> interior(convex hull {a,b,c})"
 | 
|
138  | 
shows "winding_number(linepath a b +++ linepath b c +++ linepath c a) z =  | 
|
139  | 
(if 0 < Im((b - a) * cnj (b - z)) then 1 else -1)"  | 
|
140  | 
proof -  | 
|
141  | 
  have [simp]: "{a,c,b} = {a,b,c}"  by auto
 | 
|
142  | 
have znot[simp]: "z \<notin> closed_segment a b" "z \<notin> closed_segment b c" "z \<notin> closed_segment c a"  | 
|
143  | 
using z interior_of_triangle [of a b c]  | 
|
144  | 
by (auto simp: closed_segment_def)  | 
|
145  | 
then have [simp]: "z \<notin> closed_segment b a" "z \<notin> closed_segment c b" "z \<notin> closed_segment a c"  | 
|
146  | 
using closed_segment_commute by blast+  | 
|
147  | 
have *: "winding_number (linepath a b +++ linepath b c +++ linepath c a) z =  | 
|
148  | 
winding_number (reversepath (linepath a c +++ linepath c b +++ linepath b a)) z"  | 
|
149  | 
by (simp add: reversepath_joinpaths winding_number_join not_in_path_image_join)  | 
|
150  | 
show ?thesis  | 
|
151  | 
using wn_triangle2 [OF z] apply (rule disjE)  | 
|
152  | 
apply (simp add: wn_triangle3 z)  | 
|
153  | 
apply (simp add: path_image_join winding_number_reversepath * wn_triangle3 z)  | 
|
154  | 
done  | 
|
155  | 
qed  | 
|
156  | 
||
157  | 
subsection\<open>Winding numbers for simple closed paths\<close>  | 
|
158  | 
||
159  | 
lemma winding_number_from_innerpath:  | 
|
160  | 
assumes "simple_path c1" and c1: "pathstart c1 = a" "pathfinish c1 = b"  | 
|
161  | 
and "simple_path c2" and c2: "pathstart c2 = a" "pathfinish c2 = b"  | 
|
162  | 
and "simple_path c" and c: "pathstart c = a" "pathfinish c = b"  | 
|
163  | 
      and c1c2: "path_image c1 \<inter> path_image c2 = {a,b}"
 | 
|
164  | 
      and c1c:  "path_image c1 \<inter> path_image c = {a,b}"
 | 
|
165  | 
      and c2c:  "path_image c2 \<inter> path_image c = {a,b}"
 | 
|
166  | 
      and ne_12: "path_image c \<inter> inside(path_image c1 \<union> path_image c2) \<noteq> {}"
 | 
|
167  | 
and z: "z \<in> inside(path_image c1 \<union> path_image c)"  | 
|
168  | 
and wn_d: "winding_number (c1 +++ reversepath c) z = d"  | 
|
169  | 
and "a \<noteq> b" "d \<noteq> 0"  | 
|
170  | 
obtains "z \<in> inside(path_image c1 \<union> path_image c2)" "winding_number (c1 +++ reversepath c2) z = d"  | 
|
171  | 
proof -  | 
|
172  | 
  obtain 0: "inside(path_image c1 \<union> path_image c) \<inter> inside(path_image c2 \<union> path_image c) = {}"
 | 
|
173  | 
and 1: "inside(path_image c1 \<union> path_image c) \<union> inside(path_image c2 \<union> path_image c) \<union>  | 
|
174  | 
             (path_image c - {a,b}) = inside(path_image c1 \<union> path_image c2)"
 | 
|
175  | 
by (rule split_inside_simple_closed_curve  | 
|
176  | 
[OF \<open>simple_path c1\<close> c1 \<open>simple_path c2\<close> c2 \<open>simple_path c\<close> c \<open>a \<noteq> b\<close> c1c2 c1c c2c ne_12])  | 
|
177  | 
have znot: "z \<notin> path_image c" "z \<notin> path_image c1" "z \<notin> path_image c2"  | 
|
178  | 
using union_with_outside z 1 by auto  | 
|
179  | 
have wn_cc2: "winding_number (c +++ reversepath c2) z = 0"  | 
|
180  | 
apply (rule winding_number_zero_in_outside)  | 
|
181  | 
apply (simp_all add: \<open>simple_path c2\<close> c c2 \<open>simple_path c\<close> simple_path_imp_path path_image_join)  | 
|
182  | 
by (metis "0" ComplI UnE disjoint_iff_not_equal sup.commute union_with_inside z znot)  | 
|
183  | 
show ?thesis  | 
|
184  | 
proof  | 
|
185  | 
show "z \<in> inside (path_image c1 \<union> path_image c2)"  | 
|
186  | 
using "1" z by blast  | 
|
187  | 
have "winding_number c1 z - winding_number c z = d "  | 
|
188  | 
using assms znot  | 
|
189  | 
by (metis wn_d winding_number_join simple_path_imp_path winding_number_reversepath add.commute path_image_reversepath path_reversepath pathstart_reversepath uminus_add_conv_diff)  | 
|
190  | 
then show "winding_number (c1 +++ reversepath c2) z = d"  | 
|
191  | 
using wn_cc2 by (simp add: winding_number_join simple_path_imp_path assms znot winding_number_reversepath)  | 
|
192  | 
qed  | 
|
193  | 
qed  | 
|
194  | 
||
195  | 
||
196  | 
||
197  | 
lemma simple_closed_path_wn1:  | 
|
198  | 
fixes a::complex and e::real  | 
|
199  | 
assumes "0 < e"  | 
|
200  | 
and sp_pl: "simple_path(p +++ linepath (a - e) (a + e))"  | 
|
201  | 
and psp: "pathstart p = a + e"  | 
|
202  | 
and pfp: "pathfinish p = a - e"  | 
|
203  | 
    and disj:  "ball a e \<inter> path_image p = {}"
 | 
|
204  | 
obtains z where "z \<in> inside (path_image (p +++ linepath (a - e) (a + e)))"  | 
|
205  | 
"cmod (winding_number (p +++ linepath (a - e) (a + e)) z) = 1"  | 
|
206  | 
proof -  | 
|
207  | 
have "arc p" and arc_lp: "arc (linepath (a - e) (a + e))"  | 
|
208  | 
    and pap: "path_image p \<inter> path_image (linepath (a - e) (a + e)) \<subseteq> {pathstart p, a-e}"
 | 
|
209  | 
using simple_path_join_loop_eq [of "linepath (a - e) (a + e)" p] assms by auto  | 
|
210  | 
have mid_eq_a: "midpoint (a - e) (a + e) = a"  | 
|
211  | 
by (simp add: midpoint_def)  | 
|
212  | 
then have "a \<in> path_image(p +++ linepath (a - e) (a + e))"  | 
|
213  | 
apply (simp add: assms path_image_join)  | 
|
214  | 
by (metis midpoint_in_closed_segment)  | 
|
215  | 
have "a \<in> frontier(inside (path_image(p +++ linepath (a - e) (a + e))))"  | 
|
216  | 
apply (simp add: assms Jordan_inside_outside)  | 
|
217  | 
apply (simp_all add: assms path_image_join)  | 
|
218  | 
by (metis mid_eq_a midpoint_in_closed_segment)  | 
|
219  | 
with \<open>0 < e\<close> obtain c where c: "c \<in> inside (path_image(p +++ linepath (a - e) (a + e)))"  | 
|
220  | 
and dac: "dist a c < e"  | 
|
221  | 
by (auto simp: frontier_straddle)  | 
|
222  | 
then have "c \<notin> path_image(p +++ linepath (a - e) (a + e))"  | 
|
223  | 
using inside_no_overlap by blast  | 
|
224  | 
then have "c \<notin> path_image p"  | 
|
225  | 
"c \<notin> closed_segment (a - of_real e) (a + of_real e)"  | 
|
226  | 
by (simp_all add: assms path_image_join)  | 
|
227  | 
  with \<open>0 < e\<close> dac have "c \<notin> affine hull {a - of_real e, a + of_real e}"
 | 
|
228  | 
by (simp add: segment_as_ball not_le)  | 
|
229  | 
  with \<open>0 < e\<close> have *: "~collinear{a - e, c,a + e}"
 | 
|
230  | 
using collinear_3_affine_hull [of "a-e" "a+e"] by (auto simp: insert_commute)  | 
|
231  | 
have 13: "1/3 + 1/3 + 1/3 = (1::real)" by simp  | 
|
232  | 
  have "(1/3) *\<^sub>R (a - of_real e) + (1/3) *\<^sub>R c + (1/3) *\<^sub>R (a + of_real e) \<in> interior(convex hull {a - e, c, a + e})"
 | 
|
233  | 
using interior_convex_hull_3_minimal [OF * DIM_complex]  | 
|
234  | 
by clarsimp (metis 13 zero_less_divide_1_iff zero_less_numeral)  | 
|
235  | 
  then obtain z where z: "z \<in> interior(convex hull {a - e, c, a + e})" by force
 | 
|
236  | 
have [simp]: "z \<notin> closed_segment (a - e) c"  | 
|
237  | 
by (metis DIM_complex Diff_iff IntD2 inf_sup_absorb interior_of_triangle z)  | 
|
238  | 
have [simp]: "z \<notin> closed_segment (a + e) (a - e)"  | 
|
239  | 
by (metis DIM_complex DiffD2 Un_iff interior_of_triangle z)  | 
|
240  | 
have [simp]: "z \<notin> closed_segment c (a + e)"  | 
|
241  | 
by (metis (no_types, lifting) DIM_complex DiffD2 Un_insert_right inf_sup_aci(5) insertCI interior_of_triangle mk_disjoint_insert z)  | 
|
242  | 
show thesis  | 
|
243  | 
proof  | 
|
244  | 
have "norm (winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z) = 1"  | 
|
245  | 
using winding_number_triangle [OF z] by simp  | 
|
246  | 
have zin: "z \<in> inside (path_image (linepath (a + e) (a - e)) \<union> path_image p)"  | 
|
247  | 
and zeq: "winding_number (linepath (a + e) (a - e) +++ reversepath p) z =  | 
|
248  | 
winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z"  | 
|
249  | 
proof (rule winding_number_from_innerpath  | 
|
250  | 
[of "linepath (a + e) (a - e)" "a+e" "a-e" p  | 
|
251  | 
"linepath (a + e) c +++ linepath c (a - e)" z  | 
|
252  | 
"winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z"])  | 
|
253  | 
show sp_aec: "simple_path (linepath (a + e) c +++ linepath c (a - e))"  | 
|
254  | 
proof (rule arc_imp_simple_path [OF arc_join])  | 
|
255  | 
show "arc (linepath (a + e) c)"  | 
|
256  | 
by (metis \<open>c \<notin> path_image p\<close> arc_linepath pathstart_in_path_image psp)  | 
|
257  | 
show "arc (linepath c (a - e))"  | 
|
258  | 
by (metis \<open>c \<notin> path_image p\<close> arc_linepath pathfinish_in_path_image pfp)  | 
|
259  | 
        show "path_image (linepath (a + e) c) \<inter> path_image (linepath c (a - e)) \<subseteq> {pathstart (linepath c (a - e))}"
 | 
|
260  | 
by clarsimp (metis "*" IntI Int_closed_segment closed_segment_commute singleton_iff)  | 
|
261  | 
qed auto  | 
|
262  | 
show "simple_path p"  | 
|
263  | 
using \<open>arc p\<close> arc_simple_path by blast  | 
|
264  | 
show sp_ae2: "simple_path (linepath (a + e) (a - e))"  | 
|
265  | 
using \<open>arc p\<close> arc_distinct_ends pfp psp by fastforce  | 
|
266  | 
show pa: "pathfinish (linepath (a + e) (a - e)) = a - e"  | 
|
267  | 
"pathstart (linepath (a + e) c +++ linepath c (a - e)) = a + e"  | 
|
268  | 
"pathfinish (linepath (a + e) c +++ linepath c (a - e)) = a - e"  | 
|
269  | 
"pathstart p = a + e" "pathfinish p = a - e"  | 
|
270  | 
"pathstart (linepath (a + e) (a - e)) = a + e"  | 
|
271  | 
by (simp_all add: assms)  | 
|
272  | 
      show 1: "path_image (linepath (a + e) (a - e)) \<inter> path_image p = {a + e, a - e}"
 | 
|
273  | 
proof  | 
|
274  | 
        show "path_image (linepath (a + e) (a - e)) \<inter> path_image p \<subseteq> {a + e, a - e}"
 | 
|
275  | 
using pap closed_segment_commute psp segment_convex_hull by fastforce  | 
|
276  | 
        show "{a + e, a - e} \<subseteq> path_image (linepath (a + e) (a - e)) \<inter> path_image p"
 | 
|
277  | 
using pap pathfinish_in_path_image pathstart_in_path_image pfp psp by fastforce  | 
|
278  | 
qed  | 
|
279  | 
show 2: "path_image (linepath (a + e) (a - e)) \<inter> path_image (linepath (a + e) c +++ linepath c (a - e)) =  | 
|
280  | 
               {a + e, a - e}"  (is "?lhs = ?rhs")
 | 
|
281  | 
proof  | 
|
282  | 
        have "\<not> collinear {c, a + e, a - e}"
 | 
|
283  | 
using * by (simp add: insert_commute)  | 
|
284  | 
        then have "convex hull {a + e, a - e} \<inter> convex hull {a + e, c} = {a + e}"
 | 
|
285  | 
                  "convex hull {a + e, a - e} \<inter> convex hull {c, a - e} = {a - e}"
 | 
|
286  | 
by (metis (full_types) Int_closed_segment insert_commute segment_convex_hull)+  | 
|
287  | 
then show "?lhs \<subseteq> ?rhs"  | 
|
288  | 
by (metis Int_Un_distrib equalityD1 insert_is_Un path_image_join path_image_linepath path_join_eq path_linepath segment_convex_hull simple_path_def sp_aec)  | 
|
289  | 
show "?rhs \<subseteq> ?lhs"  | 
|
290  | 
using segment_convex_hull by (simp add: path_image_join)  | 
|
291  | 
qed  | 
|
292  | 
      have "path_image p \<inter> path_image (linepath (a + e) c) \<subseteq> {a + e}"
 | 
|
293  | 
proof (clarsimp simp: path_image_join)  | 
|
294  | 
fix x  | 
|
295  | 
assume "x \<in> path_image p" and x_ac: "x \<in> closed_segment (a + e) c"  | 
|
296  | 
then have "dist x a \<ge> e"  | 
|
297  | 
by (metis IntI all_not_in_conv disj dist_commute mem_ball not_less)  | 
|
298  | 
with x_ac dac \<open>e > 0\<close> show "x = a + e"  | 
|
299  | 
by (auto simp: norm_minus_commute dist_norm closed_segment_eq_open dest: open_segment_furthest_le [where y=a])  | 
|
300  | 
qed  | 
|
301  | 
moreover  | 
|
302  | 
      have "path_image p \<inter> path_image (linepath c (a - e)) \<subseteq> {a - e}"
 | 
|
303  | 
proof (clarsimp simp: path_image_join)  | 
|
304  | 
fix x  | 
|
305  | 
assume "x \<in> path_image p" and x_ac: "x \<in> closed_segment c (a - e)"  | 
|
306  | 
then have "dist x a \<ge> e"  | 
|
307  | 
by (metis IntI all_not_in_conv disj dist_commute mem_ball not_less)  | 
|
308  | 
with x_ac dac \<open>e > 0\<close> show "x = a - e"  | 
|
309  | 
by (auto simp: norm_minus_commute dist_norm closed_segment_eq_open dest: open_segment_furthest_le [where y=a])  | 
|
310  | 
qed  | 
|
311  | 
ultimately  | 
|
312  | 
      have "path_image p \<inter> path_image (linepath (a + e) c +++ linepath c (a - e)) \<subseteq> {a + e, a - e}"
 | 
|
313  | 
by (force simp: path_image_join)  | 
|
314  | 
      then show 3: "path_image p \<inter> path_image (linepath (a + e) c +++ linepath c (a - e)) = {a + e, a - e}"
 | 
|
315  | 
apply (rule equalityI)  | 
|
316  | 
apply (clarsimp simp: path_image_join)  | 
|
317  | 
apply (metis pathstart_in_path_image psp pathfinish_in_path_image pfp)  | 
|
318  | 
done  | 
|
319  | 
show 4: "path_image (linepath (a + e) c +++ linepath c (a - e)) \<inter>  | 
|
320  | 
               inside (path_image (linepath (a + e) (a - e)) \<union> path_image p) \<noteq> {}"
 | 
|
321  | 
apply (clarsimp simp: path_image_join segment_convex_hull disjoint_iff_not_equal)  | 
|
322  | 
by (metis (no_types, hide_lams) UnI1 Un_commute c closed_segment_commute ends_in_segment(1) path_image_join  | 
|
323  | 
path_image_linepath pathstart_linepath pfp segment_convex_hull)  | 
|
324  | 
show zin_inside: "z \<in> inside (path_image (linepath (a + e) (a - e)) \<union>  | 
|
325  | 
path_image (linepath (a + e) c +++ linepath c (a - e)))"  | 
|
326  | 
apply (simp add: path_image_join)  | 
|
327  | 
by (metis z inside_of_triangle DIM_complex Un_commute closed_segment_commute)  | 
|
328  | 
show 5: "winding_number  | 
|
329  | 
(linepath (a + e) (a - e) +++ reversepath (linepath (a + e) c +++ linepath c (a - e))) z =  | 
|
330  | 
winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z"  | 
|
331  | 
by (simp add: reversepath_joinpaths path_image_join winding_number_join)  | 
|
332  | 
show 6: "winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z \<noteq> 0"  | 
|
333  | 
by (simp add: winding_number_triangle z)  | 
|
334  | 
show "winding_number (linepath (a + e) (a - e) +++ reversepath p) z =  | 
|
335  | 
winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z"  | 
|
336  | 
by (metis 1 2 3 4 5 6 pa sp_aec sp_ae2 \<open>arc p\<close> \<open>simple_path p\<close> arc_distinct_ends winding_number_from_innerpath zin_inside)  | 
|
337  | 
qed (use assms \<open>e > 0\<close> in auto)  | 
|
338  | 
show "z \<in> inside (path_image (p +++ linepath (a - e) (a + e)))"  | 
|
339  | 
using zin by (simp add: assms path_image_join Un_commute closed_segment_commute)  | 
|
340  | 
then have "cmod (winding_number (p +++ linepath (a - e) (a + e)) z) =  | 
|
341  | 
cmod ((winding_number(reversepath (p +++ linepath (a - e) (a + e))) z))"  | 
|
342  | 
apply (subst winding_number_reversepath)  | 
|
343  | 
using simple_path_imp_path sp_pl apply blast  | 
|
344  | 
apply (metis IntI emptyE inside_no_overlap)  | 
|
345  | 
by (simp add: inside_def)  | 
|
346  | 
also have "... = cmod (winding_number(linepath (a + e) (a - e) +++ reversepath p) z)"  | 
|
347  | 
by (simp add: pfp reversepath_joinpaths)  | 
|
348  | 
also have "... = cmod (winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z)"  | 
|
349  | 
by (simp add: zeq)  | 
|
350  | 
also have "... = 1"  | 
|
351  | 
using z by (simp add: interior_of_triangle winding_number_triangle)  | 
|
352  | 
finally show "cmod (winding_number (p +++ linepath (a - e) (a + e)) z) = 1" .  | 
|
353  | 
qed  | 
|
354  | 
qed  | 
|
355  | 
||
356  | 
||
357  | 
||
358  | 
lemma simple_closed_path_wn2:  | 
|
359  | 
fixes a::complex and d e::real  | 
|
360  | 
assumes "0 < d" "0 < e"  | 
|
361  | 
and sp_pl: "simple_path(p +++ linepath (a - d) (a + e))"  | 
|
362  | 
and psp: "pathstart p = a + e"  | 
|
363  | 
and pfp: "pathfinish p = a - d"  | 
|
364  | 
obtains z where "z \<in> inside (path_image (p +++ linepath (a - d) (a + e)))"  | 
|
365  | 
"cmod (winding_number (p +++ linepath (a - d) (a + e)) z) = 1"  | 
|
366  | 
proof -  | 
|
367  | 
have [simp]: "a + of_real x \<in> closed_segment (a - \<alpha>) (a - \<beta>) \<longleftrightarrow> x \<in> closed_segment (-\<alpha>) (-\<beta>)" for x \<alpha> \<beta>::real  | 
|
368  | 
using closed_segment_translation_eq [of a]  | 
|
369  | 
by (metis (no_types, hide_lams) add_uminus_conv_diff of_real_minus of_real_closed_segment)  | 
|
370  | 
have [simp]: "a - of_real x \<in> closed_segment (a + \<alpha>) (a + \<beta>) \<longleftrightarrow> -x \<in> closed_segment \<alpha> \<beta>" for x \<alpha> \<beta>::real  | 
|
371  | 
by (metis closed_segment_translation_eq diff_conv_add_uminus of_real_closed_segment of_real_minus)  | 
|
372  | 
have "arc p" and arc_lp: "arc (linepath (a - d) (a + e))" and "path p"  | 
|
373  | 
    and pap: "path_image p \<inter> closed_segment (a - d) (a + e) \<subseteq> {a+e, a-d}"
 | 
|
374  | 
using simple_path_join_loop_eq [of "linepath (a - d) (a + e)" p] assms arc_imp_path by auto  | 
|
375  | 
have "0 \<in> closed_segment (-d) e"  | 
|
376  | 
using \<open>0 < d\<close> \<open>0 < e\<close> closed_segment_eq_real_ivl by auto  | 
|
377  | 
then have "a \<in> path_image (linepath (a - d) (a + e))"  | 
|
378  | 
using of_real_closed_segment [THEN iffD2]  | 
|
379  | 
by (force dest: closed_segment_translation_eq [of a, THEN iffD2] simp del: of_real_closed_segment)  | 
|
380  | 
then have "a \<notin> path_image p"  | 
|
381  | 
using \<open>0 < d\<close> \<open>0 < e\<close> pap by auto  | 
|
382  | 
  then obtain k where "0 < k" and k: "ball a k \<inter> (path_image p) = {}"
 | 
|
383  | 
using \<open>0 < e\<close> \<open>path p\<close> not_on_path_ball by blast  | 
|
384  | 
define kde where "kde \<equiv> (min k (min d e)) / 2"  | 
|
385  | 
have "0 < kde" "kde < k" "kde < d" "kde < e"  | 
|
386  | 
using \<open>0 < k\<close> \<open>0 < d\<close> \<open>0 < e\<close> by (auto simp: kde_def)  | 
|
387  | 
let ?q = "linepath (a + kde) (a + e) +++ p +++ linepath (a - d) (a - kde)"  | 
|
388  | 
have "- kde \<in> closed_segment (-d) e"  | 
|
389  | 
using \<open>0 < kde\<close> \<open>kde < d\<close> \<open>kde < e\<close> closed_segment_eq_real_ivl by auto  | 
|
390  | 
then have a_diff_kde: "a - kde \<in> closed_segment (a - d) (a + e)"  | 
|
391  | 
using of_real_closed_segment [THEN iffD2]  | 
|
392  | 
by (force dest: closed_segment_translation_eq [of a, THEN iffD2] simp del: of_real_closed_segment)  | 
|
393  | 
then have clsub2: "closed_segment (a - d) (a - kde) \<subseteq> closed_segment (a - d) (a + e)"  | 
|
394  | 
by (simp add: subset_closed_segment)  | 
|
395  | 
  then have "path_image p \<inter> closed_segment (a - d) (a - kde) \<subseteq> {a + e, a - d}"
 | 
|
396  | 
using pap by force  | 
|
397  | 
moreover  | 
|
398  | 
have "a + e \<notin> path_image p \<inter> closed_segment (a - d) (a - kde)"  | 
|
399  | 
using \<open>0 < kde\<close> \<open>kde < d\<close> \<open>0 < e\<close> by (auto simp: closed_segment_eq_real_ivl)  | 
|
400  | 
  ultimately have sub_a_diff_d: "path_image p \<inter> closed_segment (a - d) (a - kde) \<subseteq> {a - d}"
 | 
|
401  | 
by blast  | 
|
402  | 
have "kde \<in> closed_segment (-d) e"  | 
|
403  | 
using \<open>0 < kde\<close> \<open>kde < d\<close> \<open>kde < e\<close> closed_segment_eq_real_ivl by auto  | 
|
404  | 
then have a_diff_kde: "a + kde \<in> closed_segment (a - d) (a + e)"  | 
|
405  | 
using of_real_closed_segment [THEN iffD2]  | 
|
406  | 
by (force dest: closed_segment_translation_eq [of "a", THEN iffD2] simp del: of_real_closed_segment)  | 
|
407  | 
then have clsub1: "closed_segment (a + kde) (a + e) \<subseteq> closed_segment (a - d) (a + e)"  | 
|
408  | 
by (simp add: subset_closed_segment)  | 
|
409  | 
  then have "closed_segment (a + kde) (a + e) \<inter> path_image p \<subseteq> {a + e, a - d}"
 | 
|
410  | 
using pap by force  | 
|
411  | 
moreover  | 
|
412  | 
  have "closed_segment (a + kde) (a + e) \<inter> closed_segment (a - d) (a - kde) = {}"
 | 
|
413  | 
proof (clarsimp intro!: equals0I)  | 
|
414  | 
fix y  | 
|
415  | 
assume y1: "y \<in> closed_segment (a + kde) (a + e)"  | 
|
416  | 
and y2: "y \<in> closed_segment (a - d) (a - kde)"  | 
|
417  | 
obtain u where u: "y = a + of_real u" and "0 < u"  | 
|
418  | 
using y1 \<open>0 < kde\<close> \<open>kde < e\<close> \<open>0 < e\<close> apply (clarsimp simp: in_segment)  | 
|
419  | 
apply (rule_tac u = "(1 - u)*kde + u*e" in that)  | 
|
420  | 
apply (auto simp: scaleR_conv_of_real algebra_simps)  | 
|
421  | 
by (meson le_less_trans less_add_same_cancel2 less_eq_real_def mult_left_mono)  | 
|
422  | 
moreover  | 
|
423  | 
obtain v where v: "y = a + of_real v" and "v \<le> 0"  | 
|
424  | 
using y2 \<open>0 < kde\<close> \<open>0 < d\<close> \<open>0 < e\<close> apply (clarsimp simp: in_segment)  | 
|
425  | 
apply (rule_tac v = "- ((1 - u)*d + u*kde)" in that)  | 
|
426  | 
apply (force simp: scaleR_conv_of_real algebra_simps)  | 
|
427  | 
by (meson less_eq_real_def neg_le_0_iff_le segment_bound_lemma)  | 
|
428  | 
ultimately show False  | 
|
429  | 
by auto  | 
|
430  | 
qed  | 
|
431  | 
moreover have "a - d \<notin> closed_segment (a + kde) (a + e)"  | 
|
432  | 
using \<open>0 < kde\<close> \<open>kde < d\<close> \<open>0 < e\<close> by (auto simp: closed_segment_eq_real_ivl)  | 
|
433  | 
ultimately have sub_a_plus_e:  | 
|
434  | 
"closed_segment (a + kde) (a + e) \<inter> (path_image p \<union> closed_segment (a - d) (a - kde))  | 
|
435  | 
       \<subseteq> {a + e}"
 | 
|
436  | 
by auto  | 
|
437  | 
have "kde \<in> closed_segment (-kde) e"  | 
|
438  | 
using \<open>0 < kde\<close> \<open>kde < d\<close> \<open>kde < e\<close> closed_segment_eq_real_ivl by auto  | 
|
439  | 
then have a_add_kde: "a + kde \<in> closed_segment (a - kde) (a + e)"  | 
|
440  | 
using of_real_closed_segment [THEN iffD2]  | 
|
441  | 
by (force dest: closed_segment_translation_eq [of "a", THEN iffD2] simp del: of_real_closed_segment)  | 
|
442  | 
  have "closed_segment (a - kde) (a + kde) \<inter> closed_segment (a + kde) (a + e) = {a + kde}"
 | 
|
443  | 
by (metis a_add_kde Int_closed_segment)  | 
|
444  | 
moreover  | 
|
445  | 
  have "path_image p \<inter> closed_segment (a - kde) (a + kde) = {}"
 | 
|
446  | 
proof (rule equals0I, clarify)  | 
|
447  | 
fix y assume "y \<in> path_image p" "y \<in> closed_segment (a - kde) (a + kde)"  | 
|
448  | 
with equals0D [OF k, of y] \<open>0 < kde\<close> \<open>kde < k\<close> show False  | 
|
449  | 
by (auto simp: dist_norm dest: dist_decreases_closed_segment [where c=a])  | 
|
450  | 
qed  | 
|
451  | 
moreover  | 
|
452  | 
have "- kde \<in> closed_segment (-d) kde"  | 
|
453  | 
using \<open>0 < kde\<close> \<open>kde < d\<close> \<open>kde < e\<close> closed_segment_eq_real_ivl by auto  | 
|
454  | 
then have a_diff_kde': "a - kde \<in> closed_segment (a - d) (a + kde)"  | 
|
455  | 
using of_real_closed_segment [THEN iffD2]  | 
|
456  | 
by (force dest: closed_segment_translation_eq [of a, THEN iffD2] simp del: of_real_closed_segment)  | 
|
457  | 
  then have "closed_segment (a - d) (a - kde) \<inter> closed_segment (a - kde) (a + kde) = {a - kde}"
 | 
|
458  | 
by (metis Int_closed_segment)  | 
|
459  | 
ultimately  | 
|
460  | 
  have pa_subset_pm_kde: "path_image ?q \<inter> closed_segment (a - kde) (a + kde) \<subseteq> {a - kde, a + kde}"
 | 
|
461  | 
by (auto simp: path_image_join assms)  | 
|
462  | 
have ge_kde1: "\<exists>y. x = a + y \<and> y \<ge> kde" if "x \<in> closed_segment (a + kde) (a + e)" for x  | 
|
463  | 
using that \<open>kde < e\<close> mult_le_cancel_left  | 
|
464  | 
apply (auto simp: in_segment)  | 
|
465  | 
apply (rule_tac x="(1-u)*kde + u*e" in exI)  | 
|
466  | 
apply (fastforce simp: algebra_simps scaleR_conv_of_real)  | 
|
467  | 
done  | 
|
468  | 
have ge_kde2: "\<exists>y. x = a + y \<and> y \<le> -kde" if "x \<in> closed_segment (a - d) (a - kde)" for x  | 
|
469  | 
using that \<open>kde < d\<close> affine_ineq  | 
|
470  | 
apply (auto simp: in_segment)  | 
|
471  | 
apply (rule_tac x="- ((1-u)*d + u*kde)" in exI)  | 
|
472  | 
apply (fastforce simp: algebra_simps scaleR_conv_of_real)  | 
|
473  | 
done  | 
|
474  | 
have notin_paq: "x \<notin> path_image ?q" if "dist a x < kde" for x  | 
|
475  | 
using that using \<open>0 < kde\<close> \<open>kde < d\<close> \<open>kde < k\<close>  | 
|
476  | 
apply (auto simp: path_image_join assms dist_norm dest!: ge_kde1 ge_kde2)  | 
|
477  | 
by (meson k disjoint_iff_not_equal le_less_trans less_eq_real_def mem_ball that)  | 
|
478  | 
obtain z where zin: "z \<in> inside (path_image (?q +++ linepath (a - kde) (a + kde)))"  | 
|
479  | 
and z1: "cmod (winding_number (?q +++ linepath (a - kde) (a + kde)) z) = 1"  | 
|
480  | 
proof (rule simple_closed_path_wn1 [of kde ?q a])  | 
|
481  | 
show "simple_path (?q +++ linepath (a - kde) (a + kde))"  | 
|
482  | 
proof (intro simple_path_join_loop conjI)  | 
|
483  | 
show "arc ?q"  | 
|
484  | 
proof (rule arc_join)  | 
|
485  | 
show "arc (linepath (a + kde) (a + e))"  | 
|
486  | 
using \<open>kde < e\<close> \<open>arc p\<close> by (force simp: pfp)  | 
|
487  | 
show "arc (p +++ linepath (a - d) (a - kde))"  | 
|
488  | 
using \<open>kde < d\<close> \<open>kde < e\<close> \<open>arc p\<close> sub_a_diff_d by (force simp: pfp intro: arc_join)  | 
|
489  | 
qed (auto simp: psp pfp path_image_join sub_a_plus_e)  | 
|
490  | 
show "arc (linepath (a - kde) (a + kde))"  | 
|
491  | 
using \<open>0 < kde\<close> by auto  | 
|
492  | 
qed (use pa_subset_pm_kde in auto)  | 
|
493  | 
qed (use \<open>0 < kde\<close> notin_paq in auto)  | 
|
494  | 
have eq: "path_image (?q +++ linepath (a - kde) (a + kde)) = path_image (p +++ linepath (a - d) (a + e))"  | 
|
495  | 
(is "?lhs = ?rhs")  | 
|
496  | 
proof  | 
|
497  | 
show "?lhs \<subseteq> ?rhs"  | 
|
498  | 
using clsub1 clsub2 apply (auto simp: path_image_join assms)  | 
|
499  | 
by (meson subsetCE subset_closed_segment)  | 
|
500  | 
show "?rhs \<subseteq> ?lhs"  | 
|
501  | 
apply (simp add: path_image_join assms Un_ac)  | 
|
502  | 
by (metis Un_closed_segment Un_assoc a_diff_kde a_diff_kde' le_supI2 subset_refl)  | 
|
503  | 
qed  | 
|
504  | 
show thesis  | 
|
505  | 
proof  | 
|
506  | 
show zzin: "z \<in> inside (path_image (p +++ linepath (a - d) (a + e)))"  | 
|
507  | 
by (metis eq zin)  | 
|
508  | 
then have znotin: "z \<notin> path_image p"  | 
|
509  | 
by (metis ComplD Un_iff inside_Un_outside path_image_join pathfinish_linepath pathstart_reversepath pfp reversepath_linepath)  | 
|
510  | 
have znotin_de: "z \<notin> closed_segment (a - d) (a + kde)"  | 
|
511  | 
by (metis ComplD Un_iff Un_closed_segment a_diff_kde inside_Un_outside path_image_join path_image_linepath pathstart_linepath pfp zzin)  | 
|
512  | 
have "winding_number (linepath (a - d) (a + e)) z =  | 
|
513  | 
winding_number (linepath (a - d) (a + kde)) z + winding_number (linepath (a + kde) (a + e)) z"  | 
|
514  | 
apply (rule winding_number_split_linepath)  | 
|
515  | 
apply (simp add: a_diff_kde)  | 
|
516  | 
by (metis ComplD Un_iff inside_Un_outside path_image_join path_image_linepath pathstart_linepath pfp zzin)  | 
|
517  | 
also have "... = winding_number (linepath (a + kde) (a + e)) z +  | 
|
518  | 
(winding_number (linepath (a - d) (a - kde)) z +  | 
|
519  | 
winding_number (linepath (a - kde) (a + kde)) z)"  | 
|
520  | 
by (simp add: winding_number_split_linepath [of "a-kde", symmetric] znotin_de a_diff_kde')  | 
|
521  | 
finally have "winding_number (p +++ linepath (a - d) (a + e)) z =  | 
|
522  | 
winding_number p z + winding_number (linepath (a + kde) (a + e)) z +  | 
|
523  | 
(winding_number (linepath (a - d) (a - kde)) z +  | 
|
524  | 
winding_number (linepath (a - kde) (a + kde)) z)"  | 
|
525  | 
by (metis (no_types, lifting) ComplD Un_iff \<open>arc p\<close> add.assoc arc_imp_path eq path_image_join path_join_path_ends path_linepath simple_path_imp_path sp_pl union_with_outside winding_number_join zin)  | 
|
526  | 
also have "... = winding_number ?q z + winding_number (linepath (a - kde) (a + kde)) z"  | 
|
527  | 
using \<open>path p\<close> znotin assms zzin clsub1  | 
|
528  | 
apply (subst winding_number_join, auto)  | 
|
529  | 
apply (metis (no_types, hide_lams) ComplD Un_iff contra_subsetD inside_Un_outside path_image_join path_image_linepath pathstart_linepath)  | 
|
530  | 
apply (metis Un_iff Un_closed_segment a_diff_kde' not_in_path_image_join path_image_linepath znotin_de)  | 
|
531  | 
by (metis Un_iff Un_closed_segment a_diff_kde' path_image_linepath path_linepath pathstart_linepath winding_number_join znotin_de)  | 
|
532  | 
also have "... = winding_number (?q +++ linepath (a - kde) (a + kde)) z"  | 
|
533  | 
using \<open>path p\<close> assms zin  | 
|
534  | 
apply (subst winding_number_join [symmetric], auto)  | 
|
535  | 
apply (metis ComplD Un_iff path_image_join pathfinish_join pathfinish_linepath pathstart_linepath union_with_outside)  | 
|
536  | 
by (metis Un_iff Un_closed_segment a_diff_kde' znotin_de)  | 
|
537  | 
finally have "winding_number (p +++ linepath (a - d) (a + e)) z =  | 
|
538  | 
winding_number (?q +++ linepath (a - kde) (a + kde)) z" .  | 
|
539  | 
then show "cmod (winding_number (p +++ linepath (a - d) (a + e)) z) = 1"  | 
|
540  | 
by (simp add: z1)  | 
|
541  | 
qed  | 
|
542  | 
qed  | 
|
543  | 
||
544  | 
||
545  | 
proposition simple_closed_path_wn3:  | 
|
546  | 
fixes p :: "real \<Rightarrow> complex"  | 
|
547  | 
assumes "simple_path p" and loop: "pathfinish p = pathstart p"  | 
|
548  | 
obtains z where "z \<in> inside (path_image p)" "cmod (winding_number p z) = 1"  | 
|
549  | 
proof -  | 
|
550  | 
  have ins: "inside(path_image p) \<noteq> {}" "open(inside(path_image p))"
 | 
|
551  | 
"connected(inside(path_image p))"  | 
|
552  | 
   and out: "outside(path_image p) \<noteq> {}" "open(outside(path_image p))"
 | 
|
553  | 
"connected(outside(path_image p))"  | 
|
554  | 
and bo: "bounded(inside(path_image p))" "\<not> bounded(outside(path_image p))"  | 
|
555  | 
   and ins_out: "inside(path_image p) \<inter> outside(path_image p) = {}"
 | 
|
556  | 
"inside(path_image p) \<union> outside(path_image p) = - path_image p"  | 
|
557  | 
and fro: "frontier(inside(path_image p)) = path_image p"  | 
|
558  | 
"frontier(outside(path_image p)) = path_image p"  | 
|
559  | 
using Jordan_inside_outside [OF assms] by auto  | 
|
560  | 
obtain a where a: "a \<in> inside(path_image p)"  | 
|
561  | 
    using \<open>inside (path_image p) \<noteq> {}\<close> by blast
 | 
|
562  | 
obtain d::real where "0 < d" and d_fro: "a - d \<in> frontier (inside (path_image p))"  | 
|
563  | 
and d_int: "\<And>\<epsilon>. \<lbrakk>0 \<le> \<epsilon>; \<epsilon> < d\<rbrakk> \<Longrightarrow> (a - \<epsilon>) \<in> inside (path_image p)"  | 
|
564  | 
apply (rule ray_to_frontier [of "inside (path_image p)" a "-1"])  | 
|
565  | 
using \<open>bounded (inside (path_image p))\<close> \<open>open (inside (path_image p))\<close> a interior_eq  | 
|
566  | 
apply (auto simp: of_real_def)  | 
|
567  | 
done  | 
|
568  | 
obtain e::real where "0 < e" and e_fro: "a + e \<in> frontier (inside (path_image p))"  | 
|
569  | 
and e_int: "\<And>\<epsilon>. \<lbrakk>0 \<le> \<epsilon>; \<epsilon> < e\<rbrakk> \<Longrightarrow> (a + \<epsilon>) \<in> inside (path_image p)"  | 
|
570  | 
apply (rule ray_to_frontier [of "inside (path_image p)" a 1])  | 
|
571  | 
using \<open>bounded (inside (path_image p))\<close> \<open>open (inside (path_image p))\<close> a interior_eq  | 
|
572  | 
apply (auto simp: of_real_def)  | 
|
573  | 
done  | 
|
574  | 
obtain t0 where "0 \<le> t0" "t0 \<le> 1" and pt: "p t0 = a - d"  | 
|
575  | 
using a d_fro fro by (auto simp: path_image_def)  | 
|
576  | 
obtain q where "simple_path q" and q_ends: "pathstart q = a - d" "pathfinish q = a - d"  | 
|
577  | 
and q_eq_p: "path_image q = path_image p"  | 
|
578  | 
and wn_q_eq_wn_p: "\<And>z. z \<in> inside(path_image p) \<Longrightarrow> winding_number q z = winding_number p z"  | 
|
579  | 
proof  | 
|
580  | 
show "simple_path (shiftpath t0 p)"  | 
|
581  | 
by (simp add: pathstart_shiftpath pathfinish_shiftpath  | 
|
582  | 
simple_path_shiftpath path_image_shiftpath \<open>0 \<le> t0\<close> \<open>t0 \<le> 1\<close> assms)  | 
|
583  | 
show "pathstart (shiftpath t0 p) = a - d"  | 
|
584  | 
using pt by (simp add: \<open>t0 \<le> 1\<close> pathstart_shiftpath)  | 
|
585  | 
show "pathfinish (shiftpath t0 p) = a - d"  | 
|
586  | 
by (simp add: \<open>0 \<le> t0\<close> loop pathfinish_shiftpath pt)  | 
|
587  | 
show "path_image (shiftpath t0 p) = path_image p"  | 
|
588  | 
by (simp add: \<open>0 \<le> t0\<close> \<open>t0 \<le> 1\<close> loop path_image_shiftpath)  | 
|
589  | 
show "winding_number (shiftpath t0 p) z = winding_number p z"  | 
|
590  | 
if "z \<in> inside (path_image p)" for z  | 
|
591  | 
by (metis ComplD Un_iff \<open>0 \<le> t0\<close> \<open>t0 \<le> 1\<close> \<open>simple_path p\<close> atLeastAtMost_iff inside_Un_outside  | 
|
592  | 
loop simple_path_imp_path that winding_number_shiftpath)  | 
|
593  | 
qed  | 
|
594  | 
have ad_not_ae: "a - d \<noteq> a + e"  | 
|
595  | 
by (metis \<open>0 < d\<close> \<open>0 < e\<close> add.left_inverse add_left_cancel add_uminus_conv_diff  | 
|
596  | 
le_add_same_cancel2 less_eq_real_def not_less of_real_add of_real_def of_real_eq_0_iff pt)  | 
|
597  | 
  have ad_ae_q: "{a - d, a + e} \<subseteq> path_image q"
 | 
|
598  | 
using \<open>path_image q = path_image p\<close> d_fro e_fro fro(1) by auto  | 
|
599  | 
have ada: "open_segment (a - d) a \<subseteq> inside (path_image p)"  | 
|
600  | 
proof (clarsimp simp: in_segment)  | 
|
601  | 
fix u::real assume "0 < u" "u < 1"  | 
|
602  | 
with d_int have "a - (1 - u) * d \<in> inside (path_image p)"  | 
|
603  | 
by (metis \<open>0 < d\<close> add.commute diff_add_cancel left_diff_distrib' less_add_same_cancel2 less_eq_real_def mult.left_neutral zero_less_mult_iff)  | 
|
604  | 
then show "(1 - u) *\<^sub>R (a - d) + u *\<^sub>R a \<in> inside (path_image p)"  | 
|
605  | 
by (simp add: diff_add_eq of_real_def real_vector.scale_right_diff_distrib)  | 
|
606  | 
qed  | 
|
607  | 
have aae: "open_segment a (a + e) \<subseteq> inside (path_image p)"  | 
|
608  | 
proof (clarsimp simp: in_segment)  | 
|
609  | 
fix u::real assume "0 < u" "u < 1"  | 
|
610  | 
with e_int have "a + u * e \<in> inside (path_image p)"  | 
|
611  | 
by (meson \<open>0 < e\<close> less_eq_real_def mult_less_cancel_right2 not_less zero_less_mult_iff)  | 
|
612  | 
then show "(1 - u) *\<^sub>R a + u *\<^sub>R (a + e) \<in> inside (path_image p)"  | 
|
613  | 
apply (simp add: algebra_simps)  | 
|
614  | 
by (simp add: diff_add_eq of_real_def real_vector.scale_right_diff_distrib)  | 
|
615  | 
qed  | 
|
616  | 
have "complex_of_real (d * d + (e * e + d * (e + e))) \<noteq> 0"  | 
|
617  | 
using ad_not_ae  | 
|
618  | 
by (metis \<open>0 < d\<close> \<open>0 < e\<close> add_strict_left_mono less_add_same_cancel1 not_sum_squares_lt_zero  | 
|
619  | 
of_real_eq_0_iff zero_less_double_add_iff_zero_less_single_add zero_less_mult_iff)  | 
|
620  | 
then have a_in_de: "a \<in> open_segment (a - d) (a + e)"  | 
|
621  | 
using ad_not_ae \<open>0 < d\<close> \<open>0 < e\<close>  | 
|
622  | 
apply (auto simp: in_segment algebra_simps scaleR_conv_of_real)  | 
|
623  | 
apply (rule_tac x="d / (d+e)" in exI)  | 
|
624  | 
apply (auto simp: field_simps)  | 
|
625  | 
done  | 
|
626  | 
then have "open_segment (a - d) (a + e) \<subseteq> inside (path_image p)"  | 
|
627  | 
using ada a aae Un_open_segment [of a "a-d" "a+e"] by blast  | 
|
628  | 
  then have "path_image q \<inter> open_segment (a - d) (a + e) = {}"
 | 
|
629  | 
using inside_no_overlap by (fastforce simp: q_eq_p)  | 
|
630  | 
  with ad_ae_q have paq_Int_cs: "path_image q \<inter> closed_segment (a - d) (a + e) = {a - d, a + e}"
 | 
|
631  | 
by (simp add: closed_segment_eq_open)  | 
|
632  | 
obtain t where "0 \<le> t" "t \<le> 1" and qt: "q t = a + e"  | 
|
633  | 
using a e_fro fro ad_ae_q by (auto simp: path_defs)  | 
|
634  | 
then have "t \<noteq> 0"  | 
|
635  | 
by (metis ad_not_ae pathstart_def q_ends(1))  | 
|
636  | 
then have "t \<noteq> 1"  | 
|
637  | 
by (metis ad_not_ae pathfinish_def q_ends(2) qt)  | 
|
638  | 
have q01: "q 0 = a - d" "q 1 = a - d"  | 
|
639  | 
using q_ends by (auto simp: pathstart_def pathfinish_def)  | 
|
640  | 
obtain z where zin: "z \<in> inside (path_image (subpath t 0 q +++ linepath (a - d) (a + e)))"  | 
|
641  | 
and z1: "cmod (winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z) = 1"  | 
|
642  | 
proof (rule simple_closed_path_wn2 [of d e "subpath t 0 q" a], simp_all add: q01)  | 
|
643  | 
show "simple_path (subpath t 0 q +++ linepath (a - d) (a + e))"  | 
|
644  | 
proof (rule simple_path_join_loop, simp_all add: qt q01)  | 
|
645  | 
have "inj_on q (closed_segment t 0)"  | 
|
646  | 
using \<open>0 \<le> t\<close> \<open>simple_path q\<close> \<open>t \<le> 1\<close> \<open>t \<noteq> 0\<close> \<open>t \<noteq> 1\<close>  | 
|
647  | 
by (fastforce simp: simple_path_def inj_on_def closed_segment_eq_real_ivl)  | 
|
648  | 
then show "arc (subpath t 0 q)"  | 
|
649  | 
using \<open>0 \<le> t\<close> \<open>simple_path q\<close> \<open>t \<le> 1\<close> \<open>t \<noteq> 0\<close>  | 
|
650  | 
by (simp add: arc_subpath_eq simple_path_imp_path)  | 
|
651  | 
show "arc (linepath (a - d) (a + e))"  | 
|
652  | 
by (simp add: ad_not_ae)  | 
|
653  | 
      show "path_image (subpath t 0 q) \<inter> closed_segment (a - d) (a + e) \<subseteq> {a + e, a - d}"
 | 
|
654  | 
using qt paq_Int_cs \<open>simple_path q\<close> \<open>0 \<le> t\<close> \<open>t \<le> 1\<close>  | 
|
655  | 
by (force simp: dest: rev_subsetD [OF _ path_image_subpath_subset] intro: simple_path_imp_path)  | 
|
656  | 
qed  | 
|
657  | 
qed (auto simp: \<open>0 < d\<close> \<open>0 < e\<close> qt)  | 
|
658  | 
have pa01_Un: "path_image (subpath 0 t q) \<union> path_image (subpath 1 t q) = path_image q"  | 
|
659  | 
unfolding path_image_subpath  | 
|
660  | 
using \<open>0 \<le> t\<close> \<open>t \<le> 1\<close> by (force simp: path_image_def image_iff)  | 
|
661  | 
with paq_Int_cs have pa_01q:  | 
|
662  | 
        "(path_image (subpath 0 t q) \<union> path_image (subpath 1 t q)) \<inter> closed_segment (a - d) (a + e) = {a - d, a + e}"
 | 
|
663  | 
by metis  | 
|
664  | 
have z_notin_ed: "z \<notin> closed_segment (a + e) (a - d)"  | 
|
665  | 
using zin q01 by (simp add: path_image_join closed_segment_commute inside_def)  | 
|
666  | 
have z_notin_0t: "z \<notin> path_image (subpath 0 t q)"  | 
|
667  | 
by (metis (no_types, hide_lams) IntI Un_upper1 subsetD empty_iff inside_no_overlap path_image_join  | 
|
668  | 
path_image_subpath_commute pathfinish_subpath pathstart_def pathstart_linepath q_ends(1) qt subpath_trivial zin)  | 
|
669  | 
have [simp]: "- winding_number (subpath t 0 q) z = winding_number (subpath 0 t q) z"  | 
|
670  | 
by (metis \<open>0 \<le> t\<close> \<open>simple_path q\<close> \<open>t \<le> 1\<close> atLeastAtMost_iff zero_le_one  | 
|
671  | 
path_image_subpath_commute path_subpath real_eq_0_iff_le_ge_0  | 
|
672  | 
reversepath_subpath simple_path_imp_path winding_number_reversepath z_notin_0t)  | 
|
673  | 
obtain z_in_q: "z \<in> inside(path_image q)"  | 
|
674  | 
and wn_q: "winding_number (subpath 0 t q +++ subpath t 1 q) z = - winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z"  | 
|
675  | 
proof (rule winding_number_from_innerpath  | 
|
676  | 
[of "subpath 0 t q" "a-d" "a+e" "subpath 1 t q" "linepath (a - d) (a + e)"  | 
|
677  | 
z "- winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z"],  | 
|
678  | 
simp_all add: q01 qt pa01_Un reversepath_subpath)  | 
|
679  | 
show "simple_path (subpath 0 t q)" "simple_path (subpath 1 t q)"  | 
|
680  | 
by (simp_all add: \<open>0 \<le> t\<close> \<open>simple_path q\<close> \<open>t \<le> 1\<close> \<open>t \<noteq> 0\<close> \<open>t \<noteq> 1\<close> simple_path_subpath)  | 
|
681  | 
show "simple_path (linepath (a - d) (a + e))"  | 
|
682  | 
using ad_not_ae by blast  | 
|
683  | 
    show "path_image (subpath 0 t q) \<inter> path_image (subpath 1 t q) = {a - d, a + e}"  (is "?lhs = ?rhs")
 | 
|
684  | 
proof  | 
|
685  | 
show "?lhs \<subseteq> ?rhs"  | 
|
686  | 
using \<open>0 \<le> t\<close> \<open>simple_path q\<close> \<open>t \<le> 1\<close> \<open>t \<noteq> 1\<close> q_ends qt q01  | 
|
687  | 
by (force simp: pathfinish_def qt simple_path_def path_image_subpath)  | 
|
688  | 
show "?rhs \<subseteq> ?lhs"  | 
|
689  | 
using \<open>0 \<le> t\<close> \<open>t \<le> 1\<close> q01 qt by (force simp: path_image_subpath)  | 
|
690  | 
qed  | 
|
691  | 
    show "path_image (subpath 0 t q) \<inter> closed_segment (a - d) (a + e) = {a - d, a + e}" (is "?lhs = ?rhs")
 | 
|
692  | 
proof  | 
|
693  | 
show "?lhs \<subseteq> ?rhs" using paq_Int_cs pa01_Un by fastforce  | 
|
694  | 
show "?rhs \<subseteq> ?lhs" using \<open>0 \<le> t\<close> \<open>t \<le> 1\<close> q01 qt by (force simp: path_image_subpath)  | 
|
695  | 
qed  | 
|
696  | 
    show "path_image (subpath 1 t q) \<inter> closed_segment (a - d) (a + e) = {a - d, a + e}" (is "?lhs = ?rhs")
 | 
|
697  | 
proof  | 
|
698  | 
show "?lhs \<subseteq> ?rhs" by (auto simp: pa_01q [symmetric])  | 
|
699  | 
show "?rhs \<subseteq> ?lhs" using \<open>0 \<le> t\<close> \<open>t \<le> 1\<close> q01 qt by (force simp: path_image_subpath)  | 
|
700  | 
qed  | 
|
701  | 
    show "closed_segment (a - d) (a + e) \<inter> inside (path_image q) \<noteq> {}"
 | 
|
702  | 
using a a_in_de open_closed_segment pa01_Un q_eq_p by fastforce  | 
|
703  | 
show "z \<in> inside (path_image (subpath 0 t q) \<union> closed_segment (a - d) (a + e))"  | 
|
704  | 
by (metis path_image_join path_image_linepath path_image_subpath_commute pathfinish_subpath pathstart_linepath q01(1) zin)  | 
|
705  | 
show "winding_number (subpath 0 t q +++ linepath (a + e) (a - d)) z =  | 
|
706  | 
- winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z"  | 
|
707  | 
using z_notin_ed z_notin_0t \<open>0 \<le> t\<close> \<open>simple_path q\<close> \<open>t \<le> 1\<close>  | 
|
708  | 
by (simp add: simple_path_imp_path qt q01 path_image_subpath_commute closed_segment_commute winding_number_join winding_number_reversepath [symmetric])  | 
|
| 
67135
 
1a94352812f4
Moved material from AFP to Analysis/Number_Theory
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
66884 
diff
changeset
 | 
709  | 
show "- d \<noteq> e"  | 
| 65039 | 710  | 
using ad_not_ae by auto  | 
711  | 
show "winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z \<noteq> 0"  | 
|
712  | 
using z1 by auto  | 
|
713  | 
qed  | 
|
714  | 
show ?thesis  | 
|
715  | 
proof  | 
|
716  | 
show "z \<in> inside (path_image p)"  | 
|
717  | 
using q_eq_p z_in_q by auto  | 
|
718  | 
then have [simp]: "z \<notin> path_image q"  | 
|
719  | 
by (metis disjoint_iff_not_equal inside_no_overlap q_eq_p)  | 
|
720  | 
have [simp]: "z \<notin> path_image (subpath 1 t q)"  | 
|
721  | 
using inside_def pa01_Un z_in_q by fastforce  | 
|
722  | 
have "winding_number(subpath 0 t q +++ subpath t 1 q) z = winding_number(subpath 0 1 q) z"  | 
|
723  | 
using z_notin_0t \<open>0 \<le> t\<close> \<open>simple_path q\<close> \<open>t \<le> 1\<close>  | 
|
724  | 
by (simp add: simple_path_imp_path qt path_image_subpath_commute winding_number_join winding_number_subpath_combine)  | 
|
725  | 
with wn_q have "winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z = - winding_number q z"  | 
|
726  | 
by auto  | 
|
727  | 
with z1 have "cmod (winding_number q z) = 1"  | 
|
728  | 
by simp  | 
|
729  | 
with z1 wn_q_eq_wn_p show "cmod (winding_number p z) = 1"  | 
|
730  | 
using z1 wn_q_eq_wn_p by (simp add: \<open>z \<in> inside (path_image p)\<close>)  | 
|
731  | 
qed  | 
|
732  | 
qed  | 
|
733  | 
||
734  | 
||
735  | 
theorem simple_closed_path_winding_number_inside:  | 
|
736  | 
assumes "simple_path \<gamma>"  | 
|
737  | 
obtains "\<And>z. z \<in> inside(path_image \<gamma>) \<Longrightarrow> winding_number \<gamma> z = 1"  | 
|
738  | 
| "\<And>z. z \<in> inside(path_image \<gamma>) \<Longrightarrow> winding_number \<gamma> z = -1"  | 
|
739  | 
proof (cases "pathfinish \<gamma> = pathstart \<gamma>")  | 
|
740  | 
case True  | 
|
741  | 
have "path \<gamma>"  | 
|
742  | 
by (simp add: assms simple_path_imp_path)  | 
|
| 
66884
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
743  | 
then have const: "winding_number \<gamma> constant_on inside(path_image \<gamma>)"  | 
| 65039 | 744  | 
proof (rule winding_number_constant)  | 
745  | 
show "connected (inside(path_image \<gamma>))"  | 
|
746  | 
by (simp add: Jordan_inside_outside True assms)  | 
|
747  | 
qed (use inside_no_overlap True in auto)  | 
|
748  | 
obtain z where zin: "z \<in> inside (path_image \<gamma>)" and z1: "cmod (winding_number \<gamma> z) = 1"  | 
|
749  | 
using simple_closed_path_wn3 [of \<gamma>] True assms by blast  | 
|
750  | 
have "winding_number \<gamma> z \<in> \<int>"  | 
|
751  | 
using zin integer_winding_number [OF \<open>path \<gamma>\<close> True] inside_def by blast  | 
|
752  | 
with z1 consider "winding_number \<gamma> z = 1" | "winding_number \<gamma> z = -1"  | 
|
753  | 
apply (auto simp: Ints_def abs_if split: if_split_asm)  | 
|
754  | 
by (metis of_int_1 of_int_eq_iff of_int_minus)  | 
|
| 
66884
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
755  | 
with that const zin show ?thesis  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
756  | 
unfolding constant_on_def by metis  | 
| 65039 | 757  | 
next  | 
758  | 
case False  | 
|
759  | 
then show ?thesis  | 
|
760  | 
using inside_simple_curve_imp_closed assms that(2) by blast  | 
|
761  | 
qed  | 
|
762  | 
||
763  | 
corollary simple_closed_path_abs_winding_number_inside:  | 
|
764  | 
assumes "simple_path \<gamma>" "z \<in> inside(path_image \<gamma>)"  | 
|
765  | 
shows "\<bar>Re (winding_number \<gamma> z)\<bar> = 1"  | 
|
766  | 
by (metis assms norm_minus_cancel norm_one one_complex.simps(1) real_norm_def simple_closed_path_winding_number_inside uminus_complex.simps(1))  | 
|
767  | 
||
768  | 
corollary simple_closed_path_norm_winding_number_inside:  | 
|
769  | 
assumes "simple_path \<gamma>" "z \<in> inside(path_image \<gamma>)"  | 
|
770  | 
shows "norm (winding_number \<gamma> z) = 1"  | 
|
771  | 
proof -  | 
|
772  | 
have "pathfinish \<gamma> = pathstart \<gamma>"  | 
|
773  | 
using assms inside_simple_curve_imp_closed by blast  | 
|
774  | 
with assms integer_winding_number have "winding_number \<gamma> z \<in> \<int>"  | 
|
775  | 
by (simp add: inside_def simple_path_def)  | 
|
776  | 
then show ?thesis  | 
|
777  | 
by (metis assms norm_minus_cancel norm_one simple_closed_path_winding_number_inside)  | 
|
778  | 
qed  | 
|
779  | 
||
780  | 
corollary simple_closed_path_winding_number_cases:  | 
|
781  | 
   "\<lbrakk>simple_path \<gamma>; pathfinish \<gamma> = pathstart \<gamma>; z \<notin> path_image \<gamma>\<rbrakk> \<Longrightarrow> winding_number \<gamma> z \<in> {-1,0,1}"
 | 
|
782  | 
apply (simp add: inside_Un_outside [of "path_image \<gamma>", symmetric, unfolded set_eq_iff Set.Compl_iff] del: inside_Un_outside)  | 
|
783  | 
apply (rule simple_closed_path_winding_number_inside)  | 
|
784  | 
using simple_path_def winding_number_zero_in_outside by blast+  | 
|
785  | 
||
786  | 
corollary simple_closed_path_winding_number_pos:  | 
|
787  | 
"\<lbrakk>simple_path \<gamma>; pathfinish \<gamma> = pathstart \<gamma>; z \<notin> path_image \<gamma>; 0 < Re(winding_number \<gamma> z)\<rbrakk>  | 
|
788  | 
\<Longrightarrow> winding_number \<gamma> z = 1"  | 
|
789  | 
using simple_closed_path_winding_number_cases  | 
|
790  | 
by fastforce  | 
|
791  | 
||
| 
66393
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
792  | 
|
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
793  | 
subsection \<open>Winding number for rectangular paths\<close>  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
794  | 
|
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
795  | 
(* TODO: Move *)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
796  | 
lemma closed_segmentI:  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
797  | 
  "u \<in> {0..1} \<Longrightarrow> z = (1 - u) *\<^sub>R a + u *\<^sub>R b \<Longrightarrow> z \<in> closed_segment a b"
 | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
798  | 
by (auto simp: closed_segment_def)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
799  | 
|
| 
66827
 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
66826 
diff
changeset
 | 
800  | 
lemma in_cbox_complex_iff:  | 
| 
66393
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
801  | 
  "x \<in> cbox a b \<longleftrightarrow> Re x \<in> {Re a..Re b} \<and> Im x \<in> {Im a..Im b}"
 | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
802  | 
by (cases x; cases a; cases b) (auto simp: cbox_Complex_eq)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
803  | 
|
| 
66827
 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
66826 
diff
changeset
 | 
804  | 
lemma box_Complex_eq:  | 
| 
66393
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
805  | 
"box (Complex a c) (Complex b d) = (\<lambda>(x,y). Complex x y) ` (box a b \<times> box c d)"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
806  | 
by (auto simp: box_def Basis_complex_def image_iff complex_eq_iff)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
807  | 
|
| 
66827
 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
66826 
diff
changeset
 | 
808  | 
lemma in_box_complex_iff:  | 
| 
66393
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
809  | 
  "x \<in> box a b \<longleftrightarrow> Re x \<in> {Re a<..<Re b} \<and> Im x \<in> {Im a<..<Im b}"
 | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
810  | 
by (cases x; cases a; cases b) (auto simp: box_Complex_eq)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
811  | 
(* END TODO *)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
812  | 
|
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
813  | 
lemma closed_segment_same_Re:  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
814  | 
assumes "Re a = Re b"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
815  | 
  shows   "closed_segment a b = {z. Re z = Re a \<and> Im z \<in> closed_segment (Im a) (Im b)}"
 | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
816  | 
proof safe  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
817  | 
fix z assume "z \<in> closed_segment a b"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
818  | 
  then obtain u where u: "u \<in> {0..1}" "z = a + of_real u * (b - a)"
 | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
819  | 
by (auto simp: closed_segment_def scaleR_conv_of_real algebra_simps)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
820  | 
from assms show "Re z = Re a" by (auto simp: u)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
821  | 
from u(1) show "Im z \<in> closed_segment (Im a) (Im b)"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
822  | 
by (intro closed_segmentI[of u]) (auto simp: u algebra_simps)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
823  | 
next  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
824  | 
fix z assume [simp]: "Re z = Re a" and "Im z \<in> closed_segment (Im a) (Im b)"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
825  | 
  then obtain u where u: "u \<in> {0..1}" "Im z = Im a + of_real u * (Im b - Im a)"
 | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
826  | 
by (auto simp: closed_segment_def scaleR_conv_of_real algebra_simps)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
827  | 
from u(1) show "z \<in> closed_segment a b" using assms  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
828  | 
by (intro closed_segmentI[of u]) (auto simp: u algebra_simps scaleR_conv_of_real complex_eq_iff)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
829  | 
qed  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
830  | 
|
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
831  | 
lemma closed_segment_same_Im:  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
832  | 
assumes "Im a = Im b"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
833  | 
  shows   "closed_segment a b = {z. Im z = Im a \<and> Re z \<in> closed_segment (Re a) (Re b)}"
 | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
834  | 
proof safe  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
835  | 
fix z assume "z \<in> closed_segment a b"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
836  | 
  then obtain u where u: "u \<in> {0..1}" "z = a + of_real u * (b - a)"
 | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
837  | 
by (auto simp: closed_segment_def scaleR_conv_of_real algebra_simps)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
838  | 
from assms show "Im z = Im a" by (auto simp: u)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
839  | 
from u(1) show "Re z \<in> closed_segment (Re a) (Re b)"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
840  | 
by (intro closed_segmentI[of u]) (auto simp: u algebra_simps)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
841  | 
next  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
842  | 
fix z assume [simp]: "Im z = Im a" and "Re z \<in> closed_segment (Re a) (Re b)"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
843  | 
  then obtain u where u: "u \<in> {0..1}" "Re z = Re a + of_real u * (Re b - Re a)"
 | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
844  | 
by (auto simp: closed_segment_def scaleR_conv_of_real algebra_simps)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
845  | 
from u(1) show "z \<in> closed_segment a b" using assms  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
846  | 
by (intro closed_segmentI[of u]) (auto simp: u algebra_simps scaleR_conv_of_real complex_eq_iff)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
847  | 
qed  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
848  | 
|
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
849  | 
|
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
850  | 
definition rectpath where  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
851  | 
"rectpath a1 a3 = (let a2 = Complex (Re a3) (Im a1); a4 = Complex (Re a1) (Im a3)  | 
| 
66827
 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
66826 
diff
changeset
 | 
852  | 
in linepath a1 a2 +++ linepath a2 a3 +++ linepath a3 a4 +++ linepath a4 a1)"  | 
| 
66393
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
853  | 
|
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
854  | 
lemma path_rectpath [simp, intro]: "path (rectpath a b)"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
855  | 
by (simp add: Let_def rectpath_def)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
856  | 
|
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
857  | 
lemma valid_path_rectpath [simp, intro]: "valid_path (rectpath a b)"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
858  | 
by (simp add: Let_def rectpath_def)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
859  | 
|
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
860  | 
lemma pathstart_rectpath [simp]: "pathstart (rectpath a1 a3) = a1"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
861  | 
by (simp add: rectpath_def Let_def)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
862  | 
|
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
863  | 
lemma pathfinish_rectpath [simp]: "pathfinish (rectpath a1 a3) = a1"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
864  | 
by (simp add: rectpath_def Let_def)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
865  | 
|
| 
66827
 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
66826 
diff
changeset
 | 
866  | 
lemma simple_path_rectpath [simp, intro]:  | 
| 
66393
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
867  | 
assumes "Re a1 \<noteq> Re a3" "Im a1 \<noteq> Im a3"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
868  | 
shows "simple_path (rectpath a1 a3)"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
869  | 
unfolding rectpath_def Let_def using assms  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
870  | 
by (intro simple_path_join_loop arc_join arc_linepath)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
871  | 
(auto simp: complex_eq_iff path_image_join closed_segment_same_Re closed_segment_same_Im)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
872  | 
|
| 
66827
 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
66826 
diff
changeset
 | 
873  | 
lemma path_image_rectpath:  | 
| 
66393
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
874  | 
assumes "Re a1 \<le> Re a3" "Im a1 \<le> Im a3"  | 
| 
66827
 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
66826 
diff
changeset
 | 
875  | 
shows "path_image (rectpath a1 a3) =  | 
| 
66393
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
876  | 
           {z. Re z \<in> {Re a1, Re a3} \<and> Im z \<in> {Im a1..Im a3}} \<union>
 | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
877  | 
           {z. Im z \<in> {Im a1, Im a3} \<and> Re z \<in> {Re a1..Re a3}}" (is "?lhs = ?rhs")
 | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
878  | 
proof -  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
879  | 
define a2 a4 where "a2 = Complex (Re a3) (Im a1)" and "a4 = Complex (Re a1) (Im a3)"  | 
| 
66827
 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
66826 
diff
changeset
 | 
880  | 
have "?lhs = closed_segment a1 a2 \<union> closed_segment a2 a3 \<union>  | 
| 
66393
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
881  | 
closed_segment a4 a3 \<union> closed_segment a1 a4"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
882  | 
by (simp_all add: rectpath_def Let_def path_image_join closed_segment_commute  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
883  | 
a2_def a4_def Un_assoc)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
884  | 
also have "\<dots> = ?rhs" using assms  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
885  | 
by (auto simp: rectpath_def Let_def path_image_join a2_def a4_def  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
886  | 
closed_segment_same_Re closed_segment_same_Im closed_segment_eq_real_ivl)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
887  | 
finally show ?thesis .  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
888  | 
qed  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
889  | 
|
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
890  | 
lemma path_image_rectpath_subset_cbox:  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
891  | 
assumes "Re a \<le> Re b" "Im a \<le> Im b"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
892  | 
shows "path_image (rectpath a b) \<subseteq> cbox a b"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
893  | 
using assms by (auto simp: path_image_rectpath in_cbox_complex_iff)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
894  | 
|
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
895  | 
lemma path_image_rectpath_inter_box:  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
896  | 
assumes "Re a \<le> Re b" "Im a \<le> Im b"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
897  | 
  shows   "path_image (rectpath a b) \<inter> box a b = {}"
 | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
898  | 
using assms by (auto simp: path_image_rectpath in_box_complex_iff)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
899  | 
|
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
900  | 
lemma path_image_rectpath_cbox_minus_box:  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
901  | 
assumes "Re a \<le> Re b" "Im a \<le> Im b"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
902  | 
shows "path_image (rectpath a b) = cbox a b - box a b"  | 
| 
66827
 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
66826 
diff
changeset
 | 
903  | 
using assms by (auto simp: path_image_rectpath in_cbox_complex_iff  | 
| 
66393
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
904  | 
in_box_complex_iff)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
905  | 
|
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
906  | 
lemma winding_number_rectpath:  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
907  | 
assumes "z \<in> box a1 a3"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
908  | 
shows "winding_number (rectpath a1 a3) z = 1"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
909  | 
proof -  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
910  | 
from assms have less: "Re a1 < Re a3" "Im a1 < Im a3"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
911  | 
by (auto simp: in_box_complex_iff)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
912  | 
define a2 a4 where "a2 = Complex (Re a3) (Im a1)" and "a4 = Complex (Re a1) (Im a3)"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
913  | 
let ?l1 = "linepath a1 a2" and ?l2 = "linepath a2 a3"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
914  | 
and ?l3 = "linepath a3 a4" and ?l4 = "linepath a4 a1"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
915  | 
from assms and less have "z \<notin> path_image (rectpath a1 a3)"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
916  | 
by (auto simp: path_image_rectpath_cbox_minus_box)  | 
| 
66827
 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
66826 
diff
changeset
 | 
917  | 
also have "path_image (rectpath a1 a3) =  | 
| 
66393
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
918  | 
path_image ?l1 \<union> path_image ?l2 \<union> path_image ?l3 \<union> path_image ?l4"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
919  | 
by (simp add: rectpath_def Let_def path_image_join Un_assoc a2_def a4_def)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
920  | 
finally have "z \<notin> \<dots>" .  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
921  | 
  moreover have "\<forall>l\<in>{?l1,?l2,?l3,?l4}. Re (winding_number l z) > 0"
 | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
922  | 
unfolding ball_simps HOL.simp_thms a2_def a4_def  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
923  | 
by (intro conjI; (rule winding_number_linepath_pos_lt;  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
924  | 
(insert assms, auto simp: a2_def a4_def in_box_complex_iff mult_neg_neg))+)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
925  | 
ultimately have "Re (winding_number (rectpath a1 a3) z) > 0"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
926  | 
by (simp add: winding_number_join path_image_join rectpath_def Let_def a2_def a4_def)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
927  | 
thus "winding_number (rectpath a1 a3) z = 1" using assms less  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
928  | 
by (intro simple_closed_path_winding_number_pos simple_path_rectpath)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
929  | 
(auto simp: path_image_rectpath_cbox_minus_box)  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
930  | 
qed  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
931  | 
|
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
932  | 
lemma winding_number_rectpath_outside:  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
933  | 
assumes "Re a1 \<le> Re a3" "Im a1 \<le> Im a3"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
934  | 
assumes "z \<notin> cbox a1 a3"  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
935  | 
shows "winding_number (rectpath a1 a3) z = 0"  | 
| 
66827
 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 
paulson <lp15@cam.ac.uk> 
parents: 
66826 
diff
changeset
 | 
936  | 
using assms by (intro winding_number_zero_outside[OF _ _ _ assms(3)]  | 
| 
66393
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
937  | 
path_image_rectpath_subset_cbox) simp_all  | 
| 
 
2a6371fb31f0
Winding numbers for rectangular paths
 
eberlm <eberlm@in.tum.de> 
parents: 
66304 
diff
changeset
 | 
938  | 
|
| 
66826
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
939  | 
|
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
940  | 
text\<open>A per-function version for continuous logs, a kind of monodromy\<close>  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
941  | 
|
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
942  | 
proposition winding_number_compose_exp:  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
943  | 
assumes "path p"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
944  | 
shows "winding_number (exp \<circ> p) 0 = (pathfinish p - pathstart p) / (2 * of_real pi * \<i>)"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
945  | 
proof -  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
946  | 
  obtain e where "0 < e" and e: "\<And>t. t \<in> {0..1} \<Longrightarrow> e \<le> norm(exp(p t))"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
947  | 
proof  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
948  | 
have "closed (path_image (exp \<circ> p))"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
949  | 
by (simp add: assms closed_path_image holomorphic_on_exp holomorphic_on_imp_continuous_on path_continuous_image)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
950  | 
    then show "0 < setdist {0} (path_image (exp \<circ> p))"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
951  | 
by (metis (mono_tags, lifting) compact_sing exp_not_eq_zero imageE path_image_compose  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
952  | 
path_image_nonempty setdist_eq_0_compact_closed setdist_gt_0_compact_closed setdist_eq_0_closed)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
953  | 
next  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
954  | 
fix t::real  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
955  | 
    assume "t \<in> {0..1}"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
956  | 
    have "setdist {0} (path_image (exp \<circ> p)) \<le> dist 0 (exp (p t))"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
957  | 
apply (rule setdist_le_dist)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
958  | 
      using \<open>t \<in> {0..1}\<close> path_image_def by fastforce+
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
959  | 
    then show "setdist {0} (path_image (exp \<circ> p)) \<le> cmod (exp (p t))"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
960  | 
by simp  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
961  | 
qed  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
962  | 
have "bounded (path_image p)"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
963  | 
by (simp add: assms bounded_path_image)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
964  | 
then obtain B where "0 < B" and B: "path_image p \<subseteq> cball 0 B"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
965  | 
by (meson bounded_pos mem_cball_0 subsetI)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
966  | 
let ?B = "cball (0::complex) (B+1)"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
967  | 
have "uniformly_continuous_on ?B exp"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
968  | 
using holomorphic_on_exp holomorphic_on_imp_continuous_on  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
969  | 
by (force intro: compact_uniformly_continuous)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
970  | 
then obtain d where "d > 0"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
971  | 
and d: "\<And>x x'. \<lbrakk>x\<in>?B; x'\<in>?B; dist x' x < d\<rbrakk> \<Longrightarrow> norm (exp x' - exp x) < e"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
972  | 
using \<open>e > 0\<close> by (auto simp: uniformly_continuous_on_def dist_norm)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
973  | 
then have "min 1 d > 0"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
974  | 
by force  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
975  | 
then obtain g where pfg: "polynomial_function g" and "g 0 = p 0" "g 1 = p 1"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
976  | 
           and gless: "\<And>t. t \<in> {0..1} \<Longrightarrow> norm(g t - p t) < min 1 d"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
977  | 
using path_approx_polynomial_function [OF \<open>path p\<close>] \<open>d > 0\<close> \<open>0 < e\<close>  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
978  | 
unfolding pathfinish_def pathstart_def by meson  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
979  | 
have "winding_number (exp \<circ> p) 0 = winding_number (exp \<circ> g) 0"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
980  | 
proof (rule winding_number_nearby_paths_eq [symmetric])  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
981  | 
show "path (exp \<circ> p)" "path (exp \<circ> g)"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
982  | 
by (simp_all add: pfg assms holomorphic_on_exp holomorphic_on_imp_continuous_on path_continuous_image path_polynomial_function)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
983  | 
next  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
984  | 
fix t :: "real"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
985  | 
    assume t: "t \<in> {0..1}"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
986  | 
with gless have "norm(g t - p t) < 1"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
987  | 
using min_less_iff_conj by blast  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
988  | 
moreover have ptB: "norm (p t) \<le> B"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
989  | 
using B t by (force simp: path_image_def)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
990  | 
ultimately have "cmod (g t) \<le> B + 1"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
991  | 
by (meson add_mono_thms_linordered_field(4) le_less_trans less_imp_le norm_triangle_sub)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
992  | 
with ptB gless t have "cmod ((exp \<circ> g) t - (exp \<circ> p) t) < e"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
993  | 
by (auto simp: dist_norm d)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
994  | 
with e t show "cmod ((exp \<circ> g) t - (exp \<circ> p) t) < cmod ((exp \<circ> p) t - 0)"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
995  | 
by fastforce  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
996  | 
qed (use \<open>g 0 = p 0\<close> \<open>g 1 = p 1\<close> in \<open>auto simp: pathfinish_def pathstart_def\<close>)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
997  | 
also have "... = 1 / (of_real (2 * pi) * \<i>) * contour_integral (exp \<circ> g) (\<lambda>w. 1 / (w - 0))"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
998  | 
proof (rule winding_number_valid_path)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
999  | 
have "continuous_on (path_image g) (deriv exp)"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1000  | 
by (metis DERIV_exp DERIV_imp_deriv continuous_on_cong holomorphic_on_exp holomorphic_on_imp_continuous_on)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1001  | 
then show "valid_path (exp \<circ> g)"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1002  | 
by (simp add: field_differentiable_within_exp pfg valid_path_compose valid_path_polynomial_function)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1003  | 
show "0 \<notin> path_image (exp \<circ> g)"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1004  | 
by (auto simp: path_image_def)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1005  | 
qed  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1006  | 
  also have "... = 1 / (of_real (2 * pi) * \<i>) * integral {0..1} (\<lambda>x. vector_derivative g (at x))"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1007  | 
proof (simp add: contour_integral_integral, rule integral_cong)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1008  | 
fix t :: "real"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1009  | 
    assume t: "t \<in> {0..1}"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1010  | 
show "vector_derivative (exp \<circ> g) (at t) / exp (g t) = vector_derivative g (at t)"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1011  | 
proof (simp add: divide_simps, rule vector_derivative_unique_at)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1012  | 
show "(exp \<circ> g has_vector_derivative vector_derivative (exp \<circ> g) (at t)) (at t)"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1013  | 
by (meson DERIV_exp differentiable_def field_vector_diff_chain_at has_vector_derivative_def  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1014  | 
has_vector_derivative_polynomial_function pfg vector_derivative_works)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1015  | 
show "(exp \<circ> g has_vector_derivative vector_derivative g (at t) * exp (g t)) (at t)"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1016  | 
apply (rule field_vector_diff_chain_at)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1017  | 
apply (metis has_vector_derivative_polynomial_function pfg vector_derivative_at)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1018  | 
using DERIV_exp has_field_derivative_def apply blast  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1019  | 
done  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1020  | 
qed  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1021  | 
qed  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1022  | 
also have "... = (pathfinish p - pathstart p) / (2 * of_real pi * \<i>)"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1023  | 
proof -  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1024  | 
    have "((\<lambda>x. vector_derivative g (at x)) has_integral g 1 - g 0) {0..1}"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1025  | 
apply (rule fundamental_theorem_of_calculus [OF zero_le_one])  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1026  | 
by (metis has_vector_derivative_at_within has_vector_derivative_polynomial_function pfg vector_derivative_at)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1027  | 
then show ?thesis  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1028  | 
apply (simp add: pathfinish_def pathstart_def)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1029  | 
using \<open>g 0 = p 0\<close> \<open>g 1 = p 1\<close> by auto  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1030  | 
qed  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1031  | 
finally show ?thesis .  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1032  | 
qed  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1033  | 
|
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1034  | 
|
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1035  | 
|
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1036  | 
subsection\<open>The winding number defines a continuous logarithm for the path itself\<close>  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1037  | 
|
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1038  | 
lemma winding_number_as_continuous_log:  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1039  | 
assumes "path p" and \<zeta>: "\<zeta> \<notin> path_image p"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1040  | 
obtains q where "path q"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1041  | 
"pathfinish q - pathstart q = 2 * of_real pi * \<i> * winding_number p \<zeta>"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1042  | 
                  "\<And>t. t \<in> {0..1} \<Longrightarrow> p t = \<zeta> + exp(q t)"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1043  | 
proof -  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1044  | 
let ?q = "\<lambda>t. 2 * of_real pi * \<i> * winding_number(subpath 0 t p) \<zeta> + Ln(pathstart p - \<zeta>)"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1045  | 
show ?thesis  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1046  | 
proof  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1047  | 
    have *: "continuous (at t within {0..1}) (\<lambda>x. winding_number (subpath 0 x p) \<zeta>)"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1048  | 
      if t: "t \<in> {0..1}" for t
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1049  | 
proof -  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1050  | 
let ?B = "ball (p t) (norm(p t - \<zeta>))"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1051  | 
have "p t \<noteq> \<zeta>"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1052  | 
using path_image_def that \<zeta> by blast  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1053  | 
then have "simply_connected ?B"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1054  | 
by (simp add: convex_imp_simply_connected)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1055  | 
then have "\<And>f::complex\<Rightarrow>complex. continuous_on ?B f \<and> (\<forall>\<zeta> \<in> ?B. f \<zeta> \<noteq> 0)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1056  | 
\<longrightarrow> (\<exists>g. continuous_on ?B g \<and> (\<forall>\<zeta> \<in> ?B. f \<zeta> = exp (g \<zeta>)))"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1057  | 
by (simp add: simply_connected_eq_continuous_log)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1058  | 
moreover have "continuous_on ?B (\<lambda>w. w - \<zeta>)"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1059  | 
by (intro continuous_intros)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1060  | 
moreover have "(\<forall>z \<in> ?B. z - \<zeta> \<noteq> 0)"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1061  | 
by (auto simp: dist_norm)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1062  | 
ultimately obtain g where contg: "continuous_on ?B g"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1063  | 
and geq: "\<And>z. z \<in> ?B \<Longrightarrow> z - \<zeta> = exp (g z)" by blast  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1064  | 
obtain d where "0 < d" and d:  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1065  | 
        "\<And>x. \<lbrakk>x \<in> {0..1}; dist x t < d\<rbrakk> \<Longrightarrow> dist (p x) (p t) < cmod (p t - \<zeta>)"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1066  | 
using \<open>path p\<close> t unfolding path_def continuous_on_iff  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1067  | 
by (metis \<open>p t \<noteq> \<zeta>\<close> right_minus_eq zero_less_norm_iff)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1068  | 
have "((\<lambda>x. winding_number (\<lambda>w. subpath 0 x p w - \<zeta>) 0 -  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1069  | 
winding_number (\<lambda>w. subpath 0 t p w - \<zeta>) 0) \<longlongrightarrow> 0)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1070  | 
            (at t within {0..1})"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1071  | 
proof (rule Lim_transform_within [OF _ \<open>d > 0\<close>])  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1072  | 
        have "continuous (at t within {0..1}) (g o p)"
 | 
| 
68077
 
ee8c13ae81e9
Some tidying up (mostly regarding summations from 0)
 
paulson <lp15@cam.ac.uk> 
parents: 
67135 
diff
changeset
 | 
1073  | 
proof (rule continuous_within_compose)  | 
| 
 
ee8c13ae81e9
Some tidying up (mostly regarding summations from 0)
 
paulson <lp15@cam.ac.uk> 
parents: 
67135 
diff
changeset
 | 
1074  | 
          show "continuous (at t within {0..1}) p"
 | 
| 
 
ee8c13ae81e9
Some tidying up (mostly regarding summations from 0)
 
paulson <lp15@cam.ac.uk> 
parents: 
67135 
diff
changeset
 | 
1075  | 
using \<open>path p\<close> continuous_on_eq_continuous_within path_def that by blast  | 
| 
 
ee8c13ae81e9
Some tidying up (mostly regarding summations from 0)
 
paulson <lp15@cam.ac.uk> 
parents: 
67135 
diff
changeset
 | 
1076  | 
          show "continuous (at (p t) within p ` {0..1}) g"
 | 
| 
 
ee8c13ae81e9
Some tidying up (mostly regarding summations from 0)
 
paulson <lp15@cam.ac.uk> 
parents: 
67135 
diff
changeset
 | 
1077  | 
by (metis (no_types, lifting) open_ball UNIV_I \<open>p t \<noteq> \<zeta>\<close> centre_in_ball contg continuous_on_eq_continuous_at continuous_within_topological right_minus_eq zero_less_norm_iff)  | 
| 
 
ee8c13ae81e9
Some tidying up (mostly regarding summations from 0)
 
paulson <lp15@cam.ac.uk> 
parents: 
67135 
diff
changeset
 | 
1078  | 
qed  | 
| 
66826
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1079  | 
        with LIM_zero have "((\<lambda>u. (g (subpath t u p 1) - g (subpath t u p 0))) \<longlongrightarrow> 0) (at t within {0..1})"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1080  | 
by (auto simp: subpath_def continuous_within o_def)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1081  | 
then show "((\<lambda>u. (g (subpath t u p 1) - g (subpath t u p 0)) / (2 * of_real pi * \<i>)) \<longlongrightarrow> 0)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1082  | 
           (at t within {0..1})"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1083  | 
by (simp add: tendsto_divide_zero)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1084  | 
show "(g (subpath t u p 1) - g (subpath t u p 0)) / (2 * of_real pi * \<i>) =  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1085  | 
winding_number (\<lambda>w. subpath 0 u p w - \<zeta>) 0 - winding_number (\<lambda>w. subpath 0 t p w - \<zeta>) 0"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1086  | 
          if "u \<in> {0..1}" "0 < dist u t" "dist u t < d" for u
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1087  | 
proof -  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1088  | 
          have "closed_segment t u \<subseteq> {0..1}"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1089  | 
using closed_segment_eq_real_ivl t that by auto  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1090  | 
then have piB: "path_image(subpath t u p) \<subseteq> ?B"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1091  | 
apply (clarsimp simp add: path_image_subpath_gen)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1092  | 
by (metis subsetD le_less_trans \<open>dist u t < d\<close> d dist_commute dist_in_closed_segment)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1093  | 
have *: "path (g \<circ> subpath t u p)"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1094  | 
apply (rule path_continuous_image)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1095  | 
using \<open>path p\<close> t that apply auto[1]  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1096  | 
using piB contg continuous_on_subset by blast  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1097  | 
have "(g (subpath t u p 1) - g (subpath t u p 0)) / (2 * of_real pi * \<i>)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1098  | 
= winding_number (exp \<circ> g \<circ> subpath t u p) 0"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1099  | 
using winding_number_compose_exp [OF *]  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1100  | 
by (simp add: pathfinish_def pathstart_def o_assoc)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1101  | 
also have "... = winding_number (\<lambda>w. subpath t u p w - \<zeta>) 0"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1102  | 
proof (rule winding_number_cong)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1103  | 
            have "exp(g y) = y - \<zeta>" if "y \<in> (subpath t u p) ` {0..1}" for y
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1104  | 
by (metis that geq path_image_def piB subset_eq)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1105  | 
then show "\<And>x. \<lbrakk>0 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> (exp \<circ> g \<circ> subpath t u p) x = subpath t u p x - \<zeta>"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1106  | 
by auto  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1107  | 
qed  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1108  | 
also have "... = winding_number (\<lambda>w. subpath 0 u p w - \<zeta>) 0 -  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1109  | 
winding_number (\<lambda>w. subpath 0 t p w - \<zeta>) 0"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1110  | 
apply (simp add: winding_number_offset [symmetric])  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1111  | 
            using winding_number_subpath_combine [OF \<open>path p\<close> \<zeta>, of 0 t u] \<open>t \<in> {0..1}\<close> \<open>u \<in> {0..1}\<close>
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1112  | 
by (simp add: add.commute eq_diff_eq)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1113  | 
finally show ?thesis .  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1114  | 
qed  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1115  | 
qed  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1116  | 
then show ?thesis  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1117  | 
by (subst winding_number_offset) (simp add: continuous_within LIM_zero_iff)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1118  | 
qed  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1119  | 
show "path ?q"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1120  | 
unfolding path_def  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1121  | 
by (intro continuous_intros) (simp add: continuous_on_eq_continuous_within *)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1122  | 
|
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1123  | 
have "\<zeta> \<noteq> p 0"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1124  | 
by (metis \<zeta> pathstart_def pathstart_in_path_image)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1125  | 
then show "pathfinish ?q - pathstart ?q = 2 * of_real pi * \<i> * winding_number p \<zeta>"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1126  | 
by (simp add: pathfinish_def pathstart_def)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1127  | 
    show "p t = \<zeta> + exp (?q t)" if "t \<in> {0..1}" for t
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1128  | 
proof -  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1129  | 
have "path (subpath 0 t p)"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1130  | 
using \<open>path p\<close> that by auto  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1131  | 
moreover  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1132  | 
have "\<zeta> \<notin> path_image (subpath 0 t p)"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1133  | 
using \<zeta> [unfolded path_image_def] that by (auto simp: path_image_subpath)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1134  | 
ultimately show ?thesis  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1135  | 
using winding_number_exp_2pi [of "subpath 0 t p" \<zeta>] \<open>\<zeta> \<noteq> p 0\<close>  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1136  | 
by (auto simp: exp_add algebra_simps pathfinish_def pathstart_def subpath_def)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1137  | 
qed  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1138  | 
qed  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1139  | 
qed  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1140  | 
|
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1141  | 
|
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1142  | 
subsection\<open>Winding number equality is the same as path/loop homotopy in C - {0}\<close>
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1143  | 
|
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1144  | 
lemma winding_number_homotopic_loops_null_eq:  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1145  | 
assumes "path p" and \<zeta>: "\<zeta> \<notin> path_image p"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1146  | 
  shows "winding_number p \<zeta> = 0 \<longleftrightarrow> (\<exists>a. homotopic_loops (-{\<zeta>}) p (\<lambda>t. a))"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1147  | 
(is "?lhs = ?rhs")  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1148  | 
proof  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1149  | 
assume [simp]: ?lhs  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1150  | 
obtain q where "path q"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1151  | 
and qeq: "pathfinish q - pathstart q = 2 * of_real pi * \<i> * winding_number p \<zeta>"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1152  | 
             and peq: "\<And>t. t \<in> {0..1} \<Longrightarrow> p t = \<zeta> + exp(q t)"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1153  | 
using winding_number_as_continuous_log [OF assms] by blast  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1154  | 
have *: "homotopic_with (\<lambda>r. pathfinish r = pathstart r)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1155  | 
                       {0..1} (-{\<zeta>}) ((\<lambda>w. \<zeta> + exp w) \<circ> q) ((\<lambda>w. \<zeta> + exp w) \<circ> (\<lambda>t. 0))"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1156  | 
proof (rule homotopic_with_compose_continuous_left)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1157  | 
show "homotopic_with (\<lambda>f. pathfinish ((\<lambda>w. \<zeta> + exp w) \<circ> f) = pathstart ((\<lambda>w. \<zeta> + exp w) \<circ> f))  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1158  | 
              {0..1} UNIV q (\<lambda>t. 0)"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1159  | 
proof (rule homotopic_with_mono, simp_all add: pathfinish_def pathstart_def)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1160  | 
have "homotopic_loops UNIV q (\<lambda>t. 0)"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1161  | 
by (rule homotopic_loops_linear) (use qeq \<open>path q\<close> in \<open>auto simp: continuous_on_const path_defs\<close>)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1162  | 
      then show "homotopic_with (\<lambda>h. exp (h 1) = exp (h 0)) {0..1} UNIV q (\<lambda>t. 0)"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1163  | 
by (simp add: homotopic_loops_def homotopic_with_mono pathfinish_def pathstart_def)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1164  | 
qed  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1165  | 
show "continuous_on UNIV (\<lambda>w. \<zeta> + exp w)"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1166  | 
by (rule continuous_intros)+  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1167  | 
    show "range (\<lambda>w. \<zeta> + exp w) \<subseteq> -{\<zeta>}"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1168  | 
by auto  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1169  | 
qed  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1170  | 
  then have "homotopic_with (\<lambda>r. pathfinish r = pathstart r) {0..1} (-{\<zeta>}) p (\<lambda>x. \<zeta> + 1)"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1171  | 
by (rule homotopic_with_eq) (auto simp: o_def peq pathfinish_def pathstart_def)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1172  | 
  then have "homotopic_loops (-{\<zeta>}) p (\<lambda>t. \<zeta> + 1)"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1173  | 
by (simp add: homotopic_loops_def)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1174  | 
then show ?rhs ..  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1175  | 
next  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1176  | 
assume ?rhs  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1177  | 
  then obtain a where "homotopic_loops (-{\<zeta>}) p (\<lambda>t. a)" ..
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1178  | 
then have "winding_number p \<zeta> = winding_number (\<lambda>t. a) \<zeta>" "a \<noteq> \<zeta>"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1179  | 
using winding_number_homotopic_loops homotopic_loops_imp_subset by (force simp:)+  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1180  | 
moreover have "winding_number (\<lambda>t. a) \<zeta> = 0"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1181  | 
by (metis winding_number_zero_const \<open>a \<noteq> \<zeta>\<close>)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1182  | 
ultimately show ?lhs by metis  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1183  | 
qed  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1184  | 
|
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1185  | 
|
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1186  | 
lemma winding_number_homotopic_paths_null_explicit_eq:  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1187  | 
assumes "path p" and \<zeta>: "\<zeta> \<notin> path_image p"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1188  | 
  shows "winding_number p \<zeta> = 0 \<longleftrightarrow> homotopic_paths (-{\<zeta>}) p (linepath (pathstart p) (pathstart p))"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1189  | 
(is "?lhs = ?rhs")  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1190  | 
proof  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1191  | 
assume ?lhs  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1192  | 
then show ?rhs  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1193  | 
apply (auto simp: winding_number_homotopic_loops_null_eq [OF assms])  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1194  | 
apply (rule homotopic_loops_imp_homotopic_paths_null)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1195  | 
apply (simp add: linepath_refl)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1196  | 
done  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1197  | 
next  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1198  | 
assume ?rhs  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1199  | 
then show ?lhs  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1200  | 
by (metis \<zeta> pathstart_in_path_image winding_number_homotopic_paths winding_number_trivial)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1201  | 
qed  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1202  | 
|
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1203  | 
|
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1204  | 
lemma winding_number_homotopic_paths_null_eq:  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1205  | 
assumes "path p" and \<zeta>: "\<zeta> \<notin> path_image p"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1206  | 
  shows "winding_number p \<zeta> = 0 \<longleftrightarrow> (\<exists>a. homotopic_paths (-{\<zeta>}) p (\<lambda>t. a))"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1207  | 
(is "?lhs = ?rhs")  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1208  | 
proof  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1209  | 
assume ?lhs  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1210  | 
then show ?rhs  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1211  | 
by (auto simp: winding_number_homotopic_paths_null_explicit_eq [OF assms] linepath_refl)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1212  | 
next  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1213  | 
assume ?rhs  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1214  | 
then show ?lhs  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1215  | 
by (metis \<zeta> homotopic_paths_imp_pathfinish pathfinish_def pathfinish_in_path_image winding_number_homotopic_paths winding_number_zero_const)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1216  | 
qed  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1217  | 
|
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1218  | 
|
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1219  | 
lemma winding_number_homotopic_paths_eq:  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1220  | 
assumes "path p" and \<zeta>p: "\<zeta> \<notin> path_image p"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1221  | 
and "path q" and \<zeta>q: "\<zeta> \<notin> path_image q"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1222  | 
and qp: "pathstart q = pathstart p" "pathfinish q = pathfinish p"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1223  | 
    shows "winding_number p \<zeta> = winding_number q \<zeta> \<longleftrightarrow> homotopic_paths (-{\<zeta>}) p q"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1224  | 
(is "?lhs = ?rhs")  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1225  | 
proof  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1226  | 
assume ?lhs  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1227  | 
then have "winding_number (p +++ reversepath q) \<zeta> = 0"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1228  | 
using assms by (simp add: winding_number_join winding_number_reversepath)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1229  | 
moreover  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1230  | 
have "path (p +++ reversepath q)" "\<zeta> \<notin> path_image (p +++ reversepath q)"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1231  | 
using assms by (auto simp: not_in_path_image_join)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1232  | 
  ultimately obtain a where "homotopic_paths (- {\<zeta>}) (p +++ reversepath q) (linepath a a)"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1233  | 
using winding_number_homotopic_paths_null_explicit_eq by blast  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1234  | 
then show ?rhs  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1235  | 
using homotopic_paths_imp_pathstart assms  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1236  | 
by (fastforce simp add: dest: homotopic_paths_imp_homotopic_loops homotopic_paths_loop_parts)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1237  | 
next  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1238  | 
assume ?rhs  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1239  | 
then show ?lhs  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1240  | 
by (simp add: winding_number_homotopic_paths)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1241  | 
qed  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1242  | 
|
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1243  | 
|
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1244  | 
lemma winding_number_homotopic_loops_eq:  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1245  | 
assumes "path p" and \<zeta>p: "\<zeta> \<notin> path_image p"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1246  | 
and "path q" and \<zeta>q: "\<zeta> \<notin> path_image q"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1247  | 
and loops: "pathfinish p = pathstart p" "pathfinish q = pathstart q"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1248  | 
    shows "winding_number p \<zeta> = winding_number q \<zeta> \<longleftrightarrow> homotopic_loops (-{\<zeta>}) p q"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1249  | 
(is "?lhs = ?rhs")  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1250  | 
proof  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1251  | 
assume L: ?lhs  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1252  | 
have "pathstart p \<noteq> \<zeta>" "pathstart q \<noteq> \<zeta>"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1253  | 
using \<zeta>p \<zeta>q by blast+  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1254  | 
  moreover have "path_connected (-{\<zeta>})"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1255  | 
by (simp add: path_connected_punctured_universe)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1256  | 
  ultimately obtain r where "path r" and rim: "path_image r \<subseteq> -{\<zeta>}"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1257  | 
and pas: "pathstart r = pathstart p" and paf: "pathfinish r = pathstart q"  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1258  | 
by (auto simp: path_connected_def)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1259  | 
then have "pathstart r \<noteq> \<zeta>" by blast  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1260  | 
  have "homotopic_loops (- {\<zeta>}) p (r +++ q +++ reversepath r)"
 | 
| 
68077
 
ee8c13ae81e9
Some tidying up (mostly regarding summations from 0)
 
paulson <lp15@cam.ac.uk> 
parents: 
67135 
diff
changeset
 | 
1261  | 
proof (rule homotopic_paths_imp_homotopic_loops)  | 
| 
 
ee8c13ae81e9
Some tidying up (mostly regarding summations from 0)
 
paulson <lp15@cam.ac.uk> 
parents: 
67135 
diff
changeset
 | 
1262  | 
    show "homotopic_paths (- {\<zeta>}) p (r +++ q +++ reversepath r)"
 | 
| 
 
ee8c13ae81e9
Some tidying up (mostly regarding summations from 0)
 
paulson <lp15@cam.ac.uk> 
parents: 
67135 
diff
changeset
 | 
1263  | 
by (metis (mono_tags, hide_lams) \<open>path r\<close> L \<zeta>p \<zeta>q \<open>path p\<close> \<open>path q\<close> homotopic_loops_conjugate loops not_in_path_image_join paf pas path_image_reversepath path_imp_reversepath path_join_eq pathfinish_join pathfinish_reversepath pathstart_join pathstart_reversepath rim subset_Compl_singleton winding_number_homotopic_loops winding_number_homotopic_paths_eq)  | 
| 
 
ee8c13ae81e9
Some tidying up (mostly regarding summations from 0)
 
paulson <lp15@cam.ac.uk> 
parents: 
67135 
diff
changeset
 | 
1264  | 
qed (use loops pas in auto)  | 
| 
66826
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1265  | 
  moreover have "homotopic_loops (- {\<zeta>}) (r +++ q +++ reversepath r) q"
 | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1266  | 
using rim \<zeta>q by (auto simp: homotopic_loops_conjugate paf \<open>path q\<close> \<open>path r\<close> loops)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1267  | 
ultimately show ?rhs  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1268  | 
using homotopic_loops_trans by metis  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1269  | 
next  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1270  | 
assume ?rhs  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1271  | 
then show ?lhs  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1272  | 
by (simp add: winding_number_homotopic_loops)  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1273  | 
qed  | 
| 
 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 
paulson <lp15@cam.ac.uk> 
parents: 
66393 
diff
changeset
 | 
1274  | 
|
| 65039 | 1275  | 
end  | 
1276  |