| author | paulson <lp15@cam.ac.uk> | 
| Wed, 11 Apr 2018 16:34:44 +0100 | |
| changeset 67974 | 3f352a91b45a | 
| parent 67968 | a5ad4c015d1c | 
| child 68634 | db0980691ef4 | 
| permissions | -rw-r--r-- | 
| 66826 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1 | (* Title: HOL/Analysis/Riemann_Mapping.thy | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2 | Authors: LC Paulson, based on material from HOL Light | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3 | *) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5 | section \<open>Moebius functions, Equivalents of Simply Connected Sets, Riemann Mapping Theorem\<close> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 6 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 7 | theory Riemann_Mapping | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 8 | imports Great_Picard | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 9 | begin | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 10 | |
| 67968 | 11 | subsection\<open>Moebius functions are biholomorphisms of the unit disc\<close> | 
| 66826 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 12 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 13 | definition Moebius_function :: "[real,complex,complex] \<Rightarrow> complex" where | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 14 | "Moebius_function \<equiv> \<lambda>t w z. exp(\<i> * of_real t) * (z - w) / (1 - cnj w * z)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 15 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 16 | lemma Moebius_function_simple: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 17 | "Moebius_function 0 w z = (z - w) / (1 - cnj w * z)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 18 | by (simp add: Moebius_function_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 19 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 20 | lemma Moebius_function_eq_zero: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 21 | "Moebius_function t w w = 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 22 | by (simp add: Moebius_function_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 23 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 24 | lemma Moebius_function_of_zero: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 25 | "Moebius_function t w 0 = - exp(\<i> * of_real t) * w" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 26 | by (simp add: Moebius_function_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 27 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 28 | lemma Moebius_function_norm_lt_1: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 29 | assumes w1: "norm w < 1" and z1: "norm z < 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 30 | shows "norm (Moebius_function t w z) < 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 31 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 32 | have "1 - cnj w * z \<noteq> 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 33 | by (metis complex_cnj_cnj complex_mod_sqrt_Re_mult_cnj mult.commute mult_less_cancel_right1 norm_ge_zero norm_mult norm_one order.asym right_minus_eq w1 z1) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 34 | then have VV: "1 - w * cnj z \<noteq> 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 35 | by (metis complex_cnj_cnj complex_cnj_mult complex_cnj_one right_minus_eq) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 36 | then have "1 - norm (Moebius_function t w z) ^ 2 = | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 37 | ((1 - norm w ^ 2) / (norm (1 - cnj w * z) ^ 2)) * (1 - norm z ^ 2)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 38 | apply (cases w) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 39 | apply (cases z) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 40 | apply (simp add: Moebius_function_def divide_simps norm_divide norm_mult) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 41 | apply (simp add: complex_norm complex_diff complex_mult one_complex.code complex_cnj) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 42 | apply (auto simp: algebra_simps power2_eq_square) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 43 | done | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 44 | then have "1 - (cmod (Moebius_function t w z))\<^sup>2 = (1 - cmod (w * w)) / (cmod (1 - cnj w * z))\<^sup>2 * (1 - cmod (z * z))" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 45 | by (simp add: norm_mult power2_eq_square) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 46 | moreover have "0 < 1 - cmod (z * z)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 47 | by (metis (no_types) z1 diff_gt_0_iff_gt mult.left_neutral norm_mult_less) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 48 | ultimately have "0 < 1 - norm (Moebius_function t w z) ^ 2" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 49 | using \<open>1 - cnj w * z \<noteq> 0\<close> w1 norm_mult_less by fastforce | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 50 | then show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 51 | using linorder_not_less by fastforce | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 52 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 53 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 54 | lemma Moebius_function_holomorphic: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 55 | assumes "norm w < 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 56 | shows "Moebius_function t w holomorphic_on ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 57 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 58 | have *: "1 - z * w \<noteq> 0" if "norm z < 1" for z | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 59 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 60 | have "norm (1::complex) \<noteq> norm (z * w)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 61 | using assms that norm_mult_less by fastforce | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 62 | then show ?thesis by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 63 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 64 | show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 65 | apply (simp add: Moebius_function_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 66 | apply (intro holomorphic_intros) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 67 | using assms * | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 68 | by (metis complex_cnj_cnj complex_cnj_mult complex_cnj_one complex_mod_cnj mem_ball_0 mult.commute right_minus_eq) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 69 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 70 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 71 | lemma Moebius_function_compose: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 72 | assumes meq: "-w1 = w2" and "norm w1 < 1" "norm z < 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 73 | shows "Moebius_function 0 w1 (Moebius_function 0 w2 z) = z" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 74 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 75 | have "norm w2 < 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 76 | using assms by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 77 | then have "-w1 = z" if "cnj w2 * z = 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 78 | by (metis assms(3) complex_mod_cnj less_irrefl mult.right_neutral norm_mult_less norm_one that) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 79 | moreover have "z=0" if "1 - cnj w2 * z = cnj w1 * (z - w2)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 80 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 81 | have "w2 * cnj w2 = 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 82 | using that meq by (auto simp: algebra_simps) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 83 | then show "z = 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 84 | by (metis (no_types) \<open>cmod w2 < 1\<close> complex_mod_cnj less_irrefl mult.right_neutral norm_mult_less norm_one) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 85 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 86 | moreover have "z - w2 - w1 * (1 - cnj w2 * z) = z * (1 - cnj w2 * z - cnj w1 * (z - w2))" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 87 | using meq by (fastforce simp: algebra_simps) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 88 | ultimately | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 89 | show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 90 | by (simp add: Moebius_function_def divide_simps norm_divide norm_mult) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 91 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 92 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 93 | lemma ball_biholomorphism_exists: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 94 | assumes "a \<in> ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 95 | obtains f g where "f a = 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 96 | "f holomorphic_on ball 0 1" "f ` ball 0 1 \<subseteq> ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 97 | "g holomorphic_on ball 0 1" "g ` ball 0 1 \<subseteq> ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 98 | "\<And>z. z \<in> ball 0 1 \<Longrightarrow> f (g z) = z" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 99 | "\<And>z. z \<in> ball 0 1 \<Longrightarrow> g (f z) = z" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 100 | proof | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 101 | show "Moebius_function 0 a holomorphic_on ball 0 1" "Moebius_function 0 (-a) holomorphic_on ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 102 | using Moebius_function_holomorphic assms mem_ball_0 by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 103 | show "Moebius_function 0 a a = 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 104 | by (simp add: Moebius_function_eq_zero) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 105 | show "Moebius_function 0 a ` ball 0 1 \<subseteq> ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 106 | "Moebius_function 0 (- a) ` ball 0 1 \<subseteq> ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 107 | using Moebius_function_norm_lt_1 assms by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 108 | show "Moebius_function 0 a (Moebius_function 0 (- a) z) = z" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 109 | "Moebius_function 0 (- a) (Moebius_function 0 a z) = z" if "z \<in> ball 0 1" for z | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 110 | using Moebius_function_compose assms that by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 111 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 112 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 113 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 114 | subsection\<open>A big chain of equivalents of simple connectedness for an open set\<close> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 115 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 116 | lemma biholomorphic_to_disc_aux: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 117 | assumes "open S" "connected S" "0 \<in> S" and S01: "S \<subseteq> ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 118 | and prev: "\<And>f. \<lbrakk>f holomorphic_on S; \<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0; inj_on f S\<rbrakk> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 119 | \<Longrightarrow> \<exists>g. g holomorphic_on S \<and> (\<forall>z \<in> S. f z = (g z)\<^sup>2)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 120 | shows "\<exists>f g. f holomorphic_on S \<and> g holomorphic_on ball 0 1 \<and> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 121 | (\<forall>z \<in> S. f z \<in> ball 0 1 \<and> g(f z) = z) \<and> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 122 | (\<forall>z \<in> ball 0 1. g z \<in> S \<and> f(g z) = z)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 123 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 124 |   define F where "F \<equiv> {h. h holomorphic_on S \<and> h ` S \<subseteq> ball 0 1 \<and> h 0 = 0 \<and> inj_on h S}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 125 | have idF: "id \<in> F" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 126 | using S01 by (auto simp: F_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 127 |   then have "F \<noteq> {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 128 | by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 129 |   have imF_ne: "((\<lambda>h. norm(deriv h 0)) ` F) \<noteq> {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 130 | using idF by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 131 | have holF: "\<And>h. h \<in> F \<Longrightarrow> h holomorphic_on S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 132 | by (auto simp: F_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 133 | obtain f where "f \<in> F" and normf: "\<And>h. h \<in> F \<Longrightarrow> norm(deriv h 0) \<le> norm(deriv f 0)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 134 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 135 | obtain r where "r > 0" and r: "ball 0 r \<subseteq> S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 136 | using \<open>open S\<close> \<open>0 \<in> S\<close> openE by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 137 | have bdd: "bdd_above ((\<lambda>h. norm(deriv h 0)) ` F)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 138 | proof (intro bdd_aboveI exI ballI, clarify) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 139 | show "norm (deriv f 0) \<le> 1 / r" if "f \<in> F" for f | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 140 | proof - | 
| 67399 | 141 | have r01: "( * ) (complex_of_real r) ` ball 0 1 \<subseteq> S" | 
| 66826 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 142 | using that \<open>r > 0\<close> by (auto simp: norm_mult r [THEN subsetD]) | 
| 67399 | 143 | then have "f holomorphic_on ( * ) (complex_of_real r) ` ball 0 1" | 
| 66826 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 144 | using holomorphic_on_subset [OF holF] by (simp add: that) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 145 | then have holf: "f \<circ> (\<lambda>z. (r * z)) holomorphic_on (ball 0 1)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 146 | by (intro holomorphic_intros holomorphic_on_compose) | 
| 67399 | 147 | have f0: "(f \<circ> ( * ) (complex_of_real r)) 0 = 0" | 
| 66826 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 148 | using F_def that by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 149 | have "f ` S \<subseteq> ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 150 | using F_def that by blast | 
| 67399 | 151 | with r01 have fr1: "\<And>z. norm z < 1 \<Longrightarrow> norm ((f \<circ> ( * )(of_real r))z) < 1" | 
| 66826 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 152 | by force | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 153 | have *: "((\<lambda>w. f (r * w)) has_field_derivative deriv f (r * z) * r) (at z)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 154 | if "z \<in> ball 0 1" for z::complex | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 155 | proof (rule DERIV_chain' [where g=f]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 156 | show "(f has_field_derivative deriv f (of_real r * z)) (at (of_real r * z))" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 157 | apply (rule holomorphic_derivI [OF holF \<open>open S\<close>]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 158 | apply (rule \<open>f \<in> F\<close>) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 159 | by (meson imageI r01 subset_iff that) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 160 | qed simp | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 161 | have df0: "((\<lambda>w. f (r * w)) has_field_derivative deriv f 0 * r) (at 0)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 162 | using * [of 0] by simp | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 163 | have deq: "deriv (\<lambda>x. f (complex_of_real r * x)) 0 = deriv f 0 * complex_of_real r" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 164 | using DERIV_imp_deriv df0 by blast | 
| 67399 | 165 | have "norm (deriv (f \<circ> ( * ) (complex_of_real r)) 0) \<le> 1" | 
| 66826 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 166 | by (auto intro: Schwarz_Lemma [OF holf f0 fr1, of 0]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 167 | with \<open>r > 0\<close> show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 168 | by (simp add: deq norm_mult divide_simps o_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 169 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 170 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 171 | define l where "l \<equiv> SUP h:F. norm (deriv h 0)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 172 | have eql: "norm (deriv f 0) = l" if le: "l \<le> norm (deriv f 0)" and "f \<in> F" for f | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 173 | apply (rule order_antisym [OF _ le]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 174 | using \<open>f \<in> F\<close> bdd cSUP_upper by (fastforce simp: l_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 175 | obtain \<F> where \<F>in: "\<And>n. \<F> n \<in> F" and \<F>lim: "(\<lambda>n. norm (deriv (\<F> n) 0)) \<longlonglongrightarrow> l" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 176 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 177 | have "\<exists>f. f \<in> F \<and> \<bar>norm (deriv f 0) - l\<bar> < 1 / (Suc n)" for n | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 178 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 179 | obtain f where "f \<in> F" and f: "l < norm (deriv f 0) + 1/(Suc n)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 180 | using cSup_least [OF imF_ne, of "l - 1/(Suc n)"] by (fastforce simp add: l_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 181 | then have "\<bar>norm (deriv f 0) - l\<bar> < 1 / (Suc n)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 182 | by (fastforce simp add: abs_if not_less eql) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 183 | with \<open>f \<in> F\<close> show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 184 | by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 185 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 186 | then obtain \<F> where fF: "\<And>n. (\<F> n) \<in> F" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 187 | and fless: "\<And>n. \<bar>norm (deriv (\<F> n) 0) - l\<bar> < 1 / (Suc n)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 188 | by metis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 189 | have "(\<lambda>n. norm (deriv (\<F> n) 0)) \<longlonglongrightarrow> l" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 190 | proof (rule metric_LIMSEQ_I) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 191 | fix e::real | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 192 | assume "e > 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 193 | then obtain N::nat where N: "e > 1/(Suc N)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 194 | using nat_approx_posE by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 195 | show "\<exists>N. \<forall>n\<ge>N. dist (norm (deriv (\<F> n) 0)) l < e" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 196 | proof (intro exI allI impI) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 197 | fix n assume "N \<le> n" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 198 | have "dist (norm (deriv (\<F> n) 0)) l < 1 / (Suc n)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 199 | using fless by (simp add: dist_norm) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 200 | also have "... < e" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 201 | using N \<open>N \<le> n\<close> inverse_of_nat_le le_less_trans by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 202 | finally show "dist (norm (deriv (\<F> n) 0)) l < e" . | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 203 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 204 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 205 | with fF show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 206 | using that by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 207 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 208 | have "\<And>K. \<lbrakk>compact K; K \<subseteq> S\<rbrakk> \<Longrightarrow> \<exists>B. \<forall>h\<in>F. \<forall>z\<in>K. norm (h z) \<le> B" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 209 | by (rule_tac x=1 in exI) (force simp: F_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 210 | moreover have "range \<F> \<subseteq> F" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 211 | using \<open>\<And>n. \<F> n \<in> F\<close> by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 212 | ultimately obtain f and r :: "nat \<Rightarrow> nat" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 213 | where holf: "f holomorphic_on S" and r: "strict_mono r" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 214 | and limf: "\<And>x. x \<in> S \<Longrightarrow> (\<lambda>n. \<F> (r n) x) \<longlonglongrightarrow> f x" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 215 | and ulimf: "\<And>K. \<lbrakk>compact K; K \<subseteq> S\<rbrakk> \<Longrightarrow> uniform_limit K (\<F> \<circ> r) f sequentially" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 216 | using Montel [of S F \<F>, OF \<open>open S\<close> holF] by auto+ | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 217 | have der: "\<And>n x. x \<in> S \<Longrightarrow> ((\<F> \<circ> r) n has_field_derivative ((\<lambda>n. deriv (\<F> n)) \<circ> r) n x) (at x)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 218 | using \<open>\<And>n. \<F> n \<in> F\<close> \<open>open S\<close> holF holomorphic_derivI by fastforce | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 219 | have ulim: "\<And>x. x \<in> S \<Longrightarrow> \<exists>d>0. cball x d \<subseteq> S \<and> uniform_limit (cball x d) (\<F> \<circ> r) f sequentially" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 220 | by (meson ulimf \<open>open S\<close> compact_cball open_contains_cball) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 221 | obtain f' :: "complex\<Rightarrow>complex" where f': "(f has_field_derivative f' 0) (at 0)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 222 | and tof'0: "(\<lambda>n. ((\<lambda>n. deriv (\<F> n)) \<circ> r) n 0) \<longlonglongrightarrow> f' 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 223 | using has_complex_derivative_uniform_sequence [OF \<open>open S\<close> der ulim] \<open>0 \<in> S\<close> by metis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 224 | then have derf0: "deriv f 0 = f' 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 225 | by (simp add: DERIV_imp_deriv) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 226 | have "f field_differentiable (at 0)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 227 | using field_differentiable_def f' by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 228 | have "(\<lambda>x. (norm (deriv (\<F> (r x)) 0))) \<longlonglongrightarrow> norm (deriv f 0)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 229 | using isCont_tendsto_compose [OF continuous_norm [OF continuous_ident] tof'0] derf0 by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 230 | with LIMSEQ_subseq_LIMSEQ [OF \<F>lim r] have no_df0: "norm(deriv f 0) = l" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 231 | by (force simp: o_def intro: tendsto_unique) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 232 | have nonconstf: "\<not> f constant_on S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 233 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 234 | have False if "\<And>x. x \<in> S \<Longrightarrow> f x = c" for c | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 235 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 236 | have "deriv f 0 = 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 237 | by (metis that \<open>open S\<close> \<open>0 \<in> S\<close> DERIV_imp_deriv [OF DERIV_transform_within_open [OF DERIV_const]]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 238 | with no_df0 have "l = 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 239 | by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 240 | with eql [OF _ idF] show False by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 241 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 242 | then show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 243 | by (meson constant_on_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 244 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 245 | show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 246 | proof | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 247 | show "f \<in> F" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 248 | unfolding F_def | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 249 | proof (intro CollectI conjI holf) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 250 | have "norm(f z) \<le> 1" if "z \<in> S" for z | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 251 | proof (intro Lim_norm_ubound [OF _ limf] always_eventually allI that) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 252 | fix n | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 253 | have "\<F> (r n) \<in> F" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 254 | by (simp add: \<F>in) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 255 | then show "norm (\<F> (r n) z) \<le> 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 256 | using that by (auto simp: F_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 257 | qed simp | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 258 | then have fless1: "norm(f z) < 1" if "z \<in> S" for z | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 259 | using maximum_modulus_principle [OF holf \<open>open S\<close> \<open>connected S\<close> \<open>open S\<close>] nonconstf that | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 260 | by fastforce | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 261 | then show "f ` S \<subseteq> ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 262 | by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 263 | have "(\<lambda>n. \<F> (r n) 0) \<longlonglongrightarrow> 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 264 | using \<F>in by (auto simp: F_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 265 | then show "f 0 = 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 266 | using tendsto_unique [OF _ limf ] \<open>0 \<in> S\<close> trivial_limit_sequentially by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 267 | show "inj_on f S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 268 | proof (rule Hurwitz_injective [OF \<open>open S\<close> \<open>connected S\<close> _ holf]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 269 | show "\<And>n. (\<F> \<circ> r) n holomorphic_on S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 270 | by (simp add: \<F>in holF) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 271 | show "\<And>K. \<lbrakk>compact K; K \<subseteq> S\<rbrakk> \<Longrightarrow> uniform_limit K(\<F> \<circ> r) f sequentially" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 272 | by (metis ulimf) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 273 | show "\<not> f constant_on S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 274 | using nonconstf by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 275 | show "\<And>n. inj_on ((\<F> \<circ> r) n) S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 276 | using \<F>in by (auto simp: F_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 277 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 278 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 279 | show "\<And>h. h \<in> F \<Longrightarrow> norm (deriv h 0) \<le> norm (deriv f 0)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 280 | by (metis eql le_cases no_df0) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 281 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 282 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 283 | have holf: "f holomorphic_on S" and injf: "inj_on f S" and f01: "f ` S \<subseteq> ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 284 | using \<open>f \<in> F\<close> by (auto simp: F_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 285 | obtain g where holg: "g holomorphic_on (f ` S)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 286 | and derg: "\<And>z. z \<in> S \<Longrightarrow> deriv f z * deriv g (f z) = 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 287 | and gf: "\<And>z. z \<in> S \<Longrightarrow> g(f z) = z" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 288 | using holomorphic_has_inverse [OF holf \<open>open S\<close> injf] by metis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 289 | have "ball 0 1 \<subseteq> f ` S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 290 | proof | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 291 | fix a::complex | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 292 | assume a: "a \<in> ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 293 | have False if "\<And>x. x \<in> S \<Longrightarrow> f x \<noteq> a" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 294 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 295 | obtain h k where "h a = 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 296 | and holh: "h holomorphic_on ball 0 1" and h01: "h ` ball 0 1 \<subseteq> ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 297 | and holk: "k holomorphic_on ball 0 1" and k01: "k ` ball 0 1 \<subseteq> ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 298 | and hk: "\<And>z. z \<in> ball 0 1 \<Longrightarrow> h (k z) = z" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 299 | and kh: "\<And>z. z \<in> ball 0 1 \<Longrightarrow> k (h z) = z" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 300 | using ball_biholomorphism_exists [OF a] by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 301 | have nf1: "\<And>z. z \<in> S \<Longrightarrow> norm(f z) < 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 302 | using \<open>f \<in> F\<close> by (auto simp: F_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 303 | have 1: "h \<circ> f holomorphic_on S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 304 | using F_def \<open>f \<in> F\<close> holh holomorphic_on_compose holomorphic_on_subset by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 305 | have 2: "\<And>z. z \<in> S \<Longrightarrow> (h \<circ> f) z \<noteq> 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 306 | by (metis \<open>h a = 0\<close> a comp_eq_dest_lhs nf1 kh mem_ball_0 that) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 307 | have 3: "inj_on (h \<circ> f) S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 308 | by (metis (no_types, lifting) F_def \<open>f \<in> F\<close> comp_inj_on inj_on_inverseI injf kh mem_Collect_eq subset_inj_on) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 309 | obtain \<psi> where hol\<psi>: "\<psi> holomorphic_on ((h \<circ> f) ` S)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 310 | and \<psi>2: "\<And>z. z \<in> S \<Longrightarrow> \<psi>(h (f z)) ^ 2 = h(f z)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 311 | proof (rule exE [OF prev [OF 1 2 3]], safe) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 312 | fix \<theta> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 313 | assume hol\<theta>: "\<theta> holomorphic_on S" and \<theta>2: "(\<forall>z\<in>S. (h \<circ> f) z = (\<theta> z)\<^sup>2)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 314 | show thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 315 | proof | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 316 | show "(\<theta> \<circ> g \<circ> k) holomorphic_on (h \<circ> f) ` S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 317 | proof (intro holomorphic_on_compose) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 318 | show "k holomorphic_on (h \<circ> f) ` S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 319 | apply (rule holomorphic_on_subset [OF holk]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 320 | using f01 h01 by force | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 321 | show "g holomorphic_on k ` (h \<circ> f) ` S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 322 | apply (rule holomorphic_on_subset [OF holg]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 323 | by (auto simp: kh nf1) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 324 | show "\<theta> holomorphic_on g ` k ` (h \<circ> f) ` S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 325 | apply (rule holomorphic_on_subset [OF hol\<theta>]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 326 | by (auto simp: gf kh nf1) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 327 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 328 | show "((\<theta> \<circ> g \<circ> k) (h (f z)))\<^sup>2 = h (f z)" if "z \<in> S" for z | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 329 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 330 | have "f z \<in> ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 331 | by (simp add: nf1 that) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 332 | then have "(\<theta> (g (k (h (f z)))))\<^sup>2 = (\<theta> (g (f z)))\<^sup>2" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 333 | by (metis kh) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 334 | also have "... = h (f z)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 335 | using \<theta>2 gf that by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 336 | finally show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 337 | by (simp add: o_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 338 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 339 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 340 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 341 | have norm\<psi>1: "norm(\<psi> (h (f z))) < 1" if "z \<in> S" for z | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 342 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 343 | have "norm (\<psi> (h (f z)) ^ 2) < 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 344 | by (metis (no_types) that DIM_complex \<psi>2 h01 image_subset_iff mem_ball_0 nf1) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 345 | then show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 346 | by (metis le_less_trans mult_less_cancel_left2 norm_ge_zero norm_power not_le power2_eq_square) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 347 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 348 | then have \<psi>01: "\<psi> (h (f 0)) \<in> ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 349 | by (simp add: \<open>0 \<in> S\<close>) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 350 | obtain p q where p0: "p (\<psi> (h (f 0))) = 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 351 | and holp: "p holomorphic_on ball 0 1" and p01: "p ` ball 0 1 \<subseteq> ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 352 | and holq: "q holomorphic_on ball 0 1" and q01: "q ` ball 0 1 \<subseteq> ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 353 | and pq: "\<And>z. z \<in> ball 0 1 \<Longrightarrow> p (q z) = z" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 354 | and qp: "\<And>z. z \<in> ball 0 1 \<Longrightarrow> q (p z) = z" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 355 | using ball_biholomorphism_exists [OF \<psi>01] by metis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 356 | have "p \<circ> \<psi> \<circ> h \<circ> f \<in> F" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 357 | unfolding F_def | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 358 | proof (intro CollectI conjI holf) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 359 | show "p \<circ> \<psi> \<circ> h \<circ> f holomorphic_on S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 360 | proof (intro holomorphic_on_compose holf) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 361 | show "h holomorphic_on f ` S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 362 | apply (rule holomorphic_on_subset [OF holh]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 363 | using f01 by force | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 364 | show "\<psi> holomorphic_on h ` f ` S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 365 | apply (rule holomorphic_on_subset [OF hol\<psi>]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 366 | by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 367 | show "p holomorphic_on \<psi> ` h ` f ` S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 368 | apply (rule holomorphic_on_subset [OF holp]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 369 | by (auto simp: norm\<psi>1) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 370 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 371 | show "(p \<circ> \<psi> \<circ> h \<circ> f) ` S \<subseteq> ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 372 | apply clarsimp | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 373 | by (meson norm\<psi>1 p01 image_subset_iff mem_ball_0) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 374 | show "(p \<circ> \<psi> \<circ> h \<circ> f) 0 = 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 375 | by (simp add: \<open>p (\<psi> (h (f 0))) = 0\<close>) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 376 | show "inj_on (p \<circ> \<psi> \<circ> h \<circ> f) S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 377 | unfolding inj_on_def o_def | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 378 | by (metis \<psi>2 dist_0_norm gf kh mem_ball nf1 norm\<psi>1 qp) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 379 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 380 | then have le_norm_df0: "norm (deriv (p \<circ> \<psi> \<circ> h \<circ> f) 0) \<le> norm (deriv f 0)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 381 | by (rule normf) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 382 | have 1: "k \<circ> power2 \<circ> q holomorphic_on ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 383 | proof (intro holomorphic_on_compose holq) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 384 | show "power2 holomorphic_on q ` ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 385 | using holomorphic_on_subset holomorphic_on_power | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 386 | by (blast intro: holomorphic_on_ident) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 387 | show "k holomorphic_on power2 ` q ` ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 388 | apply (rule holomorphic_on_subset [OF holk]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 389 | using q01 by (auto simp: norm_power abs_square_less_1) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 390 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 391 | have 2: "(k \<circ> power2 \<circ> q) 0 = 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 392 | using p0 F_def \<open>f \<in> F\<close> \<psi>01 \<psi>2 \<open>0 \<in> S\<close> kh qp by force | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 393 | have 3: "norm ((k \<circ> power2 \<circ> q) z) < 1" if "norm z < 1" for z | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 394 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 395 | have "norm ((power2 \<circ> q) z) < 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 396 | using that q01 by (force simp: norm_power abs_square_less_1) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 397 | with k01 show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 398 | by fastforce | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 399 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 400 | have False if c: "\<forall>z. norm z < 1 \<longrightarrow> (k \<circ> power2 \<circ> q) z = c * z" and "norm c = 1" for c | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 401 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 402 | have "c \<noteq> 0" using that by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 403 | have "norm (p(1/2)) < 1" "norm (p(-1/2)) < 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 404 | using p01 by force+ | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 405 | then have "(k \<circ> power2 \<circ> q) (p(1/2)) = c * p(1/2)" "(k \<circ> power2 \<circ> q) (p(-1/2)) = c * p(-1/2)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 406 | using c by force+ | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 407 | then have "p (1/2) = p (- (1/2))" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 408 | by (auto simp: \<open>c \<noteq> 0\<close> qp o_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 409 | then have "q (p (1/2)) = q (p (- (1/2)))" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 410 | by simp | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 411 | then have "1/2 = - (1/2::complex)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 412 | by (auto simp: qp) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 413 | then show False | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 414 | by simp | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 415 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 416 | moreover | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 417 | have False if "norm (deriv (k \<circ> power2 \<circ> q) 0) \<noteq> 1" "norm (deriv (k \<circ> power2 \<circ> q) 0) \<le> 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 418 | and le: "\<And>\<xi>. norm \<xi> < 1 \<Longrightarrow> norm ((k \<circ> power2 \<circ> q) \<xi>) \<le> norm \<xi>" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 419 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 420 | have "norm (deriv (k \<circ> power2 \<circ> q) 0) < 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 421 | using that by simp | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 422 | moreover have eq: "deriv f 0 = deriv (k \<circ> (\<lambda>z. z ^ 2) \<circ> q) 0 * deriv (p \<circ> \<psi> \<circ> h \<circ> f) 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 423 | proof (intro DERIV_imp_deriv DERIV_transform_within_open [OF DERIV_chain]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 424 | show "(k \<circ> power2 \<circ> q has_field_derivative deriv (k \<circ> power2 \<circ> q) 0) (at ((p \<circ> \<psi> \<circ> h \<circ> f) 0))" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 425 | using "1" holomorphic_derivI p0 by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 426 | show "(p \<circ> \<psi> \<circ> h \<circ> f has_field_derivative deriv (p \<circ> \<psi> \<circ> h \<circ> f) 0) (at 0)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 427 | using \<open>p \<circ> \<psi> \<circ> h \<circ> f \<in> F\<close> \<open>open S\<close> \<open>0 \<in> S\<close> holF holomorphic_derivI by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 428 | show "\<And>x. x \<in> S \<Longrightarrow> (k \<circ> power2 \<circ> q \<circ> (p \<circ> \<psi> \<circ> h \<circ> f)) x = f x" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 429 | using \<psi>2 f01 kh norm\<psi>1 qp by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 430 | qed (use assms in simp_all) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 431 | ultimately have "cmod (deriv (p \<circ> \<psi> \<circ> h \<circ> f) 0) \<le> 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 432 | using le_norm_df0 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 433 | by (metis linorder_not_le mult.commute mult_less_cancel_left2 norm_mult) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 434 | moreover have "1 \<le> norm (deriv f 0)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 435 | using normf [of id] by (simp add: idF) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 436 | ultimately show False | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 437 | by (simp add: eq) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 438 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 439 | ultimately show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 440 | using Schwarz_Lemma [OF 1 2 3] norm_one by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 441 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 442 | then show "a \<in> f ` S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 443 | by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 444 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 445 | then have "f ` S = ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 446 | using F_def \<open>f \<in> F\<close> by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 447 | then show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 448 | apply (rule_tac x=f in exI) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 449 | apply (rule_tac x=g in exI) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 450 | using holf holg derg gf by safe force+ | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 451 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 452 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 453 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 454 | locale SC_Chain = | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 455 | fixes S :: "complex set" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 456 | assumes openS: "open S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 457 | begin | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 458 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 459 | lemma winding_number_zero: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 460 | assumes "simply_connected S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 461 | shows "connected S \<and> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 462 | (\<forall>\<gamma> z. path \<gamma> \<and> path_image \<gamma> \<subseteq> S \<and> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 463 | pathfinish \<gamma> = pathstart \<gamma> \<and> z \<notin> S \<longrightarrow> winding_number \<gamma> z = 0)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 464 | using assms | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 465 | by (auto simp: simply_connected_imp_connected simply_connected_imp_winding_number_zero) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 466 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 467 | lemma contour_integral_zero: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 468 | assumes "valid_path g" "path_image g \<subseteq> S" "pathfinish g = pathstart g" "f holomorphic_on S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 469 | "\<And>\<gamma> z. \<lbrakk>path \<gamma>; path_image \<gamma> \<subseteq> S; pathfinish \<gamma> = pathstart \<gamma>; z \<notin> S\<rbrakk> \<Longrightarrow> winding_number \<gamma> z = 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 470 | shows "(f has_contour_integral 0) g" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 471 | using assms by (meson Cauchy_theorem_global openS valid_path_imp_path) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 472 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 473 | lemma global_primitive: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 474 | assumes "connected S" and holf: "f holomorphic_on S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 475 | and prev: "\<And>\<gamma> f. \<lbrakk>valid_path \<gamma>; path_image \<gamma> \<subseteq> S; pathfinish \<gamma> = pathstart \<gamma>; f holomorphic_on S\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) \<gamma>" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 476 | shows "\<exists>h. \<forall>z \<in> S. (h has_field_derivative f z) (at z)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 477 | proof (cases "S = {}")
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 478 | case True then show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 479 | by simp | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 480 | next | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 481 | case False | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 482 | then obtain a where "a \<in> S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 483 | by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 484 | show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 485 | proof (intro exI ballI) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 486 | fix x assume "x \<in> S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 487 | then obtain d where "d > 0" and d: "cball x d \<subseteq> S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 488 | using openS open_contains_cball_eq by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 489 | let ?g = "\<lambda>z. (SOME g. polynomial_function g \<and> path_image g \<subseteq> S \<and> pathstart g = a \<and> pathfinish g = z)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 490 | show "((\<lambda>z. contour_integral (?g z) f) has_field_derivative f x) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 491 | (at x)" | 
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67399diff
changeset | 492 | proof (simp add: has_field_derivative_def has_derivative_at2 bounded_linear_mult_right, rule Lim_transform) | 
| 66826 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 493 | show "(\<lambda>y. inverse(norm(y - x)) *\<^sub>R (contour_integral(linepath x y) f - f x * (y - x))) \<midarrow>x\<rightarrow> 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 494 | proof (clarsimp simp add: Lim_at) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 495 | fix e::real assume "e > 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 496 | moreover have "continuous (at x) f" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 497 | using openS \<open>x \<in> S\<close> holf continuous_on_eq_continuous_at holomorphic_on_imp_continuous_on by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 498 | ultimately obtain d1 where "d1 > 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 499 | and d1: "\<And>x'. dist x' x < d1 \<Longrightarrow> dist (f x') (f x) < e/2" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 500 | unfolding continuous_at_eps_delta | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 501 | by (metis less_divide_eq_numeral1(1) mult_zero_left) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 502 | obtain d2 where "d2 > 0" and d2: "ball x d2 \<subseteq> S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 503 | using openS \<open>x \<in> S\<close> open_contains_ball_eq by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 504 | have "inverse (norm (y - x)) * norm (contour_integral (linepath x y) f - f x * (y - x)) < e" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 505 | if "0 < d1" "0 < d2" "y \<noteq> x" "dist y x < d1" "dist y x < d2" for y | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 506 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 507 | have "f contour_integrable_on linepath x y" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 508 | proof (rule contour_integrable_continuous_linepath [OF continuous_on_subset]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 509 | show "continuous_on S f" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 510 | by (simp add: holf holomorphic_on_imp_continuous_on) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 511 | have "closed_segment x y \<subseteq> ball x d2" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 512 | by (meson dist_commute_lessI dist_in_closed_segment le_less_trans mem_ball subsetI that(5)) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 513 | with d2 show "closed_segment x y \<subseteq> S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 514 | by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 515 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 516 | then obtain z where z: "(f has_contour_integral z) (linepath x y)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 517 | by (force simp: contour_integrable_on_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 518 | have con: "((\<lambda>w. f x) has_contour_integral f x * (y - x)) (linepath x y)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 519 | using has_contour_integral_const_linepath [of "f x" y x] by metis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 520 | have "norm (z - f x * (y - x)) \<le> (e/2) * norm (y - x)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 521 | proof (rule has_contour_integral_bound_linepath) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 522 | show "((\<lambda>w. f w - f x) has_contour_integral z - f x * (y - x)) (linepath x y)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 523 | by (rule has_contour_integral_diff [OF z con]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 524 | show "\<And>w. w \<in> closed_segment x y \<Longrightarrow> norm (f w - f x) \<le> e/2" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 525 | by (metis d1 dist_norm less_le_trans not_less not_less_iff_gr_or_eq segment_bound1 that(4)) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 526 | qed (use \<open>e > 0\<close> in auto) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 527 | with \<open>e > 0\<close> have "inverse (norm (y - x)) * norm (z - f x * (y - x)) \<le> e/2" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 528 | by (simp add: divide_simps) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 529 | also have "... < e" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 530 | using \<open>e > 0\<close> by simp | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 531 | finally show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 532 | by (simp add: contour_integral_unique [OF z]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 533 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 534 | with \<open>d1 > 0\<close> \<open>d2 > 0\<close> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 535 | show "\<exists>d>0. \<forall>z. z \<noteq> x \<and> dist z x < d \<longrightarrow> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 536 | inverse (norm (z - x)) * norm (contour_integral (linepath x z) f - f x * (z - x)) < e" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 537 | by (rule_tac x="min d1 d2" in exI) auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 538 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 539 | next | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 540 | have *: "(1 / norm (y - x)) *\<^sub>R (contour_integral (?g y) f - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 541 | (contour_integral (?g x) f + f x * (y - x))) = | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 542 | (contour_integral (linepath x y) f - f x * (y - x)) /\<^sub>R norm (y - x)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 543 | if "0 < d" "y \<noteq> x" and yx: "dist y x < d" for y | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 544 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 545 | have "y \<in> S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 546 | by (metis subsetD d dist_commute less_eq_real_def mem_cball yx) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 547 | have gxy: "polynomial_function (?g x) \<and> path_image (?g x) \<subseteq> S \<and> pathstart (?g x) = a \<and> pathfinish (?g x) = x" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 548 | "polynomial_function (?g y) \<and> path_image (?g y) \<subseteq> S \<and> pathstart (?g y) = a \<and> pathfinish (?g y) = y" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 549 | using someI_ex [OF connected_open_polynomial_connected [OF openS \<open>connected S\<close> \<open>a \<in> S\<close>]] \<open>x \<in> S\<close> \<open>y \<in> S\<close> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 550 | by meson+ | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 551 | then have vp: "valid_path (?g x)" "valid_path (?g y)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 552 | by (simp_all add: valid_path_polynomial_function) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 553 | have f0: "(f has_contour_integral 0) ((?g x) +++ linepath x y +++ reversepath (?g y))" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 554 | proof (rule prev) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 555 | show "valid_path ((?g x) +++ linepath x y +++ reversepath (?g y))" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 556 | using gxy vp by (auto simp: valid_path_join) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 557 | have "closed_segment x y \<subseteq> cball x d" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 558 | using yx by (auto simp: dist_commute dest!: dist_in_closed_segment) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 559 | then have "closed_segment x y \<subseteq> S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 560 | using d by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 561 | then show "path_image ((?g x) +++ linepath x y +++ reversepath (?g y)) \<subseteq> S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 562 | using gxy by (auto simp: path_image_join) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 563 | qed (use gxy holf in auto) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 564 | then have fintxy: "f contour_integrable_on linepath x y" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 565 | by (metis (no_types, lifting) contour_integrable_joinD1 contour_integrable_joinD2 gxy(2) has_contour_integral_integrable pathfinish_linepath pathstart_reversepath valid_path_imp_reverse valid_path_join valid_path_linepath vp(2)) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 566 | have fintgx: "f contour_integrable_on (?g x)" "f contour_integrable_on (?g y)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 567 | using openS contour_integrable_holomorphic_simple gxy holf vp by blast+ | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 568 | show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 569 | apply (clarsimp simp add: divide_simps) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 570 | using contour_integral_unique [OF f0] | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 571 | apply (simp add: fintxy gxy contour_integrable_reversepath contour_integral_reversepath fintgx vp) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 572 | apply (simp add: algebra_simps) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 573 | done | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 574 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 575 | show "(\<lambda>z. (1 / norm (z - x)) *\<^sub>R | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 576 | (contour_integral (?g z) f - (contour_integral (?g x) f + f x * (z - x))) - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 577 | (contour_integral (linepath x z) f - f x * (z - x)) /\<^sub>R norm (z - x)) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 578 | \<midarrow>x\<rightarrow> 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 579 | apply (rule Lim_eventually) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 580 | apply (simp add: eventually_at) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 581 | apply (rule_tac x=d in exI) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 582 | using \<open>d > 0\<close> * by simp | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 583 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 584 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 585 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 586 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 587 | lemma holomorphic_log: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 588 | assumes "connected S" and holf: "f holomorphic_on S" and nz: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 589 | and prev: "\<And>f. f holomorphic_on S \<Longrightarrow> \<exists>h. \<forall>z \<in> S. (h has_field_derivative f z) (at z)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 590 | shows "\<exists>g. g holomorphic_on S \<and> (\<forall>z \<in> S. f z = exp(g z))" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 591 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 592 | have "(\<lambda>z. deriv f z / f z) holomorphic_on S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 593 | by (simp add: openS holf holomorphic_deriv holomorphic_on_divide nz) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 594 | then obtain g where g: "\<And>z. z \<in> S \<Longrightarrow> (g has_field_derivative deriv f z / f z) (at z)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 595 | using prev [of "\<lambda>z. deriv f z / f z"] by metis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 596 | have hfd: "\<And>x. x \<in> S \<Longrightarrow> ((\<lambda>z. exp (g z) / f z) has_field_derivative 0) (at x)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 597 | apply (rule derivative_eq_intros g| simp)+ | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 598 | apply (subst DERIV_deriv_iff_field_differentiable) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 599 | using openS holf holomorphic_on_imp_differentiable_at nz apply auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 600 | done | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 601 | obtain c where c: "\<And>x. x \<in> S \<Longrightarrow> exp (g x) / f x = c" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 602 | proof (rule DERIV_zero_connected_constant[OF \<open>connected S\<close> openS finite.emptyI]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 603 | show "continuous_on S (\<lambda>z. exp (g z) / f z)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 604 | by (metis (full_types) openS g continuous_on_divide continuous_on_exp holf holomorphic_on_imp_continuous_on holomorphic_on_open nz) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 605 |     then show "\<forall>x\<in>S - {}. ((\<lambda>z. exp (g z) / f z) has_field_derivative 0) (at x)"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 606 | using hfd by (blast intro: DERIV_zero_connected_constant [OF \<open>connected S\<close> openS finite.emptyI, of "\<lambda>z. exp(g z) / f z"]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 607 | qed auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 608 | show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 609 | proof (intro exI ballI conjI) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 610 | show "(\<lambda>z. Ln(inverse c) + g z) holomorphic_on S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 611 | apply (intro holomorphic_intros) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 612 | using openS g holomorphic_on_open by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 613 | fix z :: complex | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 614 | assume "z \<in> S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 615 | then have "exp (g z) / c = f z" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 616 | by (metis c divide_divide_eq_right exp_not_eq_zero nonzero_mult_div_cancel_left) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 617 | moreover have "1 / c \<noteq> 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 618 | using \<open>z \<in> S\<close> c nz by fastforce | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 619 | ultimately show "f z = exp (Ln (inverse c) + g z)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 620 | by (simp add: exp_add inverse_eq_divide) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 621 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 622 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 623 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 624 | lemma holomorphic_sqrt: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 625 | assumes holf: "f holomorphic_on S" and nz: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 626 | and prev: "\<And>f. \<lbrakk>f holomorphic_on S; \<forall>z \<in> S. f z \<noteq> 0\<rbrakk> \<Longrightarrow> \<exists>g. g holomorphic_on S \<and> (\<forall>z \<in> S. f z = exp(g z))" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 627 | shows "\<exists>g. g holomorphic_on S \<and> (\<forall>z \<in> S. f z = (g z)\<^sup>2)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 628 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 629 | obtain g where holg: "g holomorphic_on S" and g: "\<And>z. z \<in> S \<Longrightarrow> f z = exp (g z)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 630 | using prev [of f] holf nz by metis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 631 | show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 632 | proof (intro exI ballI conjI) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 633 | show "(\<lambda>z. exp(g z/2)) holomorphic_on S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 634 | by (intro holomorphic_intros) (auto simp: holg) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 635 | show "\<And>z. z \<in> S \<Longrightarrow> f z = (exp (g z/2))\<^sup>2" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 636 | by (metis (no_types) g exp_double nonzero_mult_div_cancel_left times_divide_eq_right zero_neq_numeral) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 637 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 638 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 639 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 640 | lemma biholomorphic_to_disc: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 641 |   assumes "connected S" and S: "S \<noteq> {}" "S \<noteq> UNIV"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 642 | and prev: "\<And>f. \<lbrakk>f holomorphic_on S; \<forall>z \<in> S. f z \<noteq> 0\<rbrakk> \<Longrightarrow> \<exists>g. g holomorphic_on S \<and> (\<forall>z \<in> S. f z = (g z)\<^sup>2)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 643 | shows "\<exists>f g. f holomorphic_on S \<and> g holomorphic_on ball 0 1 \<and> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 644 | (\<forall>z \<in> S. f z \<in> ball 0 1 \<and> g(f z) = z) \<and> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 645 | (\<forall>z \<in> ball 0 1. g z \<in> S \<and> f(g z) = z)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 646 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 647 | obtain a b where "a \<in> S" "b \<notin> S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 648 | using S by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 649 | then obtain \<delta> where "\<delta> > 0" and \<delta>: "ball a \<delta> \<subseteq> S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 650 | using openS openE by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 651 | obtain g where holg: "g holomorphic_on S" and eqg: "\<And>z. z \<in> S \<Longrightarrow> z - b = (g z)\<^sup>2" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 652 | proof (rule exE [OF prev [of "\<lambda>z. z - b"]]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 653 | show "(\<lambda>z. z - b) holomorphic_on S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 654 | by (intro holomorphic_intros) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 655 | qed (use \<open>b \<notin> S\<close> in auto) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 656 | have "\<not> g constant_on S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 657 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 658 | have "(a + \<delta>/2) \<in> ball a \<delta>" "a + (\<delta>/2) \<noteq> a" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 659 | using \<open>\<delta> > 0\<close> by (simp_all add: dist_norm) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 660 | then show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 661 | unfolding constant_on_def | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 662 | using eqg [of a] eqg [of "a + \<delta>/2"] \<open>a \<in> S\<close> \<delta> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 663 | by (metis diff_add_cancel subset_eq) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 664 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 665 | then have "open (g ` ball a \<delta>)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 666 | using open_mapping_thm [of g S "ball a \<delta>", OF holg openS \<open>connected S\<close>] \<delta> by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 667 | then obtain r where "r > 0" and r: "ball (g a) r \<subseteq> (g ` ball a \<delta>)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 668 | by (metis \<open>0 < \<delta>\<close> centre_in_ball imageI openE) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 669 | have g_not_r: "g z \<notin> ball (-(g a)) r" if "z \<in> S" for z | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 670 | proof | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 671 | assume "g z \<in> ball (-(g a)) r" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 672 | then have "- g z \<in> ball (g a) r" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 673 | by (metis add.inverse_inverse dist_minus mem_ball) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 674 | with r have "- g z \<in> (g ` ball a \<delta>)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 675 | by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 676 | then obtain w where w: "- g z = g w" "dist a w < \<delta>" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 677 | by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 678 | then have "w \<in> ball a \<delta>" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 679 | by simp | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 680 | then have "w \<in> S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 681 | using \<delta> by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 682 | then have "w = z" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 683 | by (metis diff_add_cancel eqg power_minus_Bit0 that w(1)) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 684 | then have "g z = 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 685 | using \<open>- g z = g w\<close> by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 686 | with eqg [OF that] have "z = b" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 687 | by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 688 | with that \<open>b \<notin> S\<close> show False | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 689 | by simp | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 690 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 691 | then have nz: "\<And>z. z \<in> S \<Longrightarrow> g z + g a \<noteq> 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 692 | by (metis \<open>0 < r\<close> add.commute add_diff_cancel_left' centre_in_ball diff_0) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 693 | let ?f = "\<lambda>z. (r/3) / (g z + g a) - (r/3) / (g a + g a)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 694 | obtain h where holh: "h holomorphic_on S" and "h a = 0" and h01: "h ` S \<subseteq> ball 0 1" and "inj_on h S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 695 | proof | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 696 | show "?f holomorphic_on S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 697 | by (intro holomorphic_intros holg nz) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 698 | have 3: "\<lbrakk>norm x \<le> 1/3; norm y \<le> 1/3\<rbrakk> \<Longrightarrow> norm(x - y) < 1" for x y::complex | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 699 | using norm_triangle_ineq4 [of x y] by simp | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 700 | have "norm ((r/3) / (g z + g a) - (r/3) / (g a + g a)) < 1" if "z \<in> S" for z | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 701 | apply (rule 3) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 702 | unfolding norm_divide | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 703 | using \<open>r > 0\<close> g_not_r [OF \<open>z \<in> S\<close>] g_not_r [OF \<open>a \<in> S\<close>] | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 704 | by (simp_all add: divide_simps dist_commute dist_norm) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 705 | then show "?f ` S \<subseteq> ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 706 | by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 707 | show "inj_on ?f S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 708 | using \<open>r > 0\<close> eqg apply (clarsimp simp: inj_on_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 709 | by (metis diff_add_cancel) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 710 | qed auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 711 | obtain k where holk: "k holomorphic_on (h ` S)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 712 | and derk: "\<And>z. z \<in> S \<Longrightarrow> deriv h z * deriv k (h z) = 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 713 | and kh: "\<And>z. z \<in> S \<Longrightarrow> k(h z) = z" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 714 | using holomorphic_has_inverse [OF holh openS \<open>inj_on h S\<close>] by metis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 715 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 716 | have 1: "open (h ` S)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 717 | by (simp add: \<open>inj_on h S\<close> holh openS open_mapping_thm3) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 718 | have 2: "connected (h ` S)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 719 | by (simp add: connected_continuous_image \<open>connected S\<close> holh holomorphic_on_imp_continuous_on) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 720 | have 3: "0 \<in> h ` S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 721 | using \<open>a \<in> S\<close> \<open>h a = 0\<close> by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 722 | have 4: "\<exists>g. g holomorphic_on h ` S \<and> (\<forall>z\<in>h ` S. f z = (g z)\<^sup>2)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 723 | if holf: "f holomorphic_on h ` S" and nz: "\<And>z. z \<in> h ` S \<Longrightarrow> f z \<noteq> 0" "inj_on f (h ` S)" for f | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 724 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 725 | obtain g where holg: "g holomorphic_on S" and eqg: "\<And>z. z \<in> S \<Longrightarrow> (f \<circ> h) z = (g z)\<^sup>2" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 726 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 727 | have "f \<circ> h holomorphic_on S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 728 | by (simp add: holh holomorphic_on_compose holf) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 729 | moreover have "\<forall>z\<in>S. (f \<circ> h) z \<noteq> 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 730 | by (simp add: nz) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 731 | ultimately show thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 732 | using prev that by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 733 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 734 | show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 735 | proof (intro exI conjI) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 736 | show "g \<circ> k holomorphic_on h ` S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 737 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 738 | have "k ` h ` S \<subseteq> S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 739 | by (simp add: \<open>\<And>z. z \<in> S \<Longrightarrow> k (h z) = z\<close> image_subset_iff) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 740 | then show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 741 | by (meson holg holk holomorphic_on_compose holomorphic_on_subset) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 742 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 743 | show "\<forall>z\<in>h ` S. f z = ((g \<circ> k) z)\<^sup>2" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 744 | using eqg kh by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 745 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 746 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 747 | obtain f g where f: "f holomorphic_on h ` S" and g: "g holomorphic_on ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 748 | and gf: "\<forall>z\<in>h ` S. f z \<in> ball 0 1 \<and> g (f z) = z" and fg:"\<forall>z\<in>ball 0 1. g z \<in> h ` S \<and> f (g z) = z" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 749 | using biholomorphic_to_disc_aux [OF 1 2 3 h01 4] by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 750 | show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 751 | proof (intro exI conjI) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 752 | show "f \<circ> h holomorphic_on S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 753 | by (simp add: f holh holomorphic_on_compose) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 754 | show "k \<circ> g holomorphic_on ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 755 | by (metis holomorphic_on_subset image_subset_iff fg holk g holomorphic_on_compose) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 756 | qed (use fg gf kh in auto) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 757 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 758 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 759 | lemma homeomorphic_to_disc: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 760 |   assumes S: "S \<noteq> {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 761 | and prev: "S = UNIV \<or> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 762 | (\<exists>f g. f holomorphic_on S \<and> g holomorphic_on ball 0 1 \<and> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 763 | (\<forall>z \<in> S. f z \<in> ball 0 1 \<and> g(f z) = z) \<and> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 764 | (\<forall>z \<in> ball 0 1. g z \<in> S \<and> f(g z) = z))" (is "_ \<or> ?P") | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 765 | shows "S homeomorphic ball (0::complex) 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 766 | using prev | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 767 | proof | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 768 | assume "S = UNIV" then show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 769 | using homeomorphic_ball01_UNIV homeomorphic_sym by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 770 | next | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 771 | assume ?P | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 772 | then show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 773 | unfolding homeomorphic_minimal | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 774 | using holomorphic_on_imp_continuous_on by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 775 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 776 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 777 | lemma homeomorphic_to_disc_imp_simply_connected: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 778 |   assumes "S = {} \<or> S homeomorphic ball (0::complex) 1"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 779 | shows "simply_connected S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 780 | using assms homeomorphic_simply_connected_eq convex_imp_simply_connected by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 781 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 782 | end | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 783 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 784 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 785 | proposition | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 786 | assumes "open S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 787 | shows simply_connected_eq_winding_number_zero: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 788 | "simply_connected S \<longleftrightarrow> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 789 | connected S \<and> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 790 | (\<forall>g z. path g \<and> path_image g \<subseteq> S \<and> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 791 | pathfinish g = pathstart g \<and> (z \<notin> S) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 792 | \<longrightarrow> winding_number g z = 0)" (is "?wn0") | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 793 | and simply_connected_eq_contour_integral_zero: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 794 | "simply_connected S \<longleftrightarrow> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 795 | connected S \<and> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 796 | (\<forall>g f. valid_path g \<and> path_image g \<subseteq> S \<and> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 797 | pathfinish g = pathstart g \<and> f holomorphic_on S | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 798 | \<longrightarrow> (f has_contour_integral 0) g)" (is "?ci0") | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 799 | and simply_connected_eq_global_primitive: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 800 | "simply_connected S \<longleftrightarrow> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 801 | connected S \<and> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 802 | (\<forall>f. f holomorphic_on S \<longrightarrow> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 803 | (\<exists>h. \<forall>z. z \<in> S \<longrightarrow> (h has_field_derivative f z) (at z)))" (is "?gp") | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 804 | and simply_connected_eq_holomorphic_log: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 805 | "simply_connected S \<longleftrightarrow> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 806 | connected S \<and> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 807 | (\<forall>f. f holomorphic_on S \<and> (\<forall>z \<in> S. f z \<noteq> 0) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 808 | \<longrightarrow> (\<exists>g. g holomorphic_on S \<and> (\<forall>z \<in> S. f z = exp(g z))))" (is "?log") | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 809 | and simply_connected_eq_holomorphic_sqrt: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 810 | "simply_connected S \<longleftrightarrow> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 811 | connected S \<and> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 812 | (\<forall>f. f holomorphic_on S \<and> (\<forall>z \<in> S. f z \<noteq> 0) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 813 | \<longrightarrow> (\<exists>g. g holomorphic_on S \<and> (\<forall>z \<in> S. f z = (g z)\<^sup>2)))" (is "?sqrt") | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 814 | and simply_connected_eq_biholomorphic_to_disc: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 815 | "simply_connected S \<longleftrightarrow> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 816 |            S = {} \<or> S = UNIV \<or>
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 817 | (\<exists>f g. f holomorphic_on S \<and> g holomorphic_on ball 0 1 \<and> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 818 | (\<forall>z \<in> S. f z \<in> ball 0 1 \<and> g(f z) = z) \<and> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 819 | (\<forall>z \<in> ball 0 1. g z \<in> S \<and> f(g z) = z))" (is "?bih") | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 820 | and simply_connected_eq_homeomorphic_to_disc: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 821 |           "simply_connected S \<longleftrightarrow> S = {} \<or> S homeomorphic ball (0::complex) 1" (is "?disc")
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 822 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 823 | interpret SC_Chain | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 824 | using assms by (simp add: SC_Chain_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 825 | have "?wn0 \<and> ?ci0 \<and> ?gp \<and> ?log \<and> ?sqrt \<and> ?bih \<and> ?disc" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 826 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 827 | have *: "\<lbrakk>\<alpha> \<Longrightarrow> \<beta>; \<beta> \<Longrightarrow> \<gamma>; \<gamma> \<Longrightarrow> \<delta>; \<delta> \<Longrightarrow> \<zeta>; \<zeta> \<Longrightarrow> \<eta>; \<eta> \<Longrightarrow> \<theta>; \<theta> \<Longrightarrow> \<xi>; \<xi> \<Longrightarrow> \<alpha>\<rbrakk> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 828 | \<Longrightarrow> (\<alpha> \<longleftrightarrow> \<beta>) \<and> (\<alpha> \<longleftrightarrow> \<gamma>) \<and> (\<alpha> \<longleftrightarrow> \<delta>) \<and> (\<alpha> \<longleftrightarrow> \<zeta>) \<and> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 829 | (\<alpha> \<longleftrightarrow> \<eta>) \<and> (\<alpha> \<longleftrightarrow> \<theta>) \<and> (\<alpha> \<longleftrightarrow> \<xi>)" for \<alpha> \<beta> \<gamma> \<delta> \<zeta> \<eta> \<theta> \<xi> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 830 | by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 831 | show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 832 | apply (rule *) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 833 | using winding_number_zero apply metis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 834 | using contour_integral_zero apply metis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 835 | using global_primitive apply metis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 836 | using holomorphic_log apply metis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 837 | using holomorphic_sqrt apply simp | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 838 | using biholomorphic_to_disc apply blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 839 | using homeomorphic_to_disc apply blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 840 | using homeomorphic_to_disc_imp_simply_connected apply blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 841 | done | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 842 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 843 | then show ?wn0 ?ci0 ?gp ?log ?sqrt ?bih ?disc | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 844 | by safe | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 845 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 846 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 847 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 848 | corollary contractible_eq_simply_connected_2d: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 849 | fixes S :: "complex set" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 850 | shows "open S \<Longrightarrow> (contractible S \<longleftrightarrow> simply_connected S)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 851 | apply safe | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 852 | apply (simp add: contractible_imp_simply_connected) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 853 | using convex_imp_contractible homeomorphic_contractible_eq simply_connected_eq_homeomorphic_to_disc by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 854 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 855 | |
| 67968 | 856 | subsection\<open>A further chain of equivalences about components of the complement of a simply connected set\<close> | 
| 66826 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 857 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 858 | text\<open>(following 1.35 in Burckel'S book)\<close> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 859 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 860 | context SC_Chain | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 861 | begin | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 862 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 863 | lemma frontier_properties: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 864 | assumes "simply_connected S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 865 | shows "if bounded S then connected(frontier S) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 866 | else \<forall>C \<in> components(frontier S). ~bounded C" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 867 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 868 |   have "S = {} \<or> S homeomorphic ball (0::complex) 1"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 869 | using simply_connected_eq_homeomorphic_to_disc assms openS by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 870 | then show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 871 | proof | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 872 |     assume "S = {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 873 | then have "bounded S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 874 | by simp | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 875 |     with \<open>S = {}\<close> show ?thesis
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 876 | by simp | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 877 | next | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 878 | assume S01: "S homeomorphic ball (0::complex) 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 879 | then obtain g f | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 880 | where gim: "g ` S = ball 0 1" and fg: "\<And>x. x \<in> S \<Longrightarrow> f(g x) = x" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 881 | and fim: "f ` ball 0 1 = S" and gf: "\<And>y. cmod y < 1 \<Longrightarrow> g(f y) = y" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 882 | and contg: "continuous_on S g" and contf: "continuous_on (ball 0 1) f" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 883 | by (fastforce simp: homeomorphism_def homeomorphic_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 884 | define D where "D \<equiv> \<lambda>n. ball (0::complex) (1 - 1/(of_nat n + 2))" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 885 |     define A where "A \<equiv> \<lambda>n. {z::complex. 1 - 1/(of_nat n + 2) < norm z \<and> norm z < 1}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 886 | define X where "X \<equiv> \<lambda>n::nat. closure(f ` A n)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 887 | have D01: "D n \<subseteq> ball 0 1" for n | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 888 | by (simp add: D_def ball_subset_ball_iff) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 889 | have A01: "A n \<subseteq> ball 0 1" for n | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 890 | by (auto simp: A_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 891 | have cloX: "closed(X n)" for n | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 892 | by (simp add: X_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 893 | have Xsubclo: "X n \<subseteq> closure S" for n | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 894 | unfolding X_def by (metis A01 closure_mono fim image_mono) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 895 | have connX: "connected(X n)" for n | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 896 | unfolding X_def | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 897 | apply (rule connected_imp_connected_closure) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 898 | apply (rule connected_continuous_image) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 899 | apply (simp add: continuous_on_subset [OF contf A01]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 900 | using connected_annulus [of _ "0::complex"] by (simp add: A_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 901 | have nestX: "X n \<subseteq> X m" if "m \<le> n" for m n | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 902 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 903 | have "1 - 1 / (real m + 2) \<le> 1 - 1 / (real n + 2)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 904 | using that by (auto simp: field_simps) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 905 | then show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 906 | by (auto simp: X_def A_def intro!: closure_mono) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 907 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 908 | have "closure S - S \<subseteq> (\<Inter>n. X n)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 909 | proof | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 910 | fix x | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 911 | assume "x \<in> closure S - S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 912 | then have "x \<in> closure S" "x \<notin> S" by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 913 | show "x \<in> (\<Inter>n. X n)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 914 | proof | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 915 | fix n | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 916 | have "ball 0 1 = closure (D n) \<union> A n" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 917 | by (auto simp: D_def A_def le_less_trans) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 918 | with fim have Seq: "S = f ` (closure (D n)) \<union> f ` (A n)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 919 | by (simp add: image_Un) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 920 | have "continuous_on (closure (D n)) f" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 921 | by (simp add: D_def cball_subset_ball_iff continuous_on_subset [OF contf]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 922 | moreover have "compact (closure (D n))" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 923 | by (simp add: D_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 924 | ultimately have clo_fim: "closed (f ` closure (D n))" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 925 | using compact_continuous_image compact_imp_closed by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 926 | have *: "(f ` cball 0 (1 - 1 / (real n + 2))) \<subseteq> S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 927 | by (force simp: D_def Seq) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 928 | show "x \<in> X n" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 929 | using \<open>x \<in> closure S\<close> unfolding X_def Seq | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 930 | using \<open>x \<notin> S\<close> * D_def clo_fim by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 931 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 932 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 933 | moreover have "(\<Inter>n. X n) \<subseteq> closure S - S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 934 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 935 | have "(\<Inter>n. X n) \<subseteq> closure S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 936 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 937 | have "(\<Inter>n. X n) \<subseteq> X 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 938 | by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 939 | also have "... \<subseteq> closure S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 940 | apply (simp add: X_def fim [symmetric]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 941 | apply (rule closure_mono) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 942 | by (auto simp: A_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 943 | finally show "(\<Inter>n. X n) \<subseteq> closure S" . | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 944 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 945 |       moreover have "(\<Inter>n. X n) \<inter> S \<subseteq> {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 946 | proof (clarify, clarsimp simp: X_def fim [symmetric]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 947 | fix x assume x [rule_format]: "\<forall>n. f x \<in> closure (f ` A n)" and "cmod x < 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 948 | then obtain n where n: "1 / (1 - norm x) < of_nat n" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 949 | using reals_Archimedean2 by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 950 | with \<open>cmod x < 1\<close> gr0I have XX: "1 / of_nat n < 1 - norm x" and "n > 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 951 | by (fastforce simp: divide_simps algebra_simps)+ | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 952 | have "f x \<in> f ` (D n)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 953 | using n \<open>cmod x < 1\<close> by (auto simp: divide_simps algebra_simps D_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 954 |         moreover have " f ` D n \<inter> closure (f ` A n) = {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 955 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 956 | have op_fDn: "open(f ` (D n))" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 957 | proof (rule invariance_of_domain) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 958 | show "continuous_on (D n) f" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 959 | by (rule continuous_on_subset [OF contf D01]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 960 | show "open (D n)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 961 | by (simp add: D_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 962 | show "inj_on f (D n)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 963 | unfolding inj_on_def using D01 by (metis gf mem_ball_0 subsetCE) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 964 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 965 | have injf: "inj_on f (ball 0 1)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 966 | by (metis mem_ball_0 inj_on_def gf) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 967 | have "D n \<union> A n \<subseteq> ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 968 | using D01 A01 by simp | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 969 |           moreover have "D n \<inter> A n = {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 970 | by (auto simp: D_def A_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 971 |           ultimately have "f ` D n \<inter> f ` A n = {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 972 | by (metis A01 D01 image_is_empty inj_on_image_Int injf) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 973 | then show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 974 | by (simp add: open_Int_closure_eq_empty [OF op_fDn]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 975 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 976 | ultimately show False | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 977 | using x [of n] by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 978 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 979 | ultimately | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 980 | show "(\<Inter>n. X n) \<subseteq> closure S - S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 981 | using closure_subset disjoint_iff_not_equal by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 982 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 983 | ultimately have "closure S - S = (\<Inter>n. X n)" by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 984 | then have frontierS: "frontier S = (\<Inter>n. X n)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 985 | by (simp add: frontier_def openS interior_open) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 986 | show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 987 | proof (cases "bounded S") | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 988 | case True | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 989 | have bouX: "bounded (X n)" for n | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 990 | apply (simp add: X_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 991 | apply (rule bounded_closure) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 992 | by (metis A01 fim image_mono bounded_subset [OF True]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 993 | have compaX: "compact (X n)" for n | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 994 | apply (simp add: compact_eq_bounded_closed bouX) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 995 | apply (auto simp: X_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 996 | done | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 997 | have "connected (\<Inter>n. X n)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 998 | by (metis nestX compaX connX connected_nest) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 999 | then show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1000 | by (simp add: True \<open>frontier S = (\<Inter>n. X n)\<close>) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1001 | next | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1002 | case False | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1003 | have unboundedX: "\<not> bounded(X n)" for n | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1004 | proof | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1005 | assume bXn: "bounded(X n)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1006 | have "continuous_on (cball 0 (1 - 1 / (2 + real n))) f" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1007 | by (simp add: cball_subset_ball_iff continuous_on_subset [OF contf]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1008 | then have "bounded (f ` cball 0 (1 - 1 / (2 + real n)))" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1009 | by (simp add: compact_imp_bounded [OF compact_continuous_image]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1010 | moreover have "bounded (f ` A n)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1011 | by (auto simp: X_def closure_subset image_subset_iff bounded_subset [OF bXn]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1012 | ultimately have "bounded (f ` (cball 0 (1 - 1/(2 + real n)) \<union> A n))" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1013 | by (simp add: image_Un) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1014 | then have "bounded (f ` ball 0 1)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1015 | apply (rule bounded_subset) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1016 | apply (auto simp: A_def algebra_simps) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1017 | done | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1018 | then show False | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1019 | using False by (simp add: fim [symmetric]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1020 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1021 | have clo_INTX: "closed(\<Inter>(range X))" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1022 | by (metis cloX closed_INT) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1023 | then have lcX: "locally compact (\<Inter>(range X))" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1024 | by (metis closed_imp_locally_compact) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1025 | have False if C: "C \<in> components (frontier S)" and boC: "bounded C" for C | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1026 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1027 | have "closed C" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1028 | by (metis C closed_components frontier_closed) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1029 | then have "compact C" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1030 | by (metis boC compact_eq_bounded_closed) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1031 | have Cco: "C \<in> components (\<Inter>(range X))" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1032 | by (metis frontierS C) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1033 | obtain K where "C \<subseteq> K" "compact K" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1034 | and Ksub: "K \<subseteq> \<Inter>(range X)" and clo: "closed(\<Inter>(range X) - K)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1035 |         proof (cases "{k. C \<subseteq> k \<and> compact k \<and> openin (subtopology euclidean (\<Inter>(range X))) k} = {}")
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1036 | case True | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1037 | then show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1038 | using Sura_Bura [OF lcX Cco \<open>compact C\<close>] boC | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1039 | by (simp add: True) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1040 | next | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1041 | case False | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1042 | then obtain L where "compact L" "C \<subseteq> L" and K: "openin (subtopology euclidean (\<Inter>x. X x)) L" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1043 | by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1044 | show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1045 | proof | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1046 | show "L \<subseteq> \<Inter>(range X)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1047 | by (metis K openin_imp_subset) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1048 | show "closed (\<Inter>(range X) - L)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1049 | by (metis closedin_diff closedin_self closedin_closed_trans [OF _ clo_INTX] K) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1050 | qed (use \<open>compact L\<close> \<open>C \<subseteq> L\<close> in auto) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1051 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1052 | obtain U V where "open U" and "compact (closure U)" and "open V" "K \<subseteq> U" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1053 |                      and V: "\<Inter>(range X) - K \<subseteq> V" and "U \<inter> V = {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1054 | using separation_normal_compact [OF \<open>compact K\<close> clo] by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1055 |         then have "U \<inter> (\<Inter> (range X) - K) = {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1056 | by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1057 |         have "(closure U - U) \<inter> (\<Inter>n. X n \<inter> closure U) \<noteq> {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1058 | proof (rule compact_imp_fip) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1059 | show "compact (closure U - U)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1060 | by (metis \<open>compact (closure U)\<close> \<open>open U\<close> compact_diff) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1061 | show "\<And>T. T \<in> range (\<lambda>n. X n \<inter> closure U) \<Longrightarrow> closed T" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1062 | by clarify (metis cloX closed_Int closed_closure) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1063 |           show "(closure U - U) \<inter> \<Inter>\<F> \<noteq> {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1064 | if "finite \<F>" and \<F>: "\<F> \<subseteq> range (\<lambda>n. X n \<inter> closure U)" for \<F> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1065 | proof | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1066 |             assume empty: "(closure U - U) \<inter> \<Inter>\<F> = {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1067 | obtain J where "finite J" and J: "\<F> = (\<lambda>n. X n \<inter> closure U) ` J" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1068 | using finite_subset_image [OF \<open>finite \<F>\<close> \<F>] by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1069 | show False | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1070 |             proof (cases "J = {}")
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1071 | case True | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1072 | with J empty have "closed U" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1073 | by (simp add: closure_subset_eq) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1074 |               have "C \<noteq> {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1075 | using C in_components_nonempty by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1076 |               then have "U \<noteq> {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1077 | using \<open>K \<subseteq> U\<close> \<open>C \<subseteq> K\<close> by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1078 | moreover have "U \<noteq> UNIV" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1079 | using \<open>compact (closure U)\<close> by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1080 | ultimately show False | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1081 | using \<open>open U\<close> \<open>closed U\<close> clopen by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1082 | next | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1083 | case False | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1084 | define j where "j \<equiv> Max J" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1085 | have "j \<in> J" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1086 | by (simp add: False \<open>finite J\<close> j_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1087 | have jmax: "\<And>m. m \<in> J \<Longrightarrow> m \<le> j" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1088 | by (simp add: j_def \<open>finite J\<close>) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1089 | have "\<Inter> ((\<lambda>n. X n \<inter> closure U) ` J) = X j \<inter> closure U" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1090 | using False jmax nestX \<open>j \<in> J\<close> by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1091 | then have "X j \<inter> closure U = X j \<inter> U" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1092 | apply safe | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1093 | using DiffI J empty apply auto[1] | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1094 | using closure_subset by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1095 | then have "openin (subtopology euclidean (X j)) (X j \<inter> closure U)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1096 | by (simp add: openin_open_Int \<open>open U\<close>) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1097 | moreover have "closedin (subtopology euclidean (X j)) (X j \<inter> closure U)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1098 | by (simp add: closedin_closed_Int) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1099 | moreover have "X j \<inter> closure U \<noteq> X j" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1100 | by (metis unboundedX \<open>compact (closure U)\<close> bounded_subset compact_eq_bounded_closed inf.order_iff) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1101 |               moreover have "X j \<inter> closure U \<noteq> {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1102 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1103 |                 have "C \<noteq> {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1104 | using C in_components_nonempty by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1105 | moreover have "C \<subseteq> X j \<inter> closure U" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1106 | using \<open>C \<subseteq> K\<close> \<open>K \<subseteq> U\<close> Ksub closure_subset by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1107 | ultimately show ?thesis by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1108 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1109 | ultimately show False | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1110 | using connX [of j] by (force simp: connected_clopen) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1111 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1112 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1113 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1114 | moreover have "(\<Inter>n. X n \<inter> closure U) = (\<Inter>n. X n) \<inter> closure U" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1115 | by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1116 | moreover have "x \<in> U" if "\<And>n. x \<in> X n" "x \<in> closure U" for x | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1117 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1118 | have "x \<notin> V" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1119 |             using \<open>U \<inter> V = {}\<close> \<open>open V\<close> closure_iff_nhds_not_empty that(2) by blast
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1120 | then show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1121 | by (metis (no_types) Diff_iff INT_I V \<open>K \<subseteq> U\<close> contra_subsetD that(1)) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1122 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1123 | ultimately show False | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1124 | by (auto simp: open_Int_closure_eq_empty [OF \<open>open V\<close>, of U]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1125 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1126 | then show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1127 | by (auto simp: False) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1128 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1129 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1130 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1131 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1132 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1133 | lemma unbounded_complement_components: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1134 | assumes C: "C \<in> components (- S)" and S: "connected S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1135 | and prev: "if bounded S then connected(frontier S) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1136 | else \<forall>C \<in> components(frontier S). \<not> bounded C" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1137 | shows "\<not> bounded C" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1138 | proof (cases "bounded S") | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1139 | case True | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1140 | with prev have "S \<noteq> UNIV" and confr: "connected(frontier S)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1141 | by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1142 | obtain w where C_ccsw: "C = connected_component_set (- S) w" and "w \<notin> S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1143 | using C by (auto simp: components_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1144 | show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1145 |   proof (cases "S = {}")
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1146 | case True with C show ?thesis by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1147 | next | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1148 | case False | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1149 | show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1150 | proof | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1151 | assume "bounded C" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1152 |       then have "outside C \<noteq> {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1153 | using outside_bounded_nonempty by metis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1154 | then obtain z where z: "\<not> bounded (connected_component_set (- C) z)" and "z \<notin> C" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1155 | by (auto simp: outside_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1156 | have clo_ccs: "closed (connected_component_set (- S) x)" for x | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1157 | by (simp add: closed_Compl closed_connected_component openS) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1158 | have "connected_component_set (- S) w = connected_component_set (- S) z" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1159 | proof (rule joinable_connected_component_eq [OF confr]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1160 | show "frontier S \<subseteq> - S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1161 | using openS by (auto simp: frontier_def interior_open) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1162 |         have False if "connected_component_set (- S) w \<inter> frontier (- S) = {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1163 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1164 |           have "C \<inter> frontier S = {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1165 | using that by (simp add: C_ccsw) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1166 | then show False | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1167 | by (metis C_ccsw ComplI Compl_eq_Compl_iff Diff_subset False \<open>w \<notin> S\<close> clo_ccs closure_closed compl_bot_eq connected_component_eq_UNIV connected_component_eq_empty empty_subsetI frontier_complement frontier_def frontier_not_empty frontier_of_connected_component_subset le_inf_iff subset_antisym) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1168 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1169 |         then show "connected_component_set (- S) w \<inter> frontier S \<noteq> {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1170 | by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1171 |         have *: "\<lbrakk>frontier C \<subseteq> C; frontier C \<subseteq> F; frontier C \<noteq> {}\<rbrakk> \<Longrightarrow> C \<inter> F \<noteq> {}" for C F::"complex set"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1172 | by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1173 |         have "connected_component_set (- S) z \<inter> frontier (- S) \<noteq> {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1174 | proof (rule *) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1175 | show "frontier (connected_component_set (- S) z) \<subseteq> connected_component_set (- S) z" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1176 | by (auto simp: closed_Compl closed_connected_component frontier_def openS) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1177 | show "frontier (connected_component_set (- S) z) \<subseteq> frontier (- S)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1178 | using frontier_of_connected_component_subset by fastforce | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1179 | have "\<not> bounded (-S)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1180 | by (simp add: True cobounded_imp_unbounded) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1181 |           then have "connected_component_set (- S) z \<noteq> {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1182 | apply (simp only: connected_component_eq_empty) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1183 | using confr openS \<open>bounded C\<close> \<open>w \<notin> S\<close> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1184 | apply (simp add: frontier_def interior_open C_ccsw) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1185 | by (metis ComplI Compl_eq_Diff_UNIV connected_UNIV closed_closure closure_subset connected_component_eq_self | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1186 | connected_diff_open_from_closed subset_UNIV) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1187 |           then show "frontier (connected_component_set (- S) z) \<noteq> {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1188 | apply (simp add: frontier_eq_empty connected_component_eq_UNIV) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1189 | apply (metis False compl_top_eq double_compl) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1190 | done | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1191 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1192 |         then show "connected_component_set (- S) z \<inter> frontier S \<noteq> {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1193 | by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1194 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1195 | then show False | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1196 | by (metis C_ccsw Compl_iff \<open>w \<notin> S\<close> \<open>z \<notin> C\<close> connected_component_eq_empty connected_component_idemp) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1197 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1198 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1199 | next | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1200 | case False | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1201 | obtain w where C_ccsw: "C = connected_component_set (- S) w" and "w \<notin> S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1202 | using C by (auto simp: components_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1203 | have "frontier (connected_component_set (- S) w) \<subseteq> connected_component_set (- S) w" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1204 | by (simp add: closed_Compl closed_connected_component frontier_subset_eq openS) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1205 | moreover have "frontier (connected_component_set (- S) w) \<subseteq> frontier S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1206 | using frontier_complement frontier_of_connected_component_subset by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1207 |   moreover have "frontier (connected_component_set (- S) w) \<noteq> {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1208 | by (metis C C_ccsw False bounded_empty compl_top_eq connected_component_eq_UNIV double_compl frontier_not_empty in_components_nonempty) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1209 | ultimately obtain z where zin: "z \<in> frontier S" and z: "z \<in> connected_component_set (- S) w" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1210 | by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1211 | have *: "connected_component_set (frontier S) z \<in> components(frontier S)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1212 | by (simp add: \<open>z \<in> frontier S\<close> componentsI) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1213 | with prev False have "\<not> bounded (connected_component_set (frontier S) z)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1214 | by simp | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1215 | moreover have "connected_component (- S) w = connected_component (- S) z" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1216 | using connected_component_eq [OF z] by force | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1217 | ultimately show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1218 | by (metis C_ccsw * zin bounded_subset closed_Compl closure_closed connected_component_maximal | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1219 | connected_component_refl connected_connected_component frontier_closures in_components_subset le_inf_iff mem_Collect_eq openS) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1220 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1221 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1222 | lemma empty_inside: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1223 | assumes "connected S" "\<And>C. C \<in> components (- S) \<Longrightarrow> ~bounded C" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1224 |   shows "inside S = {}"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1225 | using assms by (auto simp: components_def inside_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1226 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1227 | lemma empty_inside_imp_simply_connected: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1228 |   "\<lbrakk>connected S; inside S = {}\<rbrakk> \<Longrightarrow> simply_connected S"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1229 | by (metis ComplI inside_Un_outside openS outside_mono simply_connected_eq_winding_number_zero subsetCE sup_bot.left_neutral winding_number_zero_in_outside) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1230 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1231 | end | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1232 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1233 | proposition | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1234 | fixes S :: "complex set" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1235 | assumes "open S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1236 | shows simply_connected_eq_frontier_properties: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1237 | "simply_connected S \<longleftrightarrow> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1238 | connected S \<and> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1239 | (if bounded S then connected(frontier S) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1240 | else (\<forall>C \<in> components(frontier S). \<not>bounded C))" (is "?fp") | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1241 | and simply_connected_eq_unbounded_complement_components: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1242 | "simply_connected S \<longleftrightarrow> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1243 | connected S \<and> (\<forall>C \<in> components(- S). \<not>bounded C)" (is "?ucc") | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1244 | and simply_connected_eq_empty_inside: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1245 | "simply_connected S \<longleftrightarrow> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1246 |           connected S \<and> inside S = {}" (is "?ei")
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1247 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1248 | interpret SC_Chain | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1249 | using assms by (simp add: SC_Chain_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1250 | have "?fp \<and> ?ucc \<and> ?ei" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1251 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1252 | have *: "\<lbrakk>\<alpha> \<Longrightarrow> \<beta>; \<beta> \<Longrightarrow> \<gamma>; \<gamma> \<Longrightarrow> \<delta>; \<delta> \<Longrightarrow> \<alpha>\<rbrakk> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1253 | \<Longrightarrow> (\<alpha> \<longleftrightarrow> \<beta>) \<and> (\<alpha> \<longleftrightarrow> \<gamma>) \<and> (\<alpha> \<longleftrightarrow> \<delta>)" for \<alpha> \<beta> \<gamma> \<delta> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1254 | by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1255 | show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1256 | apply (rule *) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1257 | using frontier_properties simply_connected_imp_connected apply blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1258 | apply clarify | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1259 | using unbounded_complement_components simply_connected_imp_connected apply blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1260 | using empty_inside apply blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1261 | using empty_inside_imp_simply_connected apply blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1262 | done | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1263 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1264 | then show ?fp ?ucc ?ei | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1265 | by safe | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1266 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1267 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1268 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1269 | lemma simply_connected_iff_simple: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1270 | fixes S :: "complex set" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1271 | assumes "open S" "bounded S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1272 | shows "simply_connected S \<longleftrightarrow> connected S \<and> connected(- S)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1273 | apply (simp add: simply_connected_eq_unbounded_complement_components assms, safe) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1274 | apply (metis DIM_complex assms(2) cobounded_has_bounded_component double_compl order_refl) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1275 | by (meson assms inside_bounded_complement_connected_empty simply_connected_eq_empty_inside simply_connected_eq_unbounded_complement_components) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1276 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1277 | subsection\<open>Further equivalences based on continuous logs and sqrts\<close> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1278 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1279 | context SC_Chain | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1280 | begin | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1281 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1282 | lemma continuous_log: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1283 | fixes f :: "complex\<Rightarrow>complex" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1284 | assumes S: "simply_connected S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1285 | and contf: "continuous_on S f" and nz: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1286 | shows "\<exists>g. continuous_on S g \<and> (\<forall>z \<in> S. f z = exp(g z))" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1287 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1288 |   consider "S = {}" | "S homeomorphic ball (0::complex) 1"
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1289 | using simply_connected_eq_homeomorphic_to_disc [OF openS] S by metis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1290 | then show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1291 | proof cases | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1292 | case 1 then show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1293 | by simp | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1294 | next | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1295 | case 2 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1296 | then obtain h k :: "complex\<Rightarrow>complex" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1297 | where kh: "\<And>x. x \<in> S \<Longrightarrow> k(h x) = x" and him: "h ` S = ball 0 1" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1298 | and conth: "continuous_on S h" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1299 | and hk: "\<And>y. y \<in> ball 0 1 \<Longrightarrow> h(k y) = y" and kim: "k ` ball 0 1 = S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1300 | and contk: "continuous_on (ball 0 1) k" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1301 | unfolding homeomorphism_def homeomorphic_def by metis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1302 | obtain g where contg: "continuous_on (ball 0 1) g" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1303 | and expg: "\<And>z. z \<in> ball 0 1 \<Longrightarrow> (f \<circ> k) z = exp (g z)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1304 | proof (rule continuous_logarithm_on_ball) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1305 | show "continuous_on (ball 0 1) (f \<circ> k)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1306 | apply (rule continuous_on_compose [OF contk]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1307 | using kim continuous_on_subset [OF contf] | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1308 | by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1309 | show "\<And>z. z \<in> ball 0 1 \<Longrightarrow> (f \<circ> k) z \<noteq> 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1310 | using kim nz by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1311 | qed auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1312 | then show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1313 | by (metis comp_apply conth continuous_on_compose him imageI kh) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1314 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1315 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1316 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1317 | lemma continuous_sqrt: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1318 | fixes f :: "complex\<Rightarrow>complex" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1319 | assumes contf: "continuous_on S f" and nz: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1320 | and prev: "\<And>f::complex\<Rightarrow>complex. | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1321 | \<lbrakk>continuous_on S f; \<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0\<rbrakk> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1322 | \<Longrightarrow> \<exists>g. continuous_on S g \<and> (\<forall>z \<in> S. f z = exp(g z))" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1323 | shows "\<exists>g. continuous_on S g \<and> (\<forall>z \<in> S. f z = (g z)\<^sup>2)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1324 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1325 | obtain g where contg: "continuous_on S g" and geq: "\<And>z. z \<in> S \<Longrightarrow> f z = exp(g z)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1326 | using contf nz prev by metis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1327 | show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1328 | proof (intro exI ballI conjI) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1329 | show "continuous_on S (\<lambda>z. exp(g z/2))" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1330 | by (intro continuous_intros) (auto simp: contg) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1331 | show "\<And>z. z \<in> S \<Longrightarrow> f z = (exp (g z/2))\<^sup>2" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1332 | by (metis (no_types, lifting) divide_inverse exp_double geq mult.left_commute mult.right_neutral right_inverse zero_neq_numeral) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1333 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1334 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1335 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1336 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1337 | lemma continuous_sqrt_imp_simply_connected: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1338 | assumes "connected S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1339 | and prev: "\<And>f::complex\<Rightarrow>complex. \<lbrakk>continuous_on S f; \<forall>z \<in> S. f z \<noteq> 0\<rbrakk> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1340 | \<Longrightarrow> \<exists>g. continuous_on S g \<and> (\<forall>z \<in> S. f z = (g z)\<^sup>2)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1341 | shows "simply_connected S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1342 | proof (clarsimp simp add: simply_connected_eq_holomorphic_sqrt [OF openS] \<open>connected S\<close>) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1343 | fix f | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1344 | assume "f holomorphic_on S" and nz: "\<forall>z\<in>S. f z \<noteq> 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1345 | then obtain g where contg: "continuous_on S g" and geq: "\<And>z. z \<in> S \<Longrightarrow> f z = (g z)\<^sup>2" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1346 | by (metis holomorphic_on_imp_continuous_on prev) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1347 | show "\<exists>g. g holomorphic_on S \<and> (\<forall>z\<in>S. f z = (g z)\<^sup>2)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1348 | proof (intro exI ballI conjI) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1349 | show "g holomorphic_on S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1350 | proof (clarsimp simp add: holomorphic_on_open [OF openS]) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1351 | fix z | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1352 | assume "z \<in> S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1353 | with nz geq have "g z \<noteq> 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1354 | by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1355 | obtain \<delta> where "0 < \<delta>" "\<And>w. \<lbrakk>w \<in> S; dist w z < \<delta>\<rbrakk> \<Longrightarrow> dist (g w) (g z) < cmod (g z)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1356 | using contg [unfolded continuous_on_iff] by (metis \<open>g z \<noteq> 0\<close> \<open>z \<in> S\<close> zero_less_norm_iff) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1357 | then have \<delta>: "\<And>w. \<lbrakk>w \<in> S; w \<in> ball z \<delta>\<rbrakk> \<Longrightarrow> g w + g z \<noteq> 0" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1358 | apply (clarsimp simp: dist_norm) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1359 | by (metis \<open>g z \<noteq> 0\<close> add_diff_cancel_left' diff_0_right norm_eq_zero norm_increases_online norm_minus_commute norm_not_less_zero not_less_iff_gr_or_eq) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1360 | have *: "(\<lambda>x. (f x - f z) / (x - z) / (g x + g z)) \<midarrow>z\<rightarrow> deriv f z / (g z + g z)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1361 | apply (intro tendsto_intros) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1362 | using SC_Chain.openS SC_Chain_axioms \<open>f holomorphic_on S\<close> \<open>z \<in> S\<close> has_field_derivativeD holomorphic_derivI apply fastforce | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1363 | using \<open>z \<in> S\<close> contg continuous_on_eq_continuous_at isCont_def openS apply blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1364 | by (simp add: \<open>g z \<noteq> 0\<close>) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1365 | then have "(g has_field_derivative deriv f z / (g z + g z)) (at z)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1366 | unfolding DERIV_iff2 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1367 | proof (rule Lim_transform_within_open) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1368 | show "open (ball z \<delta> \<inter> S)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1369 | by (simp add: openS open_Int) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1370 | show "z \<in> ball z \<delta> \<inter> S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1371 | using \<open>z \<in> S\<close> \<open>0 < \<delta>\<close> by simp | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1372 | show "\<And>x. \<lbrakk>x \<in> ball z \<delta> \<inter> S; x \<noteq> z\<rbrakk> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1373 | \<Longrightarrow> (f x - f z) / (x - z) / (g x + g z) = (g x - g z) / (x - z)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1374 | using \<delta> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1375 | apply (simp add: geq \<open>z \<in> S\<close> divide_simps) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1376 | apply (auto simp: algebra_simps power2_eq_square) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1377 | done | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1378 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1379 | then show "\<exists>f'. (g has_field_derivative f') (at z)" .. | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1380 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1381 | qed (use geq in auto) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1382 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1383 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1384 | end | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1385 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1386 | proposition | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1387 | fixes S :: "complex set" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1388 | assumes "open S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1389 | shows simply_connected_eq_continuous_log: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1390 | "simply_connected S \<longleftrightarrow> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1391 | connected S \<and> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1392 | (\<forall>f::complex\<Rightarrow>complex. continuous_on S f \<and> (\<forall>z \<in> S. f z \<noteq> 0) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1393 | \<longrightarrow> (\<exists>g. continuous_on S g \<and> (\<forall>z \<in> S. f z = exp (g z))))" (is "?log") | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1394 | and simply_connected_eq_continuous_sqrt: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1395 | "simply_connected S \<longleftrightarrow> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1396 | connected S \<and> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1397 | (\<forall>f::complex\<Rightarrow>complex. continuous_on S f \<and> (\<forall>z \<in> S. f z \<noteq> 0) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1398 | \<longrightarrow> (\<exists>g. continuous_on S g \<and> (\<forall>z \<in> S. f z = (g z)\<^sup>2)))" (is "?sqrt") | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1399 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1400 | interpret SC_Chain | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1401 | using assms by (simp add: SC_Chain_def) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1402 | have "?log \<and> ?sqrt" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1403 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1404 | have *: "\<lbrakk>\<alpha> \<Longrightarrow> \<beta>; \<beta> \<Longrightarrow> \<gamma>; \<gamma> \<Longrightarrow> \<alpha>\<rbrakk> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1405 | \<Longrightarrow> (\<alpha> \<longleftrightarrow> \<beta>) \<and> (\<alpha> \<longleftrightarrow> \<gamma>)" for \<alpha> \<beta> \<gamma> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1406 | by blast | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1407 | show ?thesis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1408 | apply (rule *) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1409 | apply (simp add: local.continuous_log winding_number_zero) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1410 | apply (simp add: continuous_sqrt) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1411 | apply (simp add: continuous_sqrt_imp_simply_connected) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1412 | done | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1413 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1414 | then show ?log ?sqrt | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1415 | by safe | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1416 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1417 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1418 | |
| 66941 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1419 | subsection\<open>More Borsukian results\<close> | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1420 | |
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1421 | lemma Borsukian_componentwise_eq: | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1422 | fixes S :: "'a::euclidean_space set" | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1423 | assumes S: "locally connected S \<or> compact S" | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1424 | shows "Borsukian S \<longleftrightarrow> (\<forall>C \<in> components S. Borsukian C)" | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1425 | proof - | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1426 |   have *: "ANR(-{0::complex})"
 | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1427 | by (simp add: ANR_delete open_Compl open_imp_ANR) | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1428 | show ?thesis | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1429 | using cohomotopically_trivial_on_components [OF assms *] by (auto simp: Borsukian_alt) | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1430 | qed | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1431 | |
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1432 | lemma Borsukian_componentwise: | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1433 | fixes S :: "'a::euclidean_space set" | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1434 | assumes "locally connected S \<or> compact S" "\<And>C. C \<in> components S \<Longrightarrow> Borsukian C" | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1435 | shows "Borsukian S" | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1436 | by (metis Borsukian_componentwise_eq assms) | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1437 | |
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1438 | lemma simply_connected_eq_Borsukian: | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1439 | fixes S :: "complex set" | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1440 | shows "open S \<Longrightarrow> (simply_connected S \<longleftrightarrow> connected S \<and> Borsukian S)" | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1441 | by (auto simp: simply_connected_eq_continuous_log Borsukian_continuous_logarithm) | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1442 | |
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1443 | lemma Borsukian_eq_simply_connected: | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1444 | fixes S :: "complex set" | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1445 | shows "open S \<Longrightarrow> Borsukian S \<longleftrightarrow> (\<forall>C \<in> components S. simply_connected C)" | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1446 | apply (auto simp: Borsukian_componentwise_eq open_imp_locally_connected) | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1447 | using in_components_connected open_components simply_connected_eq_Borsukian apply blast | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1448 | using open_components simply_connected_eq_Borsukian by blast | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1449 | |
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1450 | lemma Borsukian_separation_open_closed: | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1451 | fixes S :: "complex set" | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1452 | assumes S: "open S \<or> closed S" and "bounded S" | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1453 | shows "Borsukian S \<longleftrightarrow> connected(- S)" | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1454 | using S | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1455 | proof | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1456 | assume "open S" | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1457 | show ?thesis | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1458 | unfolding Borsukian_eq_simply_connected [OF \<open>open S\<close>] | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1459 | by (meson \<open>open S\<close> \<open>bounded S\<close> bounded_subset in_components_connected in_components_subset nonseparation_by_component_eq open_components simply_connected_iff_simple) | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1460 | next | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1461 | assume "closed S" | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1462 | with \<open>bounded S\<close> show ?thesis | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1463 | by (simp add: Borsukian_separation_compact compact_eq_bounded_closed) | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1464 | qed | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1465 | |
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1466 | |
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1467 | subsection\<open>Finally, the Riemann Mapping Theorem\<close> | 
| 
c67bb79a0ceb
More topological results overlooked last time
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1468 | |
| 66826 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1469 | theorem Riemann_mapping_theorem: | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1470 | "open S \<and> simply_connected S \<longleftrightarrow> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1471 |      S = {} \<or> S = UNIV \<or>
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1472 | (\<exists>f g. f holomorphic_on S \<and> g holomorphic_on ball 0 1 \<and> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1473 | (\<forall>z \<in> S. f z \<in> ball 0 1 \<and> g(f z) = z) \<and> | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1474 | (\<forall>z \<in> ball 0 1. g z \<in> S \<and> f(g z) = z))" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1475 | (is "_ = ?rhs") | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1476 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1477 | have "simply_connected S \<longleftrightarrow> ?rhs" if "open S" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1478 | by (simp add: simply_connected_eq_biholomorphic_to_disc that) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1479 | moreover have "open S" if "?rhs" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1480 | proof - | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1481 |     { fix f g
 | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1482 | assume g: "g holomorphic_on ball 0 1" "\<forall>z\<in>ball 0 1. g z \<in> S \<and> f (g z) = z" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1483 | and "\<forall>z\<in>S. cmod (f z) < 1 \<and> g (f z) = z" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1484 | then have "S = g ` (ball 0 1)" | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1485 | by (force simp:) | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1486 | then have "open S" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: 
66826diff
changeset | 1487 | by (metis open_ball g inj_on_def open_mapping_thm3) | 
| 66826 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1488 | } | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1489 | with that show "open S" by auto | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1490 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1491 | ultimately show ?thesis by metis | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1492 | qed | 
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1493 | |
| 
0d60d2118544
Session HOL-Analysis: Moebius functions and the Riemann mapping theorem.
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1494 | end |