| author | wenzelm | 
| Sat, 18 Jan 2025 11:03:18 +0100 | |
| changeset 81904 | aa28d82d6b66 | 
| parent 80175 | 200107cdd3ac | 
| permissions | -rw-r--r-- | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1 | (* Title: HOL/Analysis/Arcwise_Connected.thy | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2 | Authors: LC Paulson, based on material from HOL Light | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3 | *) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4 | |
| 69517 | 5 | section \<open>Arcwise-Connected Sets\<close> | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 6 | |
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 7 | theory Arcwise_Connected | 
| 69517 | 8 | imports Path_Connected Ordered_Euclidean_Space "HOL-Computational_Algebra.Primes" | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 9 | begin | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 10 | |
| 71025 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70817diff
changeset | 11 | lemma path_connected_interval [simp]: | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70817diff
changeset | 12 | fixes a b::"'a::ordered_euclidean_space" | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70817diff
changeset | 13 |   shows "path_connected {a..b}"
 | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70817diff
changeset | 14 | using is_interval_cc is_interval_path_connected by blast | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70817diff
changeset | 15 | |
| 71028 
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
 immler parents: 
71025diff
changeset | 16 | lemma segment_to_closest_point: | 
| 
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
 immler parents: 
71025diff
changeset | 17 | fixes S :: "'a :: euclidean_space set" | 
| 
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
 immler parents: 
71025diff
changeset | 18 |   shows "\<lbrakk>closed S; S \<noteq> {}\<rbrakk> \<Longrightarrow> open_segment a (closest_point S a) \<inter> S = {}"
 | 
| 72531 | 19 | unfolding disjoint_iff | 
| 20 | by (metis closest_point_le dist_commute dist_in_open_segment not_le) | |
| 71028 
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
 immler parents: 
71025diff
changeset | 21 | |
| 
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
 immler parents: 
71025diff
changeset | 22 | lemma segment_to_point_exists: | 
| 
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
 immler parents: 
71025diff
changeset | 23 | fixes S :: "'a :: euclidean_space set" | 
| 
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
 immler parents: 
71025diff
changeset | 24 |     assumes "closed S" "S \<noteq> {}"
 | 
| 
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
 immler parents: 
71025diff
changeset | 25 |     obtains b where "b \<in> S" "open_segment a b \<inter> S = {}"
 | 
| 
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
 immler parents: 
71025diff
changeset | 26 | by (metis assms segment_to_closest_point closest_point_exists that) | 
| 
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
 immler parents: 
71025diff
changeset | 27 | |
| 
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
 immler parents: 
71025diff
changeset | 28 | |
| 69683 | 29 | subsection \<open>The Brouwer reduction theorem\<close> | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 30 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 31 | theorem Brouwer_reduction_theorem_gen: | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 32 | fixes S :: "'a::euclidean_space set" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 33 | assumes "closed S" "\<phi> S" | 
| 69745 | 34 | and \<phi>: "\<And>F. \<lbrakk>\<And>n. closed(F n); \<And>n. \<phi>(F n); \<And>n. F(Suc n) \<subseteq> F n\<rbrakk> \<Longrightarrow> \<phi>(\<Inter>(range F))" | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 35 | obtains T where "T \<subseteq> S" "closed T" "\<phi> T" "\<And>U. \<lbrakk>U \<subseteq> S; closed U; \<phi> U\<rbrakk> \<Longrightarrow> \<not> (U \<subset> T)" | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 36 | proof - | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 37 | obtain B :: "nat \<Rightarrow> 'a set" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 38 | where "inj B" "\<And>n. open(B n)" and open_cov: "\<And>S. open S \<Longrightarrow> \<exists>K. S = \<Union>(B ` K)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 39 | by (metis Setcompr_eq_image that univ_second_countable_sequence) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 40 |   define A where "A \<equiv> rec_nat S (\<lambda>n a. if \<exists>U. U \<subseteq> a \<and> closed U \<and> \<phi> U \<and> U \<inter> (B n) = {}
 | 
| 67613 | 41 |                                         then SOME U. U \<subseteq> a \<and> closed U \<and> \<phi> U \<and> U \<inter> (B n) = {}
 | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 42 | else a)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 43 | have [simp]: "A 0 = S" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 44 | by (simp add: A_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 45 |   have ASuc: "A(Suc n) = (if \<exists>U. U \<subseteq> A n \<and> closed U \<and> \<phi> U \<and> U \<inter> (B n) = {}
 | 
| 67613 | 46 |                           then SOME U. U \<subseteq> A n \<and> closed U \<and> \<phi> U \<and> U \<inter> (B n) = {}
 | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 47 | else A n)" for n | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 48 | by (auto simp: A_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 49 | have sub: "\<And>n. A(Suc n) \<subseteq> A n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 50 | by (auto simp: ASuc dest!: someI_ex) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 51 | have subS: "A n \<subseteq> S" for n | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 52 | by (induction n) (use sub in auto) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 53 | have clo: "closed (A n) \<and> \<phi> (A n)" for n | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 54 | by (induction n) (auto simp: assms ASuc dest!: someI_ex) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 55 | show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 56 | proof | 
| 69745 | 57 | show "\<Inter>(range A) \<subseteq> S" | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 58 | using \<open>\<And>n. A n \<subseteq> S\<close> by blast | 
| 69313 | 59 | show "closed (\<Inter>(A ` UNIV))" | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 60 | using clo by blast | 
| 69313 | 61 | show "\<phi> (\<Inter>(A ` UNIV))" | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 62 | by (simp add: clo \<phi> sub) | 
| 69313 | 63 | show "\<not> U \<subset> \<Inter>(A ` UNIV)" if "U \<subseteq> S" "closed U" "\<phi> U" for U | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 64 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 65 | have "\<exists>y. x \<notin> A y" if "x \<notin> U" and Usub: "U \<subseteq> (\<Inter>x. A x)" for x | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 66 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 67 | obtain e where "e > 0" and e: "ball x e \<subseteq> -U" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 68 | using \<open>closed U\<close> \<open>x \<notin> U\<close> openE [of "-U"] by blast | 
| 69313 | 69 | moreover obtain K where K: "ball x e = \<Union>(B ` K)" | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 70 | using open_cov [of "ball x e"] by auto | 
| 69313 | 71 | ultimately have "\<Union>(B ` K) \<subseteq> -U" | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 72 | by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 73 |         have "K \<noteq> {}"
 | 
| 69313 | 74 | using \<open>0 < e\<close> \<open>ball x e = \<Union>(B ` K)\<close> by auto | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 75 | then obtain n where "n \<in> K" "x \<in> B n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 76 | by (metis K UN_E \<open>0 < e\<close> centre_in_ball) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 77 |         then have "U \<inter> B n = {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 78 | using K e by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 79 | show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 80 |         proof (cases "\<exists>U\<subseteq>A n. closed U \<and> \<phi> U \<and> U \<inter> B n = {}")
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 81 | case True | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 82 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 83 | apply (rule_tac x="Suc n" in exI) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 84 | apply (simp add: ASuc) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 85 | apply (erule someI2_ex) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 86 | using \<open>x \<in> B n\<close> by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 87 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 88 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 89 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 90 |             by (meson Inf_lower Usub \<open>U \<inter> B n = {}\<close> \<open>\<phi> U\<close> \<open>closed U\<close> range_eqI subset_trans)
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 91 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 92 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 93 | with that show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 94 | by (meson Inter_iff psubsetE rangeI subsetI) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 95 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 96 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 97 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 98 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 99 | corollary Brouwer_reduction_theorem: | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 100 | fixes S :: "'a::euclidean_space set" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 101 |   assumes "compact S" "\<phi> S" "S \<noteq> {}"
 | 
| 69745 | 102 |       and \<phi>: "\<And>F. \<lbrakk>\<And>n. compact(F n); \<And>n. F n \<noteq> {}; \<And>n. \<phi>(F n); \<And>n. F(Suc n) \<subseteq> F n\<rbrakk> \<Longrightarrow> \<phi>(\<Inter>(range F))"
 | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 103 |   obtains T where "T \<subseteq> S" "compact T" "T \<noteq> {}" "\<phi> T"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 104 |                   "\<And>U. \<lbrakk>U \<subseteq> S; closed U; U \<noteq> {}; \<phi> U\<rbrakk> \<Longrightarrow> \<not> (U \<subset> T)"
 | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 105 | proof (rule Brouwer_reduction_theorem_gen [of S "\<lambda>T. T \<noteq> {} \<and> T \<subseteq> S \<and> \<phi> T"])
 | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 106 | fix F | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 107 | assume cloF: "\<And>n. closed (F n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 108 |      and F: "\<And>n. F n \<noteq> {} \<and> F n \<subseteq> S \<and> \<phi> (F n)" and Fsub: "\<And>n. F (Suc n) \<subseteq> F n"
 | 
| 69313 | 109 |   show "\<Inter>(F ` UNIV) \<noteq> {} \<and> \<Inter>(F ` UNIV) \<subseteq> S \<and> \<phi> (\<Inter>(F ` UNIV))"
 | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 110 | proof (intro conjI) | 
| 69313 | 111 |     show "\<Inter>(F ` UNIV) \<noteq> {}"
 | 
| 72531 | 112 | by (metis F Fsub \<open>compact S\<close> cloF closed_Int_compact compact_nest inf.orderE lift_Suc_antimono_le) | 
| 69313 | 113 | show " \<Inter>(F ` UNIV) \<subseteq> S" | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 114 | using F by blast | 
| 69313 | 115 | show "\<phi> (\<Inter>(F ` UNIV))" | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 116 | by (metis F Fsub \<phi> \<open>compact S\<close> cloF closed_Int_compact inf.orderE) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 117 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 118 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 119 |   show "S \<noteq> {} \<and> S \<subseteq> S \<and> \<phi> S"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 120 | by (simp add: assms) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 121 | qed (meson assms compact_imp_closed seq_compact_closed_subset seq_compact_eq_compact)+ | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 122 | |
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 123 | |
| 70136 | 124 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Arcwise Connections\<close>(*FIX ME this subsection is empty(?) *) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 125 | |
| 69683 | 126 | subsection\<open>Density of points with dyadic rational coordinates\<close> | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 127 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 128 | proposition closure_dyadic_rationals: | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 129 | "closure (\<Union>k. \<Union>f \<in> Basis \<rightarrow> \<int>. | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 130 |                    { \<Sum>i :: 'a :: euclidean_space \<in> Basis. (f i / 2^k) *\<^sub>R i }) = UNIV"
 | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 131 | proof - | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 132 |   have "x \<in> closure (\<Union>k. \<Union>f \<in> Basis \<rightarrow> \<int>. {\<Sum>i \<in> Basis. (f i / 2^k) *\<^sub>R i})" for x::'a
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 133 | proof (clarsimp simp: closure_approachable) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 134 | fix e::real | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 135 | assume "e > 0" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 136 |     then obtain k where k: "(1/2)^k < e/DIM('a)"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 137 | by (meson DIM_positive divide_less_eq_1_pos of_nat_0_less_iff one_less_numeral_iff real_arch_pow_inv semiring_norm(76) zero_less_divide_iff zero_less_numeral) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 138 | have "dist (\<Sum>i\<in>Basis. (real_of_int \<lfloor>2^k*(x \<bullet> i)\<rfloor> / 2^k) *\<^sub>R i) x = | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 139 | dist (\<Sum>i\<in>Basis. (real_of_int \<lfloor>2^k*(x \<bullet> i)\<rfloor> / 2^k) *\<^sub>R i) (\<Sum>i\<in>Basis. (x \<bullet> i) *\<^sub>R i)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 140 | by (simp add: euclidean_representation) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 141 | also have "... = norm ((\<Sum>i\<in>Basis. (real_of_int \<lfloor>2^k*(x \<bullet> i)\<rfloor> / 2^k) *\<^sub>R i - (x \<bullet> i) *\<^sub>R i))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 142 | by (simp add: dist_norm sum_subtractf) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 143 |     also have "... \<le> DIM('a)*((1/2)^k)"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 144 | proof (rule sum_norm_bound, simp add: algebra_simps) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 145 | fix i::'a | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 146 | assume "i \<in> Basis" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 147 | then have "norm ((real_of_int \<lfloor>x \<bullet> i*2^k\<rfloor> / 2^k) *\<^sub>R i - (x \<bullet> i) *\<^sub>R i) = | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 148 | \<bar>real_of_int \<lfloor>x \<bullet> i*2^k\<rfloor> / 2^k - x \<bullet> i\<bar>" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 149 | by (simp add: scaleR_left_diff_distrib [symmetric]) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 150 | also have "... \<le> (1/2) ^ k" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 151 | by (simp add: divide_simps) linarith | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 152 | finally show "norm ((real_of_int \<lfloor>x \<bullet> i*2^k\<rfloor> / 2^k) *\<^sub>R i - (x \<bullet> i) *\<^sub>R i) \<le> (1/2) ^ k" . | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 153 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 154 |     also have "... < DIM('a)*(e/DIM('a))"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 155 | using DIM_positive k linordered_comm_semiring_strict_class.comm_mult_strict_left_mono of_nat_0_less_iff by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 156 | also have "... = e" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 157 | by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 158 | finally have "dist (\<Sum>i\<in>Basis. (\<lfloor>2^k*(x \<bullet> i)\<rfloor> / 2^k) *\<^sub>R i) x < e" . | 
| 72531 | 159 | with Ints_of_int | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 160 | show "\<exists>k. \<exists>f \<in> Basis \<rightarrow> \<int>. dist (\<Sum>b\<in>Basis. (f b / 2^k) *\<^sub>R b) x < e" | 
| 72531 | 161 | by fastforce | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 162 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 163 | then show ?thesis by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 164 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 165 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 166 | corollary closure_rational_coordinates: | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 167 |     "closure (\<Union>f \<in> Basis \<rightarrow> \<rat>. { \<Sum>i :: 'a :: euclidean_space \<in> Basis. f i *\<^sub>R i }) = UNIV"
 | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 168 | proof - | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 169 |   have *: "(\<Union>k. \<Union>f \<in> Basis \<rightarrow> \<int>. { \<Sum>i::'a \<in> Basis. (f i / 2^k) *\<^sub>R i })
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 170 |            \<subseteq> (\<Union>f \<in> Basis \<rightarrow> \<rat>. { \<Sum>i \<in> Basis. f i *\<^sub>R i })"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 171 | proof clarsimp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 172 | fix k and f :: "'a \<Rightarrow> real" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 173 | assume f: "f \<in> Basis \<rightarrow> \<int>" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 174 | show "\<exists>x \<in> Basis \<rightarrow> \<rat>. (\<Sum>i \<in> Basis. (f i / 2^k) *\<^sub>R i) = (\<Sum>i \<in> Basis. x i *\<^sub>R i)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 175 | apply (rule_tac x="\<lambda>i. f i / 2^k" in bexI) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 176 | using Ints_subset_Rats f by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 177 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 178 | show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 179 | using closure_dyadic_rationals closure_mono [OF *] by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 180 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 181 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 182 | lemma closure_dyadic_rationals_in_convex_set: | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 183 |    "\<lbrakk>convex S; interior S \<noteq> {}\<rbrakk>
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 184 | \<Longrightarrow> closure(S \<inter> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 185 | (\<Union>k. \<Union>f \<in> Basis \<rightarrow> \<int>. | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 186 |                    { \<Sum>i :: 'a :: euclidean_space \<in> Basis. (f i / 2^k) *\<^sub>R i })) =
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 187 | closure S" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 188 | by (simp add: closure_dyadic_rationals closure_convex_Int_superset) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 189 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 190 | lemma closure_rationals_in_convex_set: | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 191 |    "\<lbrakk>convex S; interior S \<noteq> {}\<rbrakk>
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 192 |     \<Longrightarrow> closure(S \<inter> (\<Union>f \<in> Basis \<rightarrow> \<rat>. { \<Sum>i :: 'a :: euclidean_space \<in> Basis. f i *\<^sub>R i })) =
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 193 | closure S" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 194 | by (simp add: closure_rational_coordinates closure_convex_Int_superset) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 195 | |
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 196 | |
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 197 | text\<open> Every path between distinct points contains an arc, and hence | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 198 | path connection is equivalent to arcwise connection for distinct points. | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 199 | The proof is based on Whyburn's "Topological Analysis".\<close> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 200 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 201 | lemma closure_dyadic_rationals_in_convex_set_pos_1: | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 202 | fixes S :: "real set" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 203 |   assumes "convex S" and intnz: "interior S \<noteq> {}" and pos: "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> x"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 204 |     shows "closure(S \<inter> (\<Union>k m. {of_nat m / 2^k})) = closure S"
 | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 205 | proof - | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 206 | have "\<exists>m. f 1/2^k = real m / 2^k" if "(f 1) / 2^k \<in> S" "f 1 \<in> \<int>" for k and f :: "real \<Rightarrow> real" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 207 | using that by (force simp: Ints_def zero_le_divide_iff power_le_zero_eq dest: pos zero_le_imp_eq_int) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 208 |   then have "S \<inter> (\<Union>k m. {real m / 2^k}) = S \<inter>
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 209 |              (\<Union>k. \<Union>f\<in>Basis \<rightarrow> \<int>. {\<Sum>i\<in>Basis. (f i / 2^k) *\<^sub>R i})"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 210 | by force | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 211 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 212 | using closure_dyadic_rationals_in_convex_set [OF \<open>convex S\<close> intnz] by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 213 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 214 | |
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 215 | |
| 70136 | 216 | definition\<^marker>\<open>tag unimportant\<close> dyadics :: "'a::field_char_0 set" where "dyadics \<equiv> \<Union>k m. {of_nat m / 2^k}"
 | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 217 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 218 | lemma real_in_dyadics [simp]: "real m \<in> dyadics" | 
| 72531 | 219 | by (simp add: dyadics_def) (metis divide_numeral_1 numeral_One power_0) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 220 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 221 | lemma nat_neq_4k1: "of_nat m \<noteq> (4 * of_nat k + 1) / (2 * 2^n :: 'a::field_char_0)" | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 222 | proof | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 223 | assume "of_nat m = (4 * of_nat k + 1) / (2 * 2^n :: 'a)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 224 | then have "of_nat (m * (2 * 2^n)) = (of_nat (Suc (4 * k)) :: 'a)" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70136diff
changeset | 225 | by (simp add: field_split_simps) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 226 | then have "m * (2 * 2^n) = Suc (4 * k)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 227 | using of_nat_eq_iff by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 228 | then have "odd (m * (2 * 2^n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 229 | by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 230 | then show False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 231 | by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 232 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 233 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 234 | lemma nat_neq_4k3: "of_nat m \<noteq> (4 * of_nat k + 3) / (2 * 2^n :: 'a::field_char_0)" | 
| 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 235 | proof | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 236 | assume "of_nat m = (4 * of_nat k + 3) / (2 * 2^n :: 'a)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 237 | then have "of_nat (m * (2 * 2^n)) = (of_nat (4 * k + 3) :: 'a)" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70136diff
changeset | 238 | by (simp add: field_split_simps) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 239 | then have "m * (2 * 2^n) = (4 * k) + 3" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 240 | using of_nat_eq_iff by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 241 | then have "odd (m * (2 * 2^n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 242 | by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 243 | then show False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 244 | by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 245 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 246 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 247 | lemma iff_4k: | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 248 | assumes "r = real k" "odd k" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 249 | shows "(4 * real m + r) / (2 * 2^n) = (4 * real m' + r) / (2 * 2 ^ n') \<longleftrightarrow> m=m' \<and> n=n'" | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 250 | proof - | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 251 |   { assume "(4 * real m + r) / (2 * 2^n) = (4 * real m' + r) / (2 * 2 ^ n')"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 252 | then have "real ((4 * m + k) * (2 * 2 ^ n')) = real ((4 * m' + k) * (2 * 2^n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 253 | using assms by (auto simp: field_simps) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 254 | then have "(4 * m + k) * (2 * 2 ^ n') = (4 * m' + k) * (2 * 2^n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 255 | using of_nat_eq_iff by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 256 | then have "(4 * m + k) * (2 ^ n') = (4 * m' + k) * (2^n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 257 | by linarith | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 258 | then obtain "4*m + k = 4*m' + k" "n=n'" | 
| 72531 | 259 | using prime_power_cancel2 [OF two_is_prime_nat] assms | 
| 260 | by (metis even_mult_iff even_numeral odd_add) | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 261 | then have "m=m'" "n=n'" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 262 | by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 263 | } | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 264 | then show ?thesis by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 265 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 266 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 267 | lemma neq_4k1_k43: "(4 * real m + 1) / (2 * 2^n) \<noteq> (4 * real m' + 3) / (2 * 2 ^ n')" | 
| 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 268 | proof | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 269 | assume "(4 * real m + 1) / (2 * 2^n) = (4 * real m' + 3) / (2 * 2 ^ n')" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 270 | then have "real (Suc (4 * m) * (2 * 2 ^ n')) = real ((4 * m' + 3) * (2 * 2^n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 271 | by (auto simp: field_simps) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 272 | then have "Suc (4 * m) * (2 * 2 ^ n') = (4 * m' + 3) * (2 * 2^n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 273 | using of_nat_eq_iff by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 274 | then have "Suc (4 * m) * (2 ^ n') = (4 * m' + 3) * (2^n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 275 | by linarith | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 276 | then have "Suc (4 * m) = (4 * m' + 3)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 277 | by (rule prime_power_cancel2 [OF two_is_prime_nat]) auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 278 | then have "1 + 2 * m' = 2 * m" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 279 | using \<open>Suc (4 * m) = 4 * m' + 3\<close> by linarith | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 280 | then show False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 281 | using even_Suc by presburger | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 282 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 283 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 284 | lemma dyadic_413_cases: | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 285 | obtains "(of_nat m::'a::field_char_0) / 2^k \<in> Nats" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 286 | | m' k' where "k' < k" "(of_nat m:: 'a) / 2^k = of_nat (4*m' + 1) / 2^Suc k'" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 287 | | m' k' where "k' < k" "(of_nat m:: 'a) / 2^k = of_nat (4*m' + 3) / 2^Suc k'" | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 288 | proof (cases "m>0") | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 289 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 290 | then have "m=0" by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 291 | with that show ?thesis by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 292 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 293 | case True | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 294 | obtain k' m' where m': "odd m'" and k': "m = m' * 2^k'" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 295 | using prime_power_canonical [OF two_is_prime_nat True] by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 296 | then obtain q r where q: "m' = 4*q + r" and r: "r < 4" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 297 | by (metis not_add_less2 split_div zero_neq_numeral) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 298 | show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 299 | proof (cases "k \<le> k'") | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 300 | case True | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 301 | have "(of_nat m:: 'a) / 2^k = of_nat m' * (2 ^ k' / 2^k)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 302 | using k' by (simp add: field_simps) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 303 | also have "... = (of_nat m'::'a) * 2 ^ (k'-k)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 304 | using k' True by (simp add: power_diff) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 305 | also have "... \<in> \<nat>" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 306 | by (metis Nats_mult of_nat_in_Nats of_nat_numeral of_nat_power) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 307 | finally show ?thesis by (auto simp: that) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 308 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 309 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 310 | then obtain kd where kd: "Suc kd = k - k'" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 311 | using Suc_diff_Suc not_less by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 312 | have "(of_nat m:: 'a) / 2^k = of_nat m' * (2 ^ k' / 2^k)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 313 | using k' by (simp add: field_simps) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 314 | also have "... = (of_nat m'::'a) / 2 ^ (k-k')" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 315 | using k' False by (simp add: power_diff) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 316 | also have "... = ((of_nat r + 4 * of_nat q)::'a) / 2 ^ (k-k')" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 317 | using q by force | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 318 | finally have meq: "(of_nat m:: 'a) / 2^k = (of_nat r + 4 * of_nat q) / 2 ^ (k - k')" . | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 319 | have "r \<noteq> 0" "r \<noteq> 2" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 320 | using q m' by presburger+ | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 321 | with r consider "r = 1" | "r = 3" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 322 | by linarith | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 323 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 324 | proof cases | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 325 | assume "r = 1" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 326 | with meq kd that(2) [of kd q] show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 327 | by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 328 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 329 | assume "r = 3" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 330 | with meq kd that(3) [of kd q] show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 331 | by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 332 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 333 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 334 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 335 | |
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 336 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 337 | lemma dyadics_iff: | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 338 | "(dyadics :: 'a::field_char_0 set) = | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 339 |     Nats \<union> (\<Union>k m. {of_nat (4*m + 1) / 2^Suc k}) \<union> (\<Union>k m. {of_nat (4*m + 3) / 2^Suc k})"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 340 | (is "_ = ?rhs") | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 341 | proof | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 342 | show "dyadics \<subseteq> ?rhs" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 343 | unfolding dyadics_def | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 344 | apply clarify | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 345 | apply (rule dyadic_413_cases, force+) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 346 | done | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 347 | next | 
| 72531 | 348 |   have "range of_nat \<subseteq> (\<Union>k m. {(of_nat m::'a) / 2 ^ k})"
 | 
| 349 | by clarsimp (metis divide_numeral_1 numeral_One power_0) | |
| 350 | moreover have "\<And>k m. \<exists>k' m'. ((1::'a) + 4 * of_nat m) / 2 ^ Suc k = of_nat m' / 2 ^ k'" | |
| 351 | by (metis (no_types) of_nat_Suc of_nat_mult of_nat_numeral) | |
| 352 | moreover have "\<And>k m. \<exists>k' m'. (4 * of_nat m + (3::'a)) / 2 ^ Suc k = of_nat m' / 2 ^ k'" | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 353 | by (metis of_nat_add of_nat_mult of_nat_numeral) | 
| 72531 | 354 | ultimately show "?rhs \<subseteq> dyadics" | 
| 355 | by (auto simp: dyadics_def Nats_def) | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 356 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 357 | |
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 358 | |
| 70136 | 359 | function\<^marker>\<open>tag unimportant\<close> (domintros) dyad_rec :: "[nat \<Rightarrow> 'a, 'a\<Rightarrow>'a, 'a\<Rightarrow>'a, real] \<Rightarrow> 'a" where | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 360 | "dyad_rec b l r (real m) = b m" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 361 | | "dyad_rec b l r ((4 * real m + 1) / 2 ^ (Suc n)) = l (dyad_rec b l r ((2*m + 1) / 2^n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 362 | | "dyad_rec b l r ((4 * real m + 3) / 2 ^ (Suc n)) = r (dyad_rec b l r ((2*m + 1) / 2^n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 363 | | "x \<notin> dyadics \<Longrightarrow> dyad_rec b l r x = undefined" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 364 | using iff_4k [of _ 1] iff_4k [of _ 3] | 
| 72531 | 365 | apply (simp_all add: nat_neq_4k1 nat_neq_4k3 neq_4k1_k43 dyadics_iff Nats_def) | 
| 366 | by (fastforce simp: field_simps)+ | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 367 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 368 | lemma dyadics_levels: "dyadics = (\<Union>K. \<Union>k<K. \<Union> m. {of_nat m / 2^k})"
 | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 369 | unfolding dyadics_def by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 370 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 371 | lemma dyad_rec_level_termination: | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 372 | assumes "k < K" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 373 | shows "dyad_rec_dom(b, l, r, real m / 2^k)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 374 | using assms | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 375 | proof (induction K arbitrary: k m) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 376 | case 0 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 377 | then show ?case by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 378 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 379 | case (Suc K) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 380 | then consider "k = K" | "k < K" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 381 | using less_antisym by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 382 | then show ?case | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 383 | proof cases | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 384 | assume "k = K" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 385 | show ?case | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 386 | proof (rule dyadic_413_cases [of m k, where 'a=real]) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 387 | show "real m / 2^k \<in> \<nat> \<Longrightarrow> dyad_rec_dom (b, l, r, real m / 2^k)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 388 | by (force simp: Nats_def nat_neq_4k1 nat_neq_4k3 intro: dyad_rec.domintros) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 389 | show ?case if "k' < k" and eq: "real m / 2^k = real (4 * m' + 1) / 2^Suc k'" for m' k' | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 390 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 391 | have "dyad_rec_dom (b, l, r, (4 * real m' + 1) / 2^Suc k')" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 392 | proof (rule dyad_rec.domintros) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 393 | fix m n | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 394 | assume "(4 * real m' + 1) / (2 * 2 ^ k') = (4 * real m + 1) / (2 * 2^n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 395 | then have "m' = m" "k' = n" using iff_4k [of _ 1] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 396 | by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 397 | have "dyad_rec_dom (b, l, r, real (2 * m + 1) / 2 ^ k')" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 398 | using Suc.IH \<open>k = K\<close> \<open>k' < k\<close> by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 399 | then show "dyad_rec_dom (b, l, r, (2 * real m + 1) / 2^n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 400 | using \<open>k' = n\<close> by (auto simp: algebra_simps) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 401 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 402 | fix m n | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 403 | assume "(4 * real m' + 1) / (2 * 2 ^ k') = (4 * real m + 3) / (2 * 2^n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 404 | then have "False" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 405 | by (metis neq_4k1_k43) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 406 | then show "dyad_rec_dom (b, l, r, (2 * real m + 1) / 2^n)" .. | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 407 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 408 | then show ?case by (simp add: eq add_ac) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 409 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 410 | show ?case if "k' < k" and eq: "real m / 2^k = real (4 * m' + 3) / 2^Suc k'" for m' k' | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 411 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 412 | have "dyad_rec_dom (b, l, r, (4 * real m' + 3) / 2^Suc k')" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 413 | proof (rule dyad_rec.domintros) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 414 | fix m n | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 415 | assume "(4 * real m' + 3) / (2 * 2 ^ k') = (4 * real m + 1) / (2 * 2^n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 416 | then have "False" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 417 | by (metis neq_4k1_k43) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 418 | then show "dyad_rec_dom (b, l, r, (2 * real m + 1) / 2^n)" .. | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 419 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 420 | fix m n | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 421 | assume "(4 * real m' + 3) / (2 * 2 ^ k') = (4 * real m + 3) / (2 * 2^n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 422 | then have "m' = m" "k' = n" using iff_4k [of _ 3] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 423 | by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 424 | have "dyad_rec_dom (b, l, r, real (2 * m + 1) / 2 ^ k')" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 425 | using Suc.IH \<open>k = K\<close> \<open>k' < k\<close> by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 426 | then show "dyad_rec_dom (b, l, r, (2 * real m + 1) / 2^n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 427 | using \<open>k' = n\<close> by (auto simp: algebra_simps) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 428 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 429 | then show ?case by (simp add: eq add_ac) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 430 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 431 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 432 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 433 | assume "k < K" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 434 | then show ?case | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 435 | using Suc.IH by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 436 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 437 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 438 | |
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 439 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 440 | lemma dyad_rec_termination: "x \<in> dyadics \<Longrightarrow> dyad_rec_dom(b,l,r,x)" | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 441 | by (auto simp: dyadics_levels intro: dyad_rec_level_termination) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 442 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 443 | lemma dyad_rec_of_nat [simp]: "dyad_rec b l r (real m) = b m" | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 444 | by (simp add: dyad_rec.psimps dyad_rec_termination) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 445 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 446 | lemma dyad_rec_41 [simp]: "dyad_rec b l r ((4 * real m + 1) / 2 ^ (Suc n)) = l (dyad_rec b l r ((2*m + 1) / 2^n))" | 
| 72531 | 447 | proof (rule dyad_rec.psimps) | 
| 448 | show "dyad_rec_dom (b, l, r, (4 * real m + 1) / 2 ^ Suc n)" | |
| 449 | by (metis add.commute dyad_rec_level_termination lessI of_nat_Suc of_nat_mult of_nat_numeral) | |
| 450 | qed | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 451 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 452 | lemma dyad_rec_43 [simp]: "dyad_rec b l r ((4 * real m + 3) / 2 ^ (Suc n)) = r (dyad_rec b l r ((2*m + 1) / 2^n))" | 
| 72531 | 453 | proof (rule dyad_rec.psimps) | 
| 454 | show "dyad_rec_dom (b, l, r, (4 * real m + 3) / 2 ^ Suc n)" | |
| 455 | by (metis dyad_rec_level_termination lessI of_nat_add of_nat_mult of_nat_numeral) | |
| 456 | qed | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 457 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 458 | lemma dyad_rec_41_times2: | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 459 | assumes "n > 0" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 460 | shows "dyad_rec b l r (2 * ((4 * real m + 1) / 2^Suc n)) = l (dyad_rec b l r (2 * (2 * real m + 1) / 2^n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 461 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 462 | obtain n' where n': "n = Suc n'" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 463 | using assms not0_implies_Suc by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 464 | have "dyad_rec b l r (2 * ((4 * real m + 1) / 2^Suc n)) = dyad_rec b l r ((2 * (4 * real m + 1)) / (2 * 2^n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 465 | by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 466 | also have "... = dyad_rec b l r ((4 * real m + 1) / 2^n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 467 | by (subst mult_divide_mult_cancel_left) auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 468 | also have "... = l (dyad_rec b l r ((2 * real m + 1) / 2 ^ n'))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 469 | by (simp add: add.commute [of 1] n' del: power_Suc) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 470 | also have "... = l (dyad_rec b l r ((2 * (2 * real m + 1)) / (2 * 2 ^ n')))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 471 | by (subst mult_divide_mult_cancel_left) auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 472 | also have "... = l (dyad_rec b l r (2 * (2 * real m + 1) / 2^n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 473 | by (simp add: add.commute n') | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 474 | finally show ?thesis . | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 475 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 476 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 477 | lemma dyad_rec_43_times2: | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 478 | assumes "n > 0" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 479 | shows "dyad_rec b l r (2 * ((4 * real m + 3) / 2^Suc n)) = r (dyad_rec b l r (2 * (2 * real m + 1) / 2^n))" | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 480 | proof - | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 481 | obtain n' where n': "n = Suc n'" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 482 | using assms not0_implies_Suc by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 483 | have "dyad_rec b l r (2 * ((4 * real m + 3) / 2^Suc n)) = dyad_rec b l r ((2 * (4 * real m + 3)) / (2 * 2^n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 484 | by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 485 | also have "... = dyad_rec b l r ((4 * real m + 3) / 2^n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 486 | by (subst mult_divide_mult_cancel_left) auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 487 | also have "... = r (dyad_rec b l r ((2 * real m + 1) / 2 ^ n'))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 488 | by (simp add: n' del: power_Suc) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 489 | also have "... = r (dyad_rec b l r ((2 * (2 * real m + 1)) / (2 * 2 ^ n')))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 490 | by (subst mult_divide_mult_cancel_left) auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 491 | also have "... = r (dyad_rec b l r (2 * (2 * real m + 1) / 2^n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 492 | by (simp add: n') | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 493 | finally show ?thesis . | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 494 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 495 | |
| 70136 | 496 | definition\<^marker>\<open>tag unimportant\<close> dyad_rec2 | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 497 | where "dyad_rec2 u v lc rc x = | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 498 | dyad_rec (\<lambda>z. (u,v)) (\<lambda>(a,b). (a, lc a b (midpoint a b))) (\<lambda>(a,b). (rc a b (midpoint a b), b)) (2*x)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 499 | |
| 70136 | 500 | abbreviation\<^marker>\<open>tag unimportant\<close> leftrec where "leftrec u v lc rc x \<equiv> fst (dyad_rec2 u v lc rc x)" | 
| 501 | abbreviation\<^marker>\<open>tag unimportant\<close> rightrec where "rightrec u v lc rc x \<equiv> snd (dyad_rec2 u v lc rc x)" | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 502 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 503 | lemma leftrec_base: "leftrec u v lc rc (real m / 2) = u" | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 504 | by (simp add: dyad_rec2_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 505 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 506 | lemma leftrec_41: "n > 0 \<Longrightarrow> leftrec u v lc rc ((4 * real m + 1) / 2 ^ (Suc n)) = leftrec u v lc rc ((2 * real m + 1) / 2^n)" | 
| 72531 | 507 | unfolding dyad_rec2_def dyad_rec_41_times2 | 
| 508 | by (simp add: case_prod_beta) | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 509 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 510 | lemma leftrec_43: "n > 0 \<Longrightarrow> | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 511 | leftrec u v lc rc ((4 * real m + 3) / 2 ^ (Suc n)) = | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 512 | rc (leftrec u v lc rc ((2 * real m + 1) / 2^n)) (rightrec u v lc rc ((2 * real m + 1) / 2^n)) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 513 | (midpoint (leftrec u v lc rc ((2 * real m + 1) / 2^n)) (rightrec u v lc rc ((2 * real m + 1) / 2^n)))" | 
| 72531 | 514 | unfolding dyad_rec2_def dyad_rec_43_times2 | 
| 515 | by (simp add: case_prod_beta) | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 516 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 517 | lemma rightrec_base: "rightrec u v lc rc (real m / 2) = v" | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 518 | by (simp add: dyad_rec2_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 519 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 520 | lemma rightrec_41: "n > 0 \<Longrightarrow> | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 521 | rightrec u v lc rc ((4 * real m + 1) / 2 ^ (Suc n)) = | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 522 | lc (leftrec u v lc rc ((2 * real m + 1) / 2^n)) (rightrec u v lc rc ((2 * real m + 1) / 2^n)) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 523 | (midpoint (leftrec u v lc rc ((2 * real m + 1) / 2^n)) (rightrec u v lc rc ((2 * real m + 1) / 2^n)))" | 
| 72531 | 524 | unfolding dyad_rec2_def dyad_rec_41_times2 | 
| 525 | by (simp add: case_prod_beta) | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 526 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 527 | lemma rightrec_43: "n > 0 \<Longrightarrow> rightrec u v lc rc ((4 * real m + 3) / 2 ^ (Suc n)) = rightrec u v lc rc ((2 * real m + 1) / 2^n)" | 
| 72531 | 528 | unfolding dyad_rec2_def dyad_rec_43_times2 | 
| 529 | by (simp add: case_prod_beta) | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 530 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 531 | lemma dyadics_in_open_unit_interval: | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 532 |    "{0<..<1} \<inter> (\<Union>k m. {real m / 2^k}) = (\<Union>k. \<Union>m \<in> {0<..<2^k}. {real m / 2^k})"
 | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70136diff
changeset | 533 | by (auto simp: field_split_simps) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 534 | |
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 535 | |
| 77943 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 536 | lemma padic_rational_approximation_straddle: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 537 | assumes "\<epsilon> > 0" "p > 1" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 538 | obtains n q r | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 539 | where "of_int q / p^n < x" "x < of_int r / p^n" "\<bar>q / p^n - r / p^n \<bar> < \<epsilon>" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 540 | proof - | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 541 | obtain n where n: "2 / \<epsilon> < p ^ n" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 542 | using \<open>p>1\<close> real_arch_pow by blast | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 543 | define q where "q \<equiv> \<lfloor>p ^ n * x\<rfloor> - 1" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 544 | show thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 545 | proof | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 546 | show "q / p ^ n < x" "x < real_of_int (q+2) / p ^ n" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 547 | using assms by (simp_all add: q_def divide_simps floor_less_cancel mult.commute) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 548 | show "\<bar>q / p ^ n - real_of_int (q+2) / p ^ n\<bar> < \<epsilon>" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 549 | using assms n by (simp add: q_def divide_simps mult.commute) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 550 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 551 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 552 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 553 | lemma padic_rational_approximation_straddle_pos: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 554 | assumes "\<epsilon> > 0" "p > 1" "x > 0" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 555 | obtains n q r | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 556 | where "of_nat q / p^n < x" "x < of_nat r / p^n" "\<bar>q / p^n - r / p^n \<bar> < \<epsilon>" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 557 | proof - | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 558 | obtain n q r | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 559 | where *: "of_int q / p^n < x" "x < of_int r / p^n" "\<bar>q / p^n - r / p^n \<bar> < \<epsilon>" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 560 | using padic_rational_approximation_straddle assms by metis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 561 | then have "r \<ge> 0" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 562 | using assms by (smt (verit, best) divide_nonpos_pos of_int_0_le_iff zero_less_power) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 563 | show thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 564 | proof | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 565 | show "real (max 0 (nat q)) / p ^ n < x" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 566 | using * by (metis assms(3) div_0 max_nat.left_neutral nat_eq_iff2 of_nat_0 of_nat_nat) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 567 | show "x < real (nat r) / p ^ n" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 568 | using \<open>r \<ge> 0\<close> * by force | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 569 | show "\<bar>real (max 0 (nat q)) / p ^ n - real (nat r) / p ^ n\<bar> < \<epsilon>" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 570 | using * assms by (simp add: divide_simps) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 571 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 572 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 573 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 574 | lemma padic_rational_approximation_straddle_pos_le: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 575 | assumes "\<epsilon> > 0" "p > 1" "x \<ge> 0" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 576 | obtains n q r | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 577 | where "of_nat q / p^n \<le> x" "x < of_nat r / p^n" "\<bar>q / p^n - r / p^n \<bar> < \<epsilon>" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 578 | proof - | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 579 | obtain n q r | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 580 | where *: "of_int q / p^n < x" "x < of_int r / p^n" "\<bar>q / p^n - r / p^n \<bar> < \<epsilon>" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 581 | using padic_rational_approximation_straddle assms by metis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 582 | then have "r \<ge> 0" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 583 | using assms by (smt (verit, best) divide_nonpos_pos of_int_0_le_iff zero_less_power) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 584 | show thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 585 | proof | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 586 | show "real (max 0 (nat q)) / p ^ n \<le> x" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 587 | using * assms(3) nle_le by fastforce | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 588 | show "x < real (nat r) / p ^ n" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 589 | using \<open>r \<ge> 0\<close> * by force | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 590 | show "\<bar>real (max 0 (nat q)) / p ^ n - real (nat r) / p ^ n\<bar> < \<epsilon>" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 591 | using * assms by (simp add: divide_simps) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 592 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 593 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 594 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 595 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 596 | subsubsection \<open>Definition by recursion on dyadic rationals in [0,1]\<close> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 597 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 598 | lemma recursion_on_dyadic_fractions: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 599 | assumes base: "R a b" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 600 | and step: "\<And>x y. R x y \<Longrightarrow> \<exists>z. R x z \<and> R z y" and trans: "\<And>x y z. \<lbrakk>R x y; R y z\<rbrakk> \<Longrightarrow> R x z" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 601 | shows "\<exists>f :: real \<Rightarrow> 'a. f 0 = a \<and> f 1 = b \<and> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 602 |                (\<forall>x \<in> dyadics \<inter> {0..1}. \<forall>y \<in> dyadics \<inter> {0..1}. x < y \<longrightarrow> R (f x) (f y))"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 603 | proof - | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 604 | obtain mid where mid: "R x y \<Longrightarrow> R x (mid x y)" "R x y \<Longrightarrow> R (mid x y) y" for x y | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 605 | using step by metis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 606 | define g where "g \<equiv> rec_nat (\<lambda>k. if k = 0 then a else b) (\<lambda>n r k. if even k then r (k div 2) else mid (r ((k - 1) div 2)) (r ((Suc k) div 2)))" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 607 | have g0 [simp]: "g 0 = (\<lambda>k. if k = 0 then a else b)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 608 | by (simp add: g_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 609 | have gSuc [simp]: "\<And>n. g(Suc n) = (\<lambda>k. if even k then g n (k div 2) else mid (g n ((k - 1) div 2)) (g n ((Suc k) div 2)))" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 610 | by (auto simp: g_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 611 | have g_eq_g: "2 ^ d * k = k' \<Longrightarrow> g n k = g (n + d) k'" for n d k k' | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 612 | by (induction d arbitrary: k k') auto | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 613 | have "g n k = g n' k'" if "real k / 2^n = real k' / 2^n'" "n' \<le> n" for k n k' n' | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 614 | proof - | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 615 | have "real k = real k' * 2 ^ (n-n')" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 616 | using that by (simp add: power_diff divide_simps) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 617 | then have "k = k' * 2 ^ (n-n')" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 618 | using of_nat_eq_iff by fastforce | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 619 | with g_eq_g show ?thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 620 | by (metis le_add_diff_inverse mult.commute that(2)) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 621 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 622 | then have g_eq_g: "g n k = g n' k'" if "real k / 2 ^ n = real k' / 2 ^ n'" for k n k' n' | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 623 | by (metis nat_le_linear that) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 624 | then obtain f where "(\<lambda>(k,n). g n k) = f \<circ> (\<lambda>(k,n). k / 2 ^ n)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 625 | using function_factors_left by (smt (verit, del_insts) case_prod_beta') | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 626 | then have f_eq_g: "\<And>k n. f(real k / 2 ^ n) = g n k" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 627 | by (simp add: fun_eq_iff) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 628 | show ?thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 629 | proof (intro exI conjI strip) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 630 | show "f 0 = a" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 631 | by (metis f_eq_g g0 div_0 of_nat_0) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 632 | show "f 1 = b" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 633 | by (metis f_eq_g g0 div_by_1 of_nat_1_eq_iff power_0 zero_neq_one) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 634 | show "R (f x) (f y)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 635 |       if x: "x \<in> dyadics \<inter> {0..1}" and y: "y \<in> dyadics \<inter> {0..1}" and "x < y" for x y
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 636 | proof - | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 637 | obtain n1 k1 where xeq: "x = real k1 / 2^n1" "k1 \<le> 2^n1" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 638 | using x by (auto simp: dyadics_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 639 | obtain n2 k2 where yeq: "y = real k2 / 2^n2" "k2 \<le> 2^n2" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 640 | using y by (auto simp: dyadics_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 641 | have xcommon: "x = real(2^n2 * k1) / 2 ^ (n1+n2)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 642 | using xeq by (simp add: power_add) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 643 | have ycommon: "y = real(2^n1 * k2) / 2 ^ (n1+n2)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 644 | using yeq by (simp add: power_add) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 645 | have *: "R (g n j) (g n k)" if "j < k" "k \<le> 2^n" for n j k | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 646 | using that | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 647 | proof (induction n arbitrary: j k) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 648 | case 0 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 649 | then show ?case | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 650 | by (simp add: base) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 651 | next | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 652 | case (Suc n) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 653 | show ?case | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 654 | proof (cases "even j") | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 655 | case True | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 656 | then obtain a where [simp]: "j = 2*a" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 657 | by blast | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 658 | show ?thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 659 | proof (cases "even k") | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 660 | case True | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 661 | with Suc show ?thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 662 | by (auto elim!: evenE) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 663 | next | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 664 | case False | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 665 | then obtain b where [simp]: "k = Suc (2*b)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 666 | using oddE by fastforce | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 667 | show ?thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 668 | using Suc | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 669 | apply simp | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 670 | by (smt (verit, ccfv_SIG) less_Suc_eq linorder_not_le local.trans mid(1) nat_mult_less_cancel1 pos2) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 671 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 672 | next | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 673 | case False | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 674 | then obtain a where [simp]: "j = Suc (2*a)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 675 | using oddE by fastforce | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 676 | show ?thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 677 | proof (cases "even k") | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 678 | case True | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 679 | then obtain b where [simp]: "k = 2*b" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 680 | by blast | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 681 | show ?thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 682 | using Suc | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 683 | apply simp | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 684 | by (smt (verit, ccfv_SIG) Suc_leI Suc_lessD le_trans lessI linorder_neqE_nat linorder_not_le local.trans mid(2) nat_mult_less_cancel1 pos2) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 685 | next | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 686 | case False | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 687 | then obtain b where [simp]: "k = Suc (2*b)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 688 | using oddE by fastforce | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 689 | show ?thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 690 | using Suc | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 691 | apply simp | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 692 | by (smt (verit) Suc_leI le_trans lessI less_or_eq_imp_le linorder_neqE_nat linorder_not_le local.trans mid(1) mid(2) nat_mult_less_cancel1 pos2) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 693 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 694 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 695 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 696 | show ?thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 697 | unfolding xcommon ycommon f_eq_g | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 698 | proof (rule *) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 699 | show "2 ^ n2 * k1 < 2 ^ n1 * k2" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 700 | using of_nat_less_iff \<open>x < y\<close> by (fastforce simp: xeq yeq field_simps) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 701 | show "2 ^ n1 * k2 \<le> 2 ^ (n1 + n2)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 702 | by (simp add: power_add yeq) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 703 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 704 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 705 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 706 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 707 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 708 | lemma dyadics_add: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 709 | assumes "x \<in> dyadics" "y \<in> dyadics" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 710 | shows "x+y \<in> dyadics" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 711 | proof - | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 712 | obtain i j m n where x: "x = of_nat i / 2 ^ m" and y: "y = of_nat j / 2 ^ n" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 713 | using assms by (auto simp: dyadics_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 714 | have xcommon: "x = of_nat(2^n * i) / 2 ^ (m+n)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 715 | using x by (simp add: power_add) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 716 | moreover | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 717 | have ycommon: "y = of_nat(2^m * j) / 2 ^ (m+n)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 718 | using y by (simp add: power_add) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 719 | ultimately have "x+y = (of_nat(2^n * i + 2^m * j)) / 2 ^ (m+n)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 720 | by (simp add: field_simps) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 721 | then show ?thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 722 | unfolding dyadics_def by blast | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 723 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 724 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 725 | lemma dyadics_diff: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 726 | fixes x :: "'a::linordered_field" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 727 | assumes "x \<in> dyadics" "y \<in> dyadics" "y \<le> x" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 728 | shows "x-y \<in> dyadics" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 729 | proof - | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 730 | obtain i j m n where x: "x = of_nat i / 2 ^ m" and y: "y = of_nat j / 2 ^ n" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 731 | using assms by (auto simp: dyadics_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 732 | have j_le_i: "j * 2 ^ m \<le> i * 2 ^ n" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 733 | using of_nat_le_iff \<open>y \<le> x\<close> unfolding x y by (fastforce simp add: divide_simps) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 734 | have xcommon: "x = of_nat(2^n * i) / 2 ^ (m+n)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 735 | using x by (simp add: power_add) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 736 | moreover | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 737 | have ycommon: "y = of_nat(2^m * j) / 2 ^ (m+n)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 738 | using y by (simp add: power_add) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 739 | ultimately have "x-y = (of_nat(2^n * i - 2^m * j)) / 2 ^ (m+n)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 740 | by (simp add: xcommon ycommon field_simps j_le_i of_nat_diff) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 741 | then show ?thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 742 | unfolding dyadics_def by blast | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 743 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 744 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 745 | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 746 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 747 | theorem homeomorphic_monotone_image_interval: | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 748 |   fixes f :: "real \<Rightarrow> 'a::{real_normed_vector,complete_space}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 749 |   assumes cont_f: "continuous_on {0..1} f"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 750 |       and conn: "\<And>y. connected ({0..1} \<inter> f -` {y})"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 751 | and f_1not0: "f 1 \<noteq> f 0" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 752 |     shows "(f ` {0..1}) homeomorphic {0..1::real}"
 | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 753 | proof - | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 754 | have "\<exists>c d. a \<le> c \<and> c \<le> m \<and> m \<le> d \<and> d \<le> b \<and> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 755 |               (\<forall>x \<in> {c..d}. f x = f m) \<and>
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 756 |               (\<forall>x \<in> {a..<c}. (f x \<noteq> f m)) \<and>
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 757 |               (\<forall>x \<in> {d<..b}. (f x \<noteq> f m)) \<and>
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 758 |               (\<forall>x \<in> {a..<c}. \<forall>y \<in> {d<..b}. f x \<noteq> f y)"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 759 |     if m: "m \<in> {a..b}" and ab01: "{a..b} \<subseteq> {0..1}" for a b m
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 760 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 761 |     have comp: "compact (f -` {f m} \<inter> {0..1})"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 762 | by (simp add: compact_eq_bounded_closed bounded_Int closed_vimage_Int cont_f) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 763 |     obtain c0 d0 where cd0: "{0..1} \<inter> f -` {f m} = {c0..d0}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 764 |       using connected_compact_interval_1 [of "{0..1} \<inter> f -` {f m}"] conn comp
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 765 | by (metis Int_commute) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 766 | with that have "m \<in> cbox c0 d0" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 767 | by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 768 |     obtain c d where cd: "{a..b} \<inter> f -` {f m} = {c..d}"
 | 
| 72531 | 769 | using ab01 cd0 | 
| 770 | by (rule_tac c="max a c0" and d="min b d0" in that) auto | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 771 |     then have cdab: "{c..d} \<subseteq> {a..b}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 772 | by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 773 | show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 774 | proof (intro exI conjI ballI) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 775 | show "a \<le> c" "d \<le> b" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 776 | using cdab cd m by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 777 | show "c \<le> m" "m \<le> d" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 778 | using cd m by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 779 |       show "\<And>x. x \<in> {c..d} \<Longrightarrow> f x = f m"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 780 | using cd by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 781 |       show "f x \<noteq> f m" if "x \<in> {a..<c}" for x
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 782 | using that m cd [THEN equalityD1, THEN subsetD] \<open>c \<le> m\<close> by force | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 783 |       show "f x \<noteq> f m" if "x \<in> {d<..b}" for x
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 784 | using that m cd [THEN equalityD1, THEN subsetD, of x] \<open>m \<le> d\<close> by force | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 785 |       show "f x \<noteq> f y" if "x \<in> {a..<c}" "y \<in> {d<..b}" for x y
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 786 | proof (cases "f x = f m \<or> f y = f m") | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 787 | case True | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 788 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 789 |           using \<open>\<And>x. x \<in> {a..<c} \<Longrightarrow> f x \<noteq> f m\<close> that by auto
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 790 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 791 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 792 | have False if "f x = f y" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 793 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 794 | have "x \<le> m" "m \<le> y" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 795 |             using \<open>c \<le> m\<close> \<open>x \<in> {a..<c}\<close>  \<open>m \<le> d\<close> \<open>y \<in> {d<..b}\<close> by auto
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 796 |           then have "x \<in> ({0..1} \<inter> f -` {f y})" "y \<in> ({0..1} \<inter> f -` {f y})"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 797 |             using \<open>x \<in> {a..<c}\<close> \<open>y \<in> {d<..b}\<close> ab01 by (auto simp: that)
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 798 |           then have "m \<in> ({0..1} \<inter> f -` {f y})"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 799 | by (meson \<open>m \<le> y\<close> \<open>x \<le> m\<close> is_interval_connected_1 conn [of "f y"] is_interval_1) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 800 | with False show False by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 801 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 802 | then show ?thesis by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 803 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 804 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 805 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 806 | then obtain leftcut rightcut where LR: | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 807 |     "\<And>a b m. \<lbrakk>m \<in> {a..b}; {a..b} \<subseteq> {0..1}\<rbrakk> \<Longrightarrow>
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 808 | (a \<le> leftcut a b m \<and> leftcut a b m \<le> m \<and> m \<le> rightcut a b m \<and> rightcut a b m \<le> b \<and> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 809 |             (\<forall>x \<in> {leftcut a b m..rightcut a b m}. f x = f m) \<and>
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 810 |             (\<forall>x \<in> {a..<leftcut a b m}. f x \<noteq> f m) \<and>
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 811 |             (\<forall>x \<in> {rightcut a b m<..b}. f x \<noteq> f m) \<and>
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 812 |             (\<forall>x \<in> {a..<leftcut a b m}. \<forall>y \<in> {rightcut a b m<..b}. f x \<noteq> f y))"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 813 | apply atomize | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 814 | apply (clarsimp simp only: imp_conjL [symmetric] choice_iff choice_iff') | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 815 | apply (rule that, blast) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 816 | done | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 817 |   then have left_right: "\<And>a b m. \<lbrakk>m \<in> {a..b}; {a..b} \<subseteq> {0..1}\<rbrakk> \<Longrightarrow> a \<le> leftcut a b m \<and> rightcut a b m \<le> b"
 | 
| 72531 | 818 |        and  left_right_m: "\<And>a b m. \<lbrakk>m \<in> {a..b}; {a..b} \<subseteq> {0..1}\<rbrakk> \<Longrightarrow> leftcut a b m \<le> m \<and> m \<le> rightcut a b m"
 | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 819 | by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 820 |   have left_neq: "\<lbrakk>a \<le> x; x < leftcut a b m; a \<le> m; m \<le> b; {a..b} \<subseteq> {0..1}\<rbrakk> \<Longrightarrow> f x \<noteq> f m"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 821 |     and right_neq: "\<lbrakk>rightcut a b m < x; x \<le> b; a \<le> m; m \<le> b; {a..b} \<subseteq> {0..1}\<rbrakk> \<Longrightarrow> f x \<noteq> f m"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 822 |     and left_right_neq: "\<lbrakk>a \<le> x; x < leftcut a b m; rightcut a b m < y; y \<le> b; a \<le> m; m \<le> b; {a..b} \<subseteq> {0..1}\<rbrakk> \<Longrightarrow> f x \<noteq> f m"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 823 |     and feqm: "\<lbrakk>leftcut a b m \<le> x; x \<le> rightcut a b m; a \<le> m; m \<le> b; {a..b} \<subseteq> {0..1}\<rbrakk>
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 824 | \<Longrightarrow> f x = f m" for a b m x y | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 825 | by (meson atLeastAtMost_iff greaterThanAtMost_iff atLeastLessThan_iff LR)+ | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 826 | have f_eqI: "\<And>a b m x y. \<lbrakk>leftcut a b m \<le> x; x \<le> rightcut a b m; leftcut a b m \<le> y; y \<le> rightcut a b m; | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 827 |                              a \<le> m; m \<le> b; {a..b} \<subseteq> {0..1}\<rbrakk>  \<Longrightarrow> f x = f y"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 828 | by (metis feqm) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 829 | define u where "u \<equiv> rightcut 0 1 0" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 830 | have lc[simp]: "leftcut 0 1 0 = 0" and u01: "0 \<le> u" "u \<le> 1" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 831 | using LR [of 0 0 1] by (auto simp: u_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 832 |   have f0u: "\<And>x. x \<in> {0..u} \<Longrightarrow> f x = f 0"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 833 | using LR [of 0 0 1] unfolding u_def [symmetric] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 834 | by (metis \<open>leftcut 0 1 0 = 0\<close> atLeastAtMost_iff order_refl zero_le_one) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 835 |   have fu1: "\<And>x. x \<in> {u<..1} \<Longrightarrow> f x \<noteq> f 0"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 836 | using LR [of 0 0 1] unfolding u_def [symmetric] by fastforce | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 837 | define v where "v \<equiv> leftcut u 1 1" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 838 | have rc[simp]: "rightcut u 1 1 = 1" and v01: "u \<le> v" "v \<le> 1" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 839 | using LR [of 1 u 1] u01 by (auto simp: v_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 840 |   have fuv: "\<And>x. x \<in> {u..<v} \<Longrightarrow> f x \<noteq> f 1"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 841 | using LR [of 1 u 1] u01 v_def by fastforce | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 842 |   have f0v: "\<And>x. x \<in> {0..<v} \<Longrightarrow> f x \<noteq> f 1"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 843 | by (metis f_1not0 atLeastAtMost_iff atLeastLessThan_iff f0u fuv linear) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 844 |   have fv1: "\<And>x. x \<in> {v..1} \<Longrightarrow> f x = f 1"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 845 | using LR [of 1 u 1] u01 v_def by (metis atLeastAtMost_iff atLeastatMost_subset_iff order_refl rc) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 846 | define a where "a \<equiv> leftrec u v leftcut rightcut" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 847 | define b where "b \<equiv> rightrec u v leftcut rightcut" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 848 | define c where "c \<equiv> \<lambda>x. midpoint (a x) (b x)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 849 | have a_real [simp]: "a (real j) = u" for j | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 850 | using a_def leftrec_base | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 851 | by (metis nonzero_mult_div_cancel_right of_nat_mult of_nat_numeral zero_neq_numeral) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 852 | have b_real [simp]: "b (real j) = v" for j | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 853 | using b_def rightrec_base | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 854 | by (metis nonzero_mult_div_cancel_right of_nat_mult of_nat_numeral zero_neq_numeral) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 855 | have a41: "a ((4 * real m + 1) / 2^Suc n) = a ((2 * real m + 1) / 2^n)" if "n > 0" for m n | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 856 | using that a_def leftrec_41 by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 857 | have b41: "b ((4 * real m + 1) / 2^Suc n) = | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 858 | leftcut (a ((2 * real m + 1) / 2^n)) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 859 | (b ((2 * real m + 1) / 2^n)) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 860 | (c ((2 * real m + 1) / 2^n))" if "n > 0" for m n | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 861 | using that a_def b_def c_def rightrec_41 by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 862 | have a43: "a ((4 * real m + 3) / 2^Suc n) = | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 863 | rightcut (a ((2 * real m + 1) / 2^n)) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 864 | (b ((2 * real m + 1) / 2^n)) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 865 | (c ((2 * real m + 1) / 2^n))" if "n > 0" for m n | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 866 | using that a_def b_def c_def leftrec_43 by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 867 | have b43: "b ((4 * real m + 3) / 2^Suc n) = b ((2 * real m + 1) / 2^n)" if "n > 0" for m n | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 868 | using that b_def rightrec_43 by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 869 | have uabv: "u \<le> a (real m / 2 ^ n) \<and> a (real m / 2 ^ n) \<le> b (real m / 2 ^ n) \<and> b (real m / 2 ^ n) \<le> v" for m n | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 870 | proof (induction n arbitrary: m) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 871 | case 0 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 872 | then show ?case by (simp add: v01) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 873 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 874 | case (Suc n p) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 875 | show ?case | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 876 | proof (cases "even p") | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 877 | case True | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 878 | then obtain m where "p = 2*m" by (metis evenE) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 879 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 880 | by (simp add: Suc.IH) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 881 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 882 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 883 | then obtain m where m: "p = 2*m + 1" by (metis oddE) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 884 | show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 885 | proof (cases n) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 886 | case 0 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 887 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 888 | by (simp add: a_def b_def leftrec_base rightrec_base v01) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 889 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 890 | case (Suc n') | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 891 | then have "n > 0" by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 892 | have a_le_c: "a (real m / 2^n) \<le> c (real m / 2^n)" for m | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 893 | unfolding c_def by (metis Suc.IH ge_midpoint_1) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 894 | have c_le_b: "c (real m / 2^n) \<le> b (real m / 2^n)" for m | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 895 | unfolding c_def by (metis Suc.IH le_midpoint_1) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 896 | have c_ge_u: "c (real m / 2^n) \<ge> u" for m | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 897 | using Suc.IH a_le_c order_trans by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 898 | have c_le_v: "c (real m / 2^n) \<le> v" for m | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 899 | using Suc.IH c_le_b order_trans by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 900 | have a_ge_0: "0 \<le> a (real m / 2^n)" for m | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 901 | using Suc.IH order_trans u01(1) by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 902 | have b_le_1: "b (real m / 2^n) \<le> 1" for m | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 903 | using Suc.IH order_trans v01(2) by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 904 | have left_le: "leftcut (a ((real m) / 2^n)) (b ((real m) / 2^n)) (c ((real m) / 2^n)) \<le> c ((real m) / 2^n)" for m | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 905 | by (simp add: LR a_ge_0 a_le_c b_le_1 c_le_b) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 906 | have right_ge: "rightcut (a ((real m) / 2^n)) (b ((real m) / 2^n)) (c ((real m) / 2^n)) \<ge> c ((real m) / 2^n)" for m | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 907 | by (simp add: LR a_ge_0 a_le_c b_le_1 c_le_b) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 908 | show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 909 | proof (cases "even m") | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 910 | case True | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 911 | then obtain r where r: "m = 2*r" by (metis evenE) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 912 | show ?thesis | 
| 72531 | 913 | using order_trans [OF left_le c_le_v, of "1+2*r"] Suc.IH [of "m+1"] | 
| 914 | using a_le_c [of "m+1"] c_le_b [of "m+1"] a_ge_0 [of "m+1"] b_le_1 [of "m+1"] left_right \<open>n > 0\<close> | |
| 915 | by (simp_all add: r m add.commute [of 1] a41 b41 del: power_Suc) | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 916 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 917 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 918 | then obtain r where r: "m = 2*r + 1" by (metis oddE) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 919 | show ?thesis | 
| 72531 | 920 | using order_trans [OF c_ge_u right_ge, of "1+2*r"] Suc.IH [of "m"] | 
| 921 | using a_le_c [of "m"] c_le_b [of "m"] a_ge_0 [of "m"] b_le_1 [of "m"] left_right \<open>n > 0\<close> | |
| 922 | apply (simp_all add: r m add.commute [of 3] a43 b43 del: power_Suc) | |
| 923 | by (simp add: add.commute) | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 924 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 925 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 926 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 927 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 928 | have a_ge_0 [simp]: "0 \<le> a(m / 2^n)" and b_le_1 [simp]: "b(m / 2^n) \<le> 1" for m::nat and n | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 929 | using uabv order_trans u01 v01 by blast+ | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 930 | then have b_ge_0 [simp]: "0 \<le> b(m / 2^n)" and a_le_1 [simp]: "a(m / 2^n) \<le> 1" for m::nat and n | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 931 | using uabv order_trans by blast+ | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 932 | have alec [simp]: "a(m / 2^n) \<le> c(m / 2^n)" and cleb [simp]: "c(m / 2^n) \<le> b(m / 2^n)" for m::nat and n | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 933 | by (auto simp: c_def ge_midpoint_1 le_midpoint_1 uabv) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 934 | have c_ge_0 [simp]: "0 \<le> c(m / 2^n)" and c_le_1 [simp]: "c(m / 2^n) \<le> 1" for m::nat and n | 
| 72531 | 935 | using a_ge_0 alec b_le_1 cleb order_trans by blast+ | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 936 | have "\<lbrakk>d = m-n; odd j; \<bar>real i / 2^m - real j / 2^n\<bar> < 1/2 ^ n\<rbrakk> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 937 | \<Longrightarrow> (a(j / 2^n)) \<le> (c(i / 2^m)) \<and> (c(i / 2^m)) \<le> (b(j / 2^n))" for d i j m n | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 938 | proof (induction d arbitrary: j n rule: less_induct) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 939 | case (less d j n) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 940 | show ?case | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 941 | proof (cases "m \<le> n") | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 942 | case True | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 943 | have "\<bar>2^n\<bar> * \<bar>real i / 2^m - real j / 2^n\<bar> = 0" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 944 | proof (rule Ints_nonzero_abs_less1) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 945 | have "(real i * 2^n - real j * 2^m) / 2^m = (real i * 2^n) / 2^m - (real j * 2^m) / 2^m" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 946 | using diff_divide_distrib by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 947 | also have "... = (real i * 2 ^ (n-m)) - (real j)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 948 | using True by (auto simp: power_diff field_simps) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 949 | also have "... \<in> \<int>" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 950 | by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 951 | finally have "(real i * 2^n - real j * 2^m) / 2^m \<in> \<int>" . | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 952 | with True Ints_abs show "\<bar>2^n\<bar> * \<bar>real i / 2^m - real j / 2^n\<bar> \<in> \<int>" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70136diff
changeset | 953 | by (fastforce simp: field_split_simps) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 954 | show "\<bar>\<bar>2^n\<bar> * \<bar>real i / 2^m - real j / 2^n\<bar>\<bar> < 1" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70136diff
changeset | 955 | using less.prems by (auto simp: field_split_simps) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 956 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 957 | then have "real i / 2^m = real j / 2^n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 958 | by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 959 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 960 | by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 961 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 962 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 963 | then have "n < m" by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 964 | obtain k where k: "j = Suc (2*k)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 965 | using \<open>odd j\<close> oddE by fastforce | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 966 | show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 967 | proof (cases "n > 0") | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 968 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 969 | then have "a (real j / 2^n) = u" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 970 | by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 971 | also have "... \<le> c (real i / 2^m)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 972 | using alec uabv by (blast intro: order_trans) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 973 | finally have ac: "a (real j / 2^n) \<le> c (real i / 2^m)" . | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 974 | have "c (real i / 2^m) \<le> v" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 975 | using cleb uabv by (blast intro: order_trans) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 976 | also have "... = b (real j / 2^n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 977 | using False by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 978 | finally show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 979 | by (auto simp: ac) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 980 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 981 | case True show ?thesis | 
| 72531 | 982 | proof (cases "i / 2^m" "j / 2^n" rule: linorder_cases) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 983 | case less | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 984 | moreover have "real (4 * k + 1) / 2 ^ Suc n + 1 / (2 ^ Suc n) = real j / 2 ^ n" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70136diff
changeset | 985 | using k by (force simp: field_split_simps) | 
| 72531 | 986 | moreover have "\<bar>real i / 2 ^ m - j / 2 ^ n\<bar> < 2 / (2 ^ Suc n)" | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 987 | using less.prems by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 988 | ultimately have closer: "\<bar>real i / 2 ^ m - real (4 * k + 1) / 2 ^ Suc n\<bar> < 1 / (2 ^ Suc n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 989 | using less.prems by linarith | 
| 72531 | 990 | have "a (real (4 * k + 1) / 2 ^ Suc n) \<le> c (i / 2 ^ m) \<and> | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 991 | c (real i / 2 ^ m) \<le> b (real (4 * k + 1) / 2 ^ Suc n)" | 
| 72531 | 992 | proof (rule less.IH [OF _ refl]) | 
| 993 | show "m - Suc n < d" | |
| 994 | using \<open>n < m\<close> diff_less_mono2 less.prems(1) lessI by presburger | |
| 995 | show "\<bar>real i / 2 ^ m - real (4 * k + 1) / 2 ^ Suc n\<bar> < 1 / 2 ^ Suc n" | |
| 996 | using closer \<open>n < m\<close> \<open>d = m - n\<close> by (auto simp: field_split_simps \<open>n < m\<close> diff_less_mono2) | |
| 997 | qed auto | |
| 998 | then show ?thesis | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 999 | using LR [of "c((2*k + 1) / 2^n)" "a((2*k + 1) / 2^n)" "b((2*k + 1) / 2^n)"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1000 | using alec [of "2*k+1"] cleb [of "2*k+1"] a_ge_0 [of "2*k+1"] b_le_1 [of "2*k+1"] | 
| 72531 | 1001 | using k a41 b41 \<open>0 < n\<close> | 
| 1002 | by (simp add: add.commute) | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1003 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1004 | case equal then show ?thesis by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1005 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1006 | case greater | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1007 | moreover have "real (4 * k + 3) / 2 ^ Suc n - 1 / (2 ^ Suc n) = real j / 2 ^ n" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70136diff
changeset | 1008 | using k by (force simp: field_split_simps) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1009 | moreover have "\<bar>real i / 2 ^ m - real j / 2 ^ n\<bar> < 2 * 1 / (2 ^ Suc n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1010 | using less.prems by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1011 | ultimately have closer: "\<bar>real i / 2 ^ m - real (4 * k + 3) / 2 ^ Suc n\<bar> < 1 / (2 ^ Suc n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1012 | using less.prems by linarith | 
| 72531 | 1013 | have "a (real (4 * k + 3) / 2 ^ Suc n) \<le> c (real i / 2 ^ m) \<and> | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1014 | c (real i / 2 ^ m) \<le> b (real (4 * k + 3) / 2 ^ Suc n)" | 
| 72531 | 1015 | proof (rule less.IH [OF _ refl]) | 
| 1016 | show "m - Suc n < d" | |
| 1017 | using \<open>n < m\<close> diff_less_mono2 less.prems(1) by blast | |
| 1018 | show "\<bar>real i / 2 ^ m - real (4 * k + 3) / 2 ^ Suc n\<bar> < 1 / 2 ^ Suc n" | |
| 1019 | using closer \<open>n < m\<close> \<open>d = m - n\<close> by (auto simp: field_split_simps \<open>n < m\<close> diff_less_mono2) | |
| 1020 | qed auto | |
| 1021 | then show ?thesis | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1022 | using LR [of "c((2*k + 1) / 2^n)" "a((2*k + 1) / 2^n)" "b((2*k + 1) / 2^n)"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1023 | using alec [of "2*k+1"] cleb [of "2*k+1"] a_ge_0 [of "2*k+1"] b_le_1 [of "2*k+1"] | 
| 72531 | 1024 | using k a43 b43 \<open>0 < n\<close> | 
| 1025 | by (simp add: add.commute) | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1026 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1027 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1028 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1029 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1030 | then have aj_le_ci: "a (real j / 2 ^ n) \<le> c (real i / 2 ^ m)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1031 | and ci_le_bj: "c (real i / 2 ^ m) \<le> b (real j / 2 ^ n)" if "odd j" "\<bar>real i / 2^m - real j / 2^n\<bar> < 1/2 ^ n" for i j m n | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1032 | using that by blast+ | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1033 | have close_ab: "odd m \<Longrightarrow> \<bar>a (real m / 2 ^ n) - b (real m / 2 ^ n)\<bar> \<le> 2 / 2^n" for m n | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1034 | proof (induction n arbitrary: m) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1035 | case 0 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1036 | with u01 v01 show ?case by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1037 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1038 | case (Suc n m) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1039 | with oddE obtain k where k: "m = Suc (2*k)" by fastforce | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1040 | show ?case | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1041 | proof (cases "n > 0") | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1042 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1043 | with u01 v01 show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1044 | by (simp add: a_def b_def leftrec_base rightrec_base) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1045 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1046 | case True | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1047 | show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1048 | proof (cases "even k") | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1049 | case True | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1050 | then obtain j where j: "k = 2*j" by (metis evenE) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1051 | have "\<bar>a ((2 * real j + 1) / 2 ^ n) - (b ((2 * real j + 1) / 2 ^ n))\<bar> \<le> 2/2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1052 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1053 | have "odd (Suc k)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1054 | using True by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1055 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1056 | by (metis (no_types) Groups.add_ac(2) Suc.IH j of_nat_Suc of_nat_mult of_nat_numeral) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1057 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1058 | moreover have "a ((2 * real j + 1) / 2 ^ n) \<le> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1059 | leftcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1060 | using alec [of "2*j+1"] cleb [of "2*j+1"] a_ge_0 [of "2*j+1"] b_le_1 [of "2*j+1"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1061 | by (auto simp: add.commute left_right) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1062 | moreover have "leftcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n)) \<le> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1063 | c ((2 * real j + 1) / 2 ^ n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1064 | using alec [of "2*j+1"] cleb [of "2*j+1"] a_ge_0 [of "2*j+1"] b_le_1 [of "2*j+1"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1065 | by (auto simp: add.commute left_right_m) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1066 | ultimately have "\<bar>a ((2 * real j + 1) / 2 ^ n) - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1067 | leftcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n))\<bar> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1068 | \<le> 2/2 ^ Suc n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1069 | by (simp add: c_def midpoint_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1070 | with j k \<open>n > 0\<close> show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1071 | by (simp add: add.commute [of 1] a41 b41 del: power_Suc) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1072 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1073 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1074 | then obtain j where j: "k = 2*j + 1" by (metis oddE) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1075 | have "\<bar>a ((2 * real j + 1) / 2 ^ n) - (b ((2 * real j + 1) / 2 ^ n))\<bar> \<le> 2/2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1076 | using Suc.IH [OF False] j by (auto simp: algebra_simps) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1077 | moreover have "c ((2 * real j + 1) / 2 ^ n) \<le> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1078 | rightcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1079 | using alec [of "2*j+1"] cleb [of "2*j+1"] a_ge_0 [of "2*j+1"] b_le_1 [of "2*j+1"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1080 | by (auto simp: add.commute left_right_m) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1081 | moreover have "rightcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n)) \<le> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1082 | b ((2 * real j + 1) / 2 ^ n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1083 | using alec [of "2*j+1"] cleb [of "2*j+1"] a_ge_0 [of "2*j+1"] b_le_1 [of "2*j+1"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1084 | by (auto simp: add.commute left_right) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1085 | ultimately have "\<bar>rightcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n)) - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1086 | b ((2 * real j + 1) / 2 ^ n)\<bar> \<le> 2/2 ^ Suc n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1087 | by (simp add: c_def midpoint_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1088 | with j k \<open>n > 0\<close> show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1089 | by (simp add: add.commute [of 3] a43 b43 del: power_Suc) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1090 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1091 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1092 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1093 | have m1_to_3: "4 * real k - 1 = real (4 * (k-1)) + 3" if "0 < k" for k | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1094 | using that by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1095 | have fb_eq_fa: "\<lbrakk>0 < j; 2*j < 2 ^ n\<rbrakk> \<Longrightarrow> f(b((2 * real j - 1) / 2^n)) = f(a((2 * real j + 1) / 2^n))" for n j | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1096 | proof (induction n arbitrary: j) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1097 | case 0 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1098 | then show ?case by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1099 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1100 | case (Suc n j) show ?case | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1101 | proof (cases "n > 0") | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1102 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1103 | with Suc.prems show ?thesis by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1104 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1105 | case True | 
| 80175 
200107cdd3ac
Some new simprules – and patches for proofs
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1106 | show ?thesis | 
| 
200107cdd3ac
Some new simprules – and patches for proofs
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1107 | proof (cases "even j") | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1108 | case True | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1109 | then obtain k where k: "j = 2*k" by (metis evenE) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1110 | with \<open>0 < j\<close> have "k > 0" "2 * k < 2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1111 | using Suc.prems(2) k by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1112 | with k \<open>0 < n\<close> Suc.IH [of k] show ?thesis | 
| 80175 
200107cdd3ac
Some new simprules – and patches for proofs
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1113 | apply (simp add: m1_to_3 a41 b43 del: power_Suc of_nat_diff) | 
| 
200107cdd3ac
Some new simprules – and patches for proofs
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1114 | by simp | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1115 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1116 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1117 | then obtain k where k: "j = 2*k + 1" by (metis oddE) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1118 | have "f (leftcut (a ((2 * k + 1) / 2^n)) (b ((2 * k + 1) / 2^n)) (c ((2 * k + 1) / 2^n))) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1119 | = f (c ((2 * k + 1) / 2^n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1120 | "f (c ((2 * k + 1) / 2^n)) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1121 | = f (rightcut (a ((2 * k + 1) / 2^n)) (b ((2 * k + 1) / 2^n)) (c ((2 * k + 1) / 2^n)))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1122 | using alec [of "2*k+1" n] cleb [of "2*k+1" n] a_ge_0 [of "2*k+1" n] b_le_1 [of "2*k+1" n] k | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1123 | using left_right_m [of "c((2*k + 1) / 2^n)" "a((2*k + 1) / 2^n)" "b((2*k + 1) / 2^n)"] | 
| 72531 | 1124 | by (auto simp: add.commute feqm [OF order_refl] feqm [OF _ order_refl, symmetric]) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1125 | then | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1126 | show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1127 | by (simp add: k add.commute [of 1] add.commute [of 3] a43 b41\<open>0 < n\<close> del: power_Suc) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1128 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1129 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1130 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1131 | have f_eq_fc: "\<lbrakk>0 < j; j < 2 ^ n\<rbrakk> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1132 | \<Longrightarrow> f(b((2*j - 1) / 2 ^ (Suc n))) = f(c(j / 2^n)) \<and> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1133 | f(a((2*j + 1) / 2 ^ (Suc n))) = f(c(j / 2^n))" for n and j::nat | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1134 | proof (induction n arbitrary: j) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1135 | case 0 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1136 | then show ?case by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1137 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1138 | case (Suc n) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1139 | show ?case | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1140 | proof (cases "even j") | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1141 | case True | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1142 | then obtain k where k: "j = 2*k" by (metis evenE) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1143 | then have less2n: "k < 2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1144 | using Suc.prems(2) by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1145 | have "0 < k" using \<open>0 < j\<close> k by linarith | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1146 | then have m1_to_3: "real (4 * k - Suc 0) = real (4 * (k-1)) + 3" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1147 | by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1148 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1149 | using Suc.IH [of k] k \<open>0 < k\<close> | 
| 80175 
200107cdd3ac
Some new simprules – and patches for proofs
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 1150 | by (simp add: less2n add.commute [of 1] m1_to_3 a41 b43 del: power_Suc of_nat_diff) auto | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1151 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1152 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1153 | then obtain k where k: "j = 2*k + 1" by (metis oddE) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1154 | with Suc.prems have "k < 2^n" by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1155 | show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1156 | using alec [of "2*k+1" "Suc n"] cleb [of "2*k+1" "Suc n"] a_ge_0 [of "2*k+1" "Suc n"] b_le_1 [of "2*k+1" "Suc n"] k | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1157 | using left_right_m [of "c((2*k + 1) / 2 ^ Suc n)" "a((2*k + 1) / 2 ^ Suc n)" "b((2*k + 1) / 2 ^ Suc n)"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1158 | apply (simp add: add.commute [of 1] add.commute [of 3] m1_to_3 b41 a43 del: power_Suc) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1159 | apply (force intro: feqm) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1160 | done | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1161 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1162 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1163 |   define D01 where "D01 \<equiv> {0<..<1} \<inter> (\<Union>k m. {real m / 2^k})"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1164 |   have cloD01 [simp]: "closure D01 = {0..1}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1165 | unfolding D01_def | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1166 | by (subst closure_dyadic_rationals_in_convex_set_pos_1) auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1167 | have "uniformly_continuous_on D01 (f \<circ> c)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1168 | proof (clarsimp simp: uniformly_continuous_on_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1169 | fix e::real | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1170 | assume "0 < e" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1171 |     have ucontf: "uniformly_continuous_on {0..1} f"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1172 | by (simp add: compact_uniformly_continuous [OF cont_f]) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1173 |     then obtain d where "0 < d" and d: "\<And>x x'. \<lbrakk>x \<in> {0..1}; x' \<in> {0..1}; norm (x' - x) < d\<rbrakk> \<Longrightarrow> norm (f x' - f x) < e/2"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1174 | unfolding uniformly_continuous_on_def dist_norm | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1175 | by (metis \<open>0 < e\<close> less_divide_eq_numeral1(1) mult_zero_left) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1176 | obtain n where n: "1/2^n < min d 1" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1177 | by (metis \<open>0 < d\<close> divide_less_eq_1 less_numeral_extra(1) min_def one_less_numeral_iff power_one_over real_arch_pow_inv semiring_norm(76) zero_less_numeral) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1178 | with gr0I have "n > 0" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70136diff
changeset | 1179 | by (force simp: field_split_simps) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1180 | show "\<exists>d>0. \<forall>x\<in>D01. \<forall>x'\<in>D01. dist x' x < d \<longrightarrow> dist (f (c x')) (f (c x)) < e" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1181 | proof (intro exI ballI impI conjI) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1182 | show "(0::real) < 1/2^n" by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1183 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1184 | have dist_fc_close: "dist (f(c(real i / 2^m))) (f(c(real j / 2^n))) < e/2" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1185 | if i: "0 < i" "i < 2 ^ m" and j: "0 < j" "j < 2 ^ n" and clo: "abs(i / 2^m - j / 2^n) < 1/2 ^ n" for i j m | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1186 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1187 | have abs3: "\<bar>x - a\<bar> < e \<Longrightarrow> x = a \<or> \<bar>x - (a - e/2)\<bar> < e/2 \<or> \<bar>x - (a + e/2)\<bar> < e/2" for x a e::real | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1188 | by linarith | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1189 | consider "i / 2 ^ m = j / 2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1190 | | "\<bar>i / 2 ^ m - (2 * j - 1) / 2 ^ Suc n\<bar> < 1/2 ^ Suc n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1191 | | "\<bar>i / 2 ^ m - (2 * j + 1) / 2 ^ Suc n\<bar> < 1/2 ^ Suc n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1192 | using abs3 [OF clo] j by (auto simp: field_simps of_nat_diff) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1193 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1194 | proof cases | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1195 | case 1 with \<open>0 < e\<close> show ?thesis by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1196 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1197 | case 2 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1198 | have *: "abs(a - b) \<le> 1/2 ^ n \<and> 1/2 ^ n < d \<and> a \<le> c \<and> c \<le> b \<Longrightarrow> b - c < d" for a b c | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1199 | by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1200 | have "norm (c (real i / 2 ^ m) - b (real (2 * j - 1) / 2 ^ Suc n)) < d" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1201 | using 2 j n close_ab [of "2*j-1" "Suc n"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1202 | using b_ge_0 [of "2*j-1" "Suc n"] b_le_1 [of "2*j-1" "Suc n"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1203 | using aj_le_ci [of "2*j-1" i m "Suc n"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1204 | using ci_le_bj [of "2*j-1" i m "Suc n"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1205 | apply (simp add: divide_simps of_nat_diff del: power_Suc) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1206 | apply (auto simp: divide_simps intro!: *) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1207 | done | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1208 | moreover have "f(c(j / 2^n)) = f(b ((2*j - 1) / 2 ^ (Suc n)))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1209 | using f_eq_fc [OF j] by metis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1210 | ultimately show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1211 | by (metis dist_norm atLeastAtMost_iff b_ge_0 b_le_1 c_ge_0 c_le_1 d) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1212 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1213 | case 3 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1214 | have *: "abs(a - b) \<le> 1/2 ^ n \<and> 1/2 ^ n < d \<and> a \<le> c \<and> c \<le> b \<Longrightarrow> c - a < d" for a b c | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1215 | by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1216 | have "norm (c (real i / 2 ^ m) - a (real (2 * j + 1) / 2 ^ Suc n)) < d" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1217 | using 3 j n close_ab [of "2*j+1" "Suc n"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1218 | using b_ge_0 [of "2*j+1" "Suc n"] b_le_1 [of "2*j+1" "Suc n"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1219 | using aj_le_ci [of "2*j+1" i m "Suc n"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1220 | using ci_le_bj [of "2*j+1" i m "Suc n"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1221 | apply (simp add: divide_simps of_nat_diff del: power_Suc) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1222 | apply (auto simp: divide_simps intro!: *) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1223 | done | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1224 | moreover have "f(c(j / 2^n)) = f(a ((2*j + 1) / 2 ^ (Suc n)))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1225 | using f_eq_fc [OF j] by metis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1226 | ultimately show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1227 | by (metis dist_norm a_ge_0 atLeastAtMost_iff a_ge_0 a_le_1 c_ge_0 c_le_1 d) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1228 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1229 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1230 | show "dist (f (c x')) (f (c x)) < e" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1231 | if "x \<in> D01" "x' \<in> D01" "dist x' x < 1/2^n" for x x' | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1232 | using that unfolding D01_def dyadics_in_open_unit_interval | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1233 | proof clarsimp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1234 | fix i k::nat and m p | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1235 | assume i: "0 < i" "i < 2 ^ m" and k: "0<k" "k < 2 ^ p" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1236 | assume clo: "dist (real k / 2 ^ p) (real i / 2 ^ m) < 1/2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1237 | obtain j::nat where "0 < j" "j < 2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1238 | and clo_ij: "abs(i / 2^m - j / 2^n) < 1/2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1239 | and clo_kj: "abs(k / 2^p - j / 2^n) < 1/2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1240 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1241 | have "max (2^n * i / 2^m) (2^n * k / 2^p) \<ge> 0" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1242 | by (auto simp: le_max_iff_disj) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1243 | then obtain j where "floor (max (2^n*i / 2^m) (2^n*k / 2^p)) = int j" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1244 | using zero_le_floor zero_le_imp_eq_int by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1245 | then have j_le: "real j \<le> max (2^n * i / 2^m) (2^n * k / 2^p)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1246 | and less_j1: "max (2^n * i / 2^m) (2^n * k / 2^p) < real j + 1" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1247 | using floor_correct [of "max (2^n * i / 2^m) (2^n * k / 2^p)"] by linarith+ | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1248 | show thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1249 | proof (cases "j = 0") | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1250 | case True | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1251 | show thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1252 | proof | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1253 | show "(1::nat) < 2 ^ n" | 
| 72531 | 1254 | by (metis Suc_1 \<open>0 < n\<close> lessI one_less_power) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1255 | show "\<bar>real i / 2 ^ m - real 1/2 ^ n\<bar> < 1/2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1256 | using i less_j1 by (simp add: dist_norm field_simps True) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1257 | show "\<bar>real k / 2 ^ p - real 1/2 ^ n\<bar> < 1/2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1258 | using k less_j1 by (simp add: dist_norm field_simps True) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1259 | qed simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1260 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1261 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1262 | have 1: "real j * 2 ^ m < real i * 2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1263 | if j: "real j * 2 ^ p \<le> real k * 2 ^ n" and k: "real k * 2 ^ m < real i * 2 ^ p" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1264 | for i k m p | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1265 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1266 | have "real j * 2 ^ p * 2 ^ m \<le> real k * 2 ^ n * 2 ^ m" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1267 | using j by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1268 | moreover have "real k * 2 ^ m * 2 ^ n < real i * 2 ^ p * 2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1269 | using k by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1270 | ultimately have "real j * 2 ^ p * 2 ^ m < real i * 2 ^ p * 2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1271 | by (simp only: mult_ac) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1272 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1273 | by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1274 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1275 | have 2: "real j * 2 ^ m < 2 ^ m + real i * 2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1276 | if j: "real j * 2 ^ p \<le> real k * 2 ^ n" and k: "real k * (2 ^ m * 2 ^ n) < 2 ^ m * 2 ^ p + real i * (2 ^ n * 2 ^ p)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1277 | for i k m p | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1278 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1279 | have "real j * 2 ^ p * 2 ^ m \<le> real k * (2 ^ m * 2 ^ n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1280 | using j by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1281 | also have "... < 2 ^ m * 2 ^ p + real i * (2 ^ n * 2 ^ p)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1282 | by (rule k) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1283 | finally have "(real j * 2 ^ m) * 2 ^ p < (2 ^ m + real i * 2 ^ n) * 2 ^ p" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1284 | by (simp add: algebra_simps) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1285 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1286 | by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1287 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1288 | have 3: "real j * 2 ^ p < 2 ^ p + real k * 2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1289 | if j: "real j * 2 ^ m \<le> real i * 2 ^ n" and i: "real i * 2 ^ p \<le> real k * 2 ^ m" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1290 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1291 | have "real j * 2 ^ m * 2 ^ p \<le> real i * 2 ^ n * 2 ^ p" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1292 | using j by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1293 | moreover have "real i * 2 ^ p * 2 ^ n \<le> real k * 2 ^ m * 2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1294 | using i by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1295 | ultimately have "real j * 2 ^ m * 2 ^ p \<le> real k * 2 ^ m * 2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1296 | by (simp only: mult_ac) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1297 | then have "real j * 2 ^ p \<le> real k * 2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1298 | by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1299 | also have "... < 2 ^ p + real k * 2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1300 | by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1301 | finally show ?thesis by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1302 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1303 | show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1304 | proof | 
| 72531 | 1305 | have "2 ^ n * real i / 2 ^ m < 2 ^ n" "2 ^ n * real k / 2 ^ p < 2 ^ n" | 
| 1306 | using i k by (auto simp: field_simps) | |
| 1307 | then have "max (2^n * i / 2^m) (2^n * k / 2^p) < 2^n" | |
| 1308 | by simp | |
| 1309 | with j_le have "real j < 2 ^ n" by linarith | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1310 | then show "j < 2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1311 | by auto | 
| 72531 | 1312 | have "\<bar>real i * 2 ^ n - real j * 2 ^ m\<bar> < 2 ^ m" | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1313 | using clo less_j1 j_le | 
| 72531 | 1314 | by (auto simp: le_max_iff_disj field_split_simps dist_norm abs_if split: if_split_asm dest: 1 2) | 
| 1315 | then show "\<bar>real i / 2 ^ m - real j / 2 ^ n\<bar> < 1/2 ^ n" | |
| 1316 | by (auto simp: field_split_simps) | |
| 1317 | have "\<bar>real k * 2 ^ n - real j * 2 ^ p\<bar> < 2 ^ p" | |
| 1318 | using clo less_j1 j_le | |
| 1319 | by (auto simp: le_max_iff_disj field_split_simps dist_norm abs_if split: if_split_asm dest: 3 2) | |
| 1320 | then show "\<bar>real k / 2 ^ p - real j / 2 ^ n\<bar> < 1/2 ^ n" | |
| 1321 | by (auto simp: le_max_iff_disj field_split_simps dist_norm) | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1322 | qed (use False in simp) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1323 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1324 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1325 | show "dist (f (c (real k / 2 ^ p))) (f (c (real i / 2 ^ m))) < e" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1326 | proof (rule dist_triangle_half_l) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1327 | show "dist (f (c (real k / 2 ^ p))) (f(c(j / 2^n))) < e/2" | 
| 72531 | 1328 | using \<open>0 < j\<close> \<open>j < 2 ^ n\<close> k clo_kj | 
| 1329 | by (intro dist_fc_close) auto | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1330 | show "dist (f (c (real i / 2 ^ m))) (f (c (real j / 2 ^ n))) < e/2" | 
| 72531 | 1331 | using \<open>0 < j\<close> \<open>j < 2 ^ n\<close> i clo_ij | 
| 1332 | by (intro dist_fc_close) auto | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1333 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1334 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1335 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1336 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1337 |   then obtain h where ucont_h: "uniformly_continuous_on {0..1} h"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1338 | and fc_eq: "\<And>x. x \<in> D01 \<Longrightarrow> (f \<circ> c) x = h x" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1339 | proof (rule uniformly_continuous_on_extension_on_closure [of D01 "f \<circ> c"]) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1340 | qed (use closure_subset [of D01] in \<open>auto intro!: that\<close>) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1341 |   then have cont_h: "continuous_on {0..1} h"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1342 | using uniformly_continuous_imp_continuous by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1343 | have h_eq: "h (real k / 2 ^ m) = f (c (real k / 2 ^ m))" if "0 < k" "k < 2^m" for k m | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1344 | using fc_eq that by (force simp: D01_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1345 |   have "h ` {0..1} = f ` {0..1}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1346 | proof | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1347 |     have "h ` (closure D01) \<subseteq> f ` {0..1}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1348 | proof (rule image_closure_subset) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1349 | show "continuous_on (closure D01) h" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1350 | using cont_h by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1351 |       show "closed (f ` {0..1})"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1352 | using compact_continuous_image [OF cont_f] compact_imp_closed by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1353 |       show "h ` D01 \<subseteq> f ` {0..1}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1354 | by (force simp: dyadics_in_open_unit_interval D01_def h_eq) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1355 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1356 |     with cloD01 show "h ` {0..1} \<subseteq> f ` {0..1}" by simp
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1357 | have a12 [simp]: "a (1/2) = u" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1358 | by (metis a_def leftrec_base numeral_One of_nat_numeral) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1359 | have b12 [simp]: "b (1/2) = v" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1360 | by (metis b_def rightrec_base numeral_One of_nat_numeral) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1361 |     have "f ` {0..1} \<subseteq> closure(h ` D01)"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1362 | proof (clarsimp simp: closure_approachable dyadics_in_open_unit_interval D01_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1363 | fix x e::real | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1364 | assume "0 \<le> x" "x \<le> 1" "0 < e" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1365 |       have ucont_f: "uniformly_continuous_on {0..1} f"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1366 | using compact_uniformly_continuous cont_f by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1367 | then obtain \<delta> where "\<delta> > 0" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1368 |         and \<delta>: "\<And>x x'. \<lbrakk>x \<in> {0..1}; x' \<in> {0..1}; dist x' x < \<delta>\<rbrakk> \<Longrightarrow> norm (f x' - f x) < e"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1369 | using \<open>0 < e\<close> by (auto simp: uniformly_continuous_on_def dist_norm) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1370 |       have *: "\<exists>m::nat. \<exists>y. odd m \<and> 0 < m \<and> m < 2 ^ n \<and> y \<in> {a(m / 2^n) .. b(m / 2^n)} \<and> f y = f x"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1371 | if "n \<noteq> 0" for n | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1372 | using that | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1373 | proof (induction n) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1374 | case 0 then show ?case by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1375 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1376 | case (Suc n) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1377 | show ?case | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1378 | proof (cases "n=0") | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1379 | case True | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1380 |           consider "x \<in> {0..u}" | "x \<in> {u..v}" | "x \<in> {v..1}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1381 | using \<open>0 \<le> x\<close> \<open>x \<le> 1\<close> by force | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1382 | then have "\<exists>y\<ge>a (real 1/2). y \<le> b (real 1/2) \<and> f y = f x" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1383 | proof cases | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1384 | case 1 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1385 | then show ?thesis | 
| 72531 | 1386 | using uabv [of 1 1] f0u [of u] f0u [of x] by force | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1387 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1388 | case 2 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1389 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1390 | by (rule_tac x=x in exI) auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1391 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1392 | case 3 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1393 | then show ?thesis | 
| 72531 | 1394 | using uabv [of 1 1] fv1 [of v] fv1 [of x] by force | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1395 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1396 | with \<open>n=0\<close> show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1397 | by (rule_tac x=1 in exI) auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1398 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1399 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1400 | with Suc obtain m y | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1401 | where "odd m" "0 < m" and mless: "m < 2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1402 |               and y: "y \<in> {a (real m / 2 ^ n)..b (real m / 2 ^ n)}" and feq: "f y = f x"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1403 | by metis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1404 | then obtain j where j: "m = 2*j + 1" by (metis oddE) | 
| 72531 | 1405 | have j4: "4 * j + 1 < 2 ^ Suc n" | 
| 1406 | using mless j by (simp add: algebra_simps) | |
| 1407 | ||
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1408 |           consider "y \<in> {a((2*j + 1) / 2^n) .. b((4*j + 1) / 2 ^ (Suc n))}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1409 |             | "y \<in> {b((4*j + 1) / 2 ^ (Suc n)) .. a((4*j + 3) / 2 ^ (Suc n))}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1410 |             | "y \<in> {a((4*j + 3) / 2 ^ (Suc n)) .. b((2*j + 1) / 2^n)}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1411 | using y j by force | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1412 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1413 | proof cases | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1414 | case 1 | 
| 72531 | 1415 | show ?thesis | 
| 1416 | proof (intro exI conjI) | |
| 1417 |               show "y \<in> {a (real (4 * j + 1) / 2 ^ Suc n)..b (real (4 * j + 1) / 2 ^ Suc n)}"
 | |
| 1418 | using mless j \<open>n \<noteq> 0\<close> 1 by (simp add: a41 b41 add.commute [of 1] del: power_Suc) | |
| 1419 | qed (use feq j4 in auto) | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1420 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1421 | case 2 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1422 | show ?thesis | 
| 72531 | 1423 | proof (intro exI conjI) | 
| 1424 |               show "b (real (4 * j + 1) / 2 ^ Suc n) \<in> {a (real (4 * j + 1) / 2 ^ Suc n)..b (real (4 * j + 1) / 2 ^ Suc n)}"
 | |
| 1425 | using \<open>n \<noteq> 0\<close> alec [of "2*j+1" n] cleb [of "2*j+1" n] a_ge_0 [of "2*j+1" n] b_le_1 [of "2*j+1" n] | |
| 1426 | using left_right [of "c((2*j + 1) / 2^n)" "a((2*j + 1) / 2^n)" "b((2*j + 1) / 2^n)"] | |
| 1427 | by (simp add: a41 b41 add.commute [of 1] del: power_Suc) | |
| 1428 | show "f (b (real (4 * j + 1) / 2 ^ Suc n)) = f x" | |
| 1429 | using \<open>n \<noteq> 0\<close> 2 | |
| 1430 | using alec [of "2*j+1" n] cleb [of "2*j+1" n] a_ge_0 [of "2*j+1" n] b_le_1 [of "2*j+1" n] | |
| 1431 | by (force simp add: b41 a43 add.commute [of 1] feq [symmetric] simp del: power_Suc intro: f_eqI) | |
| 1432 | qed (use j4 in auto) | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1433 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1434 | case 3 | 
| 72531 | 1435 | show ?thesis | 
| 1436 | proof (intro exI conjI) | |
| 1437 | show "4 * j + 3 < 2 ^ Suc n" | |
| 1438 | using mless j by simp | |
| 1439 | show "f y = f x" | |
| 1440 | by fact | |
| 1441 |               show "y \<in> {a (real (4 * j + 3) / 2 ^ Suc n) .. b (real (4 * j + 3) / 2 ^ Suc n)}"
 | |
| 1442 | using 3 False b43 [of n j] by (simp add: add.commute) | |
| 1443 | qed (use 3 in auto) | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1444 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1445 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1446 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1447 | obtain n where n: "1/2^n < min (\<delta> / 2) 1" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1448 | by (metis \<open>0 < \<delta>\<close> divide_less_eq_1 less_numeral_extra(1) min_less_iff_conj one_less_numeral_iff power_one_over real_arch_pow_inv semiring_norm(76) zero_less_divide_iff zero_less_numeral) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1449 | with gr0I have "n \<noteq> 0" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1450 | by fastforce | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1451 | with * obtain m::nat and y | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1452 | where "odd m" "0 < m" and mless: "m < 2 ^ n" | 
| 72531 | 1453 | and y: "a(m / 2^n) \<le> y \<and> y \<le> b(m / 2^n)" and feq: "f x = f y" | 
| 1454 | by (metis atLeastAtMost_iff) | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1455 | then have "0 \<le> y" "y \<le> 1" | 
| 72531 | 1456 | by (meson a_ge_0 b_le_1 order.trans)+ | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1457 | moreover have "y < \<delta> + c (real m / 2 ^ n)" "c (real m / 2 ^ n) < \<delta> + y" | 
| 72531 | 1458 | using y alec [of m n] cleb [of m n] n field_sum_of_halves close_ab [OF \<open>odd m\<close>, of n] | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1459 | by linarith+ | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1460 | moreover note \<open>0 < m\<close> mless \<open>0 \<le> x\<close> \<open>x \<le> 1\<close> | 
| 72531 | 1461 | ultimately have "dist (h (real m / 2 ^ n)) (f x) < e" | 
| 1462 | by (auto simp: dist_norm h_eq feq \<delta>) | |
| 1463 |       then show "\<exists>k. \<exists>m\<in>{0<..<2 ^ k}. dist (h (real m / 2 ^ k)) (f x) < e"
 | |
| 1464 | using \<open>0 < m\<close> greaterThanLessThan_iff mless by blast | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1465 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1466 |     also have "... \<subseteq> h ` {0..1}"
 | 
| 72531 | 1467 | proof (rule closure_minimal) | 
| 1468 |       show "h ` D01 \<subseteq> h ` {0..1}"
 | |
| 1469 | using cloD01 closure_subset by blast | |
| 1470 |       show "closed (h ` {0..1})"
 | |
| 1471 | using compact_continuous_image [OF cont_h] compact_imp_closed by auto | |
| 1472 | qed | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1473 |     finally show "f ` {0..1} \<subseteq> h ` {0..1}" .
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1474 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1475 |   moreover have "inj_on h {0..1}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1476 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1477 | have "u < v" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1478 | by (metis atLeastAtMost_iff f0u f_1not0 fv1 order.not_eq_order_implies_strict u01(1) u01(2) v01(1)) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1479 | have f_not_fu: "\<And>x. \<lbrakk>u < x; x \<le> v\<rbrakk> \<Longrightarrow> f x \<noteq> f u" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1480 | by (metis atLeastAtMost_iff f0u fu1 greaterThanAtMost_iff order_refl order_trans u01(1) v01(2)) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1481 | have f_not_fv: "\<And>x. \<lbrakk>u \<le> x; x < v\<rbrakk> \<Longrightarrow> f x \<noteq> f v" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1482 | by (metis atLeastAtMost_iff order_refl order_trans v01(2) atLeastLessThan_iff fuv fv1) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1483 | have a_less_b: | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1484 | "a(j / 2^n) < b(j / 2^n) \<and> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1485 | (\<forall>x. a(j / 2^n) < x \<longrightarrow> x \<le> b(j / 2^n) \<longrightarrow> f x \<noteq> f(a(j / 2^n))) \<and> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1486 | (\<forall>x. a(j / 2^n) \<le> x \<longrightarrow> x < b(j / 2^n) \<longrightarrow> f x \<noteq> f(b(j / 2^n)))" for n and j::nat | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1487 | proof (induction n arbitrary: j) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1488 | case 0 then show ?case | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1489 | by (simp add: \<open>u < v\<close> f_not_fu f_not_fv) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1490 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1491 | case (Suc n j) show ?case | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1492 | proof (cases "n > 0") | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1493 | case False then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1494 | by (auto simp: a_def b_def leftrec_base rightrec_base \<open>u < v\<close> f_not_fu f_not_fv) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1495 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1496 | case True show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1497 | proof (cases "even j") | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1498 | case True | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1499 | with \<open>0 < n\<close> Suc.IH show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1500 | by (auto elim!: evenE) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1501 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1502 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1503 | then obtain k where k: "j = 2*k + 1" by (metis oddE) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1504 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1505 | proof (cases "even k") | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1506 | case True | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1507 | then obtain m where m: "k = 2*m" by (metis evenE) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1508 | have fleft: "f (leftcut (a ((2*m + 1) / 2^n)) (b ((2*m + 1) / 2^n)) (c ((2*m + 1) / 2^n))) = | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1509 | f (c((2*m + 1) / 2^n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1510 | using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1511 | using left_right_m [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1512 | by (auto intro: f_eqI) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1513 | show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1514 | proof (intro conjI impI notI allI) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1515 | have False if "b (real j / 2 ^ Suc n) \<le> a (real j / 2 ^ Suc n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1516 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1517 | have "f (c ((1 + real m * 2) / 2 ^ n)) = f (a ((1 + real m * 2) / 2 ^ n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1518 | using k m \<open>0 < n\<close> fleft that a41 [of n m] b41 [of n m] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1519 | using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1520 | using left_right [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1521 | by (auto simp: algebra_simps) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1522 | moreover have "a (real (1 + m * 2) / 2 ^ n) < c (real (1 + m * 2) / 2 ^ n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1523 | using Suc.IH [of "1 + m * 2"] by (simp add: c_def midpoint_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1524 | moreover have "c (real (1 + m * 2) / 2 ^ n) \<le> b (real (1 + m * 2) / 2 ^ n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1525 | using cleb by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1526 | ultimately show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1527 | using Suc.IH [of "1 + m * 2"] by force | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1528 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1529 | then show "a (real j / 2 ^ Suc n) < b (real j / 2 ^ Suc n)" by force | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1530 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1531 | fix x | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1532 | assume "a (real j / 2 ^ Suc n) < x" "x \<le> b (real j / 2 ^ Suc n)" "f x = f (a (real j / 2 ^ Suc n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1533 | then show False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1534 | using Suc.IH [of "1 + m * 2", THEN conjunct2, THEN conjunct1] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1535 | using k m \<open>0 < n\<close> a41 [of n m] b41 [of n m] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1536 | using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1537 | using left_right_m [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1538 | by (auto simp: algebra_simps) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1539 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1540 | fix x | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1541 | assume "a (real j / 2 ^ Suc n) \<le> x" "x < b (real j / 2 ^ Suc n)" "f x = f (b (real j / 2 ^ Suc n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1542 | then show False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1543 | using k m \<open>0 < n\<close> a41 [of n m] b41 [of n m] fleft left_neq | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1544 | using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1545 | by (auto simp: algebra_simps) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1546 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1547 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1548 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1549 | with oddE obtain m where m: "k = Suc (2*m)" by fastforce | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1550 | have fright: "f (rightcut (a ((2*m + 1) / 2^n)) (b ((2*m + 1) / 2^n)) (c ((2*m + 1) / 2^n))) = f (c((2*m + 1) / 2^n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1551 | using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1552 | using left_right_m [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1553 | by (auto intro: f_eqI [OF _ order_refl]) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1554 | show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1555 | proof (intro conjI impI notI allI) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1556 | have False if "b (real j / 2 ^ Suc n) \<le> a (real j / 2 ^ Suc n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1557 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1558 | have "f (c ((1 + real m * 2) / 2 ^ n)) = f (b ((1 + real m * 2) / 2 ^ n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1559 | using k m \<open>0 < n\<close> fright that a43 [of n m] b43 [of n m] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1560 | using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1561 | using left_right [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1562 | by (auto simp: algebra_simps) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1563 | moreover have "a (real (1 + m * 2) / 2 ^ n) \<le> c (real (1 + m * 2) / 2 ^ n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1564 | using alec by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1565 | moreover have "c (real (1 + m * 2) / 2 ^ n) < b (real (1 + m * 2) / 2 ^ n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1566 | using Suc.IH [of "1 + m * 2"] by (simp add: c_def midpoint_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1567 | ultimately show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1568 | using Suc.IH [of "1 + m * 2"] by force | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1569 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1570 | then show "a (real j / 2 ^ Suc n) < b (real j / 2 ^ Suc n)" by force | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1571 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1572 | fix x | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1573 | assume "a (real j / 2 ^ Suc n) < x" "x \<le> b (real j / 2 ^ Suc n)" "f x = f (a (real j / 2 ^ Suc n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1574 | then show False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1575 | using k m \<open>0 < n\<close> a43 [of n m] b43 [of n m] fright right_neq | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1576 | using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1577 | by (auto simp: algebra_simps) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1578 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1579 | fix x | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1580 | assume "a (real j / 2 ^ Suc n) \<le> x" "x < b (real j / 2 ^ Suc n)" "f x = f (b (real j / 2 ^ Suc n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1581 | then show False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1582 | using Suc.IH [of "1 + m * 2", THEN conjunct2, THEN conjunct2] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1583 | using k m \<open>0 < n\<close> a43 [of n m] b43 [of n m] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1584 | using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1585 | using left_right_m [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1586 | by (auto simp: algebra_simps fright simp del: power_Suc) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1587 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1588 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1589 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1590 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1591 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1592 | have c_gt_0 [simp]: "0 < c(m / 2^n)" and c_less_1 [simp]: "c(m / 2^n) < 1" for m::nat and n | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1593 | using a_less_b [of m n] apply (simp_all add: c_def midpoint_def) | 
| 72531 | 1594 | using a_ge_0 [of m n] b_le_1 [of m n] by linarith+ | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1595 | have approx: "\<exists>j n. odd j \<and> n \<noteq> 0 \<and> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1596 | real i / 2^m \<le> real j / 2^n \<and> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1597 | real j / 2^n \<le> real k / 2^p \<and> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1598 | \<bar>real i / 2 ^ m - real j / 2 ^ n\<bar> < 1/2^n \<and> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1599 | \<bar>real k / 2 ^ p - real j / 2 ^ n\<bar> < 1/2^n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1600 | if "0 < i" "i < 2 ^ m" "0 < k" "k < 2 ^ p" "i / 2^m < k / 2^p" "m + p = N" for N m p i k | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1601 | using that | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1602 | proof (induction N arbitrary: m p i k rule: less_induct) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1603 | case (less N) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1604 | then consider "i / 2^m \<le> 1/2" "1/2 \<le> k / 2^p" | "k / 2^p < 1/2" | "k / 2^p \<ge> 1/2" "1/2 < i / 2^m" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1605 | by linarith | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1606 | then show ?case | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1607 | proof cases | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1608 | case 1 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1609 | with less.prems show ?thesis | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70136diff
changeset | 1610 | by (rule_tac x=1 in exI)+ (fastforce simp: field_split_simps) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1611 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1612 | case 2 show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1613 | proof (cases m) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1614 | case 0 with less.prems show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1615 | by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1616 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1617 | case (Suc m') show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1618 | proof (cases p) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1619 | case 0 with less.prems show ?thesis by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1620 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1621 | case (Suc p') | 
| 72531 | 1622 | have \<section>: False if "real i * 2 ^ p' < real k * 2 ^ m'" "k < 2 ^ p'" "2 ^ m' \<le> i" | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1623 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1624 | have "real k * 2 ^ m' < 2 ^ p' * 2 ^ m'" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1625 | using that by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1626 | then have "real i * 2 ^ p' < 2 ^ p' * 2 ^ m'" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1627 | using that by linarith | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1628 | with that show ?thesis by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1629 | qed | 
| 72531 | 1630 | moreover have *: "real i / 2 ^ m' < real k / 2^p'" "k < 2 ^ p'" | 
| 1631 | using less.prems \<open>m = Suc m'\<close> 2 Suc by (force simp: field_split_simps)+ | |
| 1632 | moreover have "i < 2 ^ m' " | |
| 1633 | using \<section> * by (clarsimp simp: divide_simps linorder_not_le) (meson linorder_not_le) | |
| 1634 | ultimately show ?thesis | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1635 | using less.IH [of "m'+p'" i m' k p'] less.prems \<open>m = Suc m'\<close> 2 Suc | 
| 72531 | 1636 | by (force simp: field_split_simps) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1637 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1638 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1639 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1640 | case 3 show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1641 | proof (cases m) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1642 | case 0 with less.prems show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1643 | by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1644 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1645 | case (Suc m') show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1646 | proof (cases p) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1647 | case 0 with less.prems show ?thesis by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1648 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1649 | case (Suc p') | 
| 72531 | 1650 | have "real (i - 2 ^ m') / 2 ^ m' < real (k - 2 ^ p') / 2 ^ p'" | 
| 1651 | using less.prems \<open>m = Suc m'\<close> Suc 3 by (auto simp: field_simps of_nat_diff) | |
| 1652 | moreover have "k - 2 ^ p' < 2 ^ p'" "i - 2 ^ m' < 2 ^ m'" | |
| 1653 | using less.prems Suc \<open>m = Suc m'\<close> by auto | |
| 1654 | moreover | |
| 1655 | have "2 ^ p' \<le> k" "2 ^ p' \<noteq> k" | |
| 1656 | using less.prems \<open>m = Suc m'\<close> Suc 3 by auto | |
| 1657 | then have "2 ^ p' < k" | |
| 1658 | by linarith | |
| 1659 | ultimately show ?thesis | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1660 | using less.IH [of "m'+p'" "i - 2^m'" m' "k - 2 ^ p'" p'] less.prems \<open>m = Suc m'\<close> Suc 3 | 
| 72531 | 1661 | apply (clarsimp simp: field_simps of_nat_diff) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1662 | apply (rule_tac x="2 ^ n + j" in exI, simp) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1663 | apply (rule_tac x="Suc n" in exI) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1664 | apply (auto simp: field_simps) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1665 | done | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1666 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1667 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1668 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1669 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1670 | have clec: "c(real i / 2^m) \<le> c(real j / 2^n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1671 | if i: "0 < i" "i < 2 ^ m" and j: "0 < j" "j < 2 ^ n" and ij: "i / 2^m < j / 2^n" for m i n j | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1672 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1673 | obtain j' n' where "odd j'" "n' \<noteq> 0" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1674 | and i_le_j: "real i / 2 ^ m \<le> real j' / 2 ^ n'" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1675 | and j_le_j: "real j' / 2 ^ n' \<le> real j / 2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1676 | and clo_ij: "\<bar>real i / 2 ^ m - real j' / 2 ^ n'\<bar> < 1/2 ^ n'" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1677 | and clo_jj: "\<bar>real j / 2 ^ n - real j' / 2 ^ n'\<bar> < 1/2 ^ n'" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1678 | using approx [of i m j n "m+n"] that i j ij by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1679 | with oddE obtain q where q: "j' = Suc (2*q)" by fastforce | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1680 | have "c (real i / 2 ^ m) \<le> c((2*q + 1) / 2^n')" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1681 | proof (cases "i / 2^m = (2*q + 1) / 2^n'") | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1682 | case True then show ?thesis by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1683 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1684 | case False | 
| 72531 | 1685 | with i_le_j clo_ij q have "\<bar>real i / 2 ^ m - real (4 * q + 1) / 2 ^ Suc n'\<bar> < 1 / 2 ^ Suc n'" | 
| 1686 | by (auto simp: field_split_simps) | |
| 1687 | then have "c(i / 2^m) \<le> b(real(4 * q + 1) / 2 ^ (Suc n'))" | |
| 1688 | by (meson ci_le_bj even_mult_iff even_numeral even_plus_one_iff) | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1689 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1690 | using alec [of "2*q+1" n'] cleb [of "2*q+1" n'] a_ge_0 [of "2*q+1" n'] b_le_1 [of "2*q+1" n'] b41 [of n' q] \<open>n' \<noteq> 0\<close> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1691 | using left_right_m [of "c((2*q + 1) / 2^n')" "a((2*q + 1) / 2^n')" "b((2*q + 1) / 2^n')"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1692 | by (auto simp: algebra_simps) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1693 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1694 | also have "... \<le> c(real j / 2^n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1695 | proof (cases "j / 2^n = (2*q + 1) / 2^n'") | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1696 | case True | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1697 | then show ?thesis by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1698 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1699 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1700 | with j_le_j q have less: "(2*q + 1) / 2^n' < j / 2^n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1701 | by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1702 | have *: "\<lbrakk>q < i; abs(i - q) < s*2; r = q + s\<rbrakk> \<Longrightarrow> abs(i - r) < s" for i q s r::real | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1703 | by auto | 
| 72531 | 1704 | have "\<bar>real j / 2 ^ n - real (4 * q + 3) / 2 ^ Suc n'\<bar> < 1 / 2 ^ Suc n'" | 
| 1705 | by (rule * [OF less]) (use j_le_j clo_jj q in \<open>auto simp: field_split_simps\<close>) | |
| 1706 | then have "a(real(4*q + 3) / 2 ^ (Suc n')) \<le> c(j / 2^n)" | |
| 1707 | by (metis Suc3_eq_add_3 add.commute aj_le_ci even_Suc even_mult_iff even_numeral) | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1708 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1709 | using alec [of "2*q+1" n'] cleb [of "2*q+1" n'] a_ge_0 [of "2*q+1" n'] b_le_1 [of "2*q+1" n'] a43 [of n' q] \<open>n' \<noteq> 0\<close> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1710 | using left_right_m [of "c((2*q + 1) / 2^n')" "a((2*q + 1) / 2^n')" "b((2*q + 1) / 2^n')"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1711 | by (auto simp: algebra_simps) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1712 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1713 | finally show ?thesis . | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1714 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1715 | have "x = y" if "0 \<le> x" "x \<le> 1" "0 \<le> y" "y \<le> 1" "h x = h y" for x y | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1716 | using that | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1717 | proof (induction x y rule: linorder_class.linorder_less_wlog) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1718 | case (less x1 x2) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1719 | obtain m n where m: "0 < m" "m < 2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1720 | and x12: "x1 < m / 2^n" "m / 2^n < x2" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1721 | and neq: "h x1 \<noteq> h (real m / 2^n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1722 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1723 | have "(x1 + x2) / 2 \<in> closure D01" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1724 | using cloD01 less.hyps less.prems by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1725 | with less obtain y where "y \<in> D01" and dist_y: "dist y ((x1 + x2) / 2) < (x2 - x1) / 64" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1726 | unfolding closure_approachable | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1727 | by (metis diff_gt_0_iff_gt less_divide_eq_numeral1(1) mult_zero_left) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1728 | obtain m n where m: "0 < m" "m < 2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1729 | and clo: "\<bar>real m / 2 ^ n - (x1 + x2) / 2\<bar> < (x2 - x1) / 64" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1730 | and n: "1/2^n < (x2 - x1) / 128" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1731 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1732 | have "min 1 ((x2 - x1) / 128) > 0" "1/2 < (1::real)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1733 | using less by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1734 | then obtain N where N: "1/2^N < min 1 ((x2 - x1) / 128)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1735 | by (metis power_one_over real_arch_pow_inv) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1736 | then have "N > 0" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1737 | using less_divide_eq_1 by force | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1738 | obtain p q where p: "p < 2 ^ q" "p \<noteq> 0" and yeq: "y = real p / 2 ^ q" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1739 | using \<open>y \<in> D01\<close> by (auto simp: zero_less_divide_iff D01_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1740 | show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1741 | proof | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1742 | show "0 < 2^N * p" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1743 | using p by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1744 | show "2 ^ N * p < 2 ^ (N+q)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1745 | by (simp add: p power_add) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1746 | have "\<bar>real (2 ^ N * p) / 2 ^ (N + q) - (x1 + x2) / 2\<bar> = \<bar>real p / 2 ^ q - (x1 + x2) / 2\<bar>" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1747 | by (simp add: power_add) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1748 | also have "... = \<bar>y - (x1 + x2) / 2\<bar>" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1749 | by (simp add: yeq) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1750 | also have "... < (x2 - x1) / 64" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1751 | using dist_y by (simp add: dist_norm) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1752 | finally show "\<bar>real (2 ^ N * p) / 2 ^ (N + q) - (x1 + x2) / 2\<bar> < (x2 - x1) / 64" . | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1753 | have "(1::real) / 2 ^ (N + q) \<le> 1/2^N" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1754 | by (simp add: field_simps) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1755 | also have "... < (x2 - x1) / 128" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1756 | using N by force | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1757 | finally show "1/2 ^ (N + q) < (x2 - x1) / 128" . | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1758 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1759 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1760 | obtain m' n' m'' n'' where "0 < m'" "m' < 2 ^ n'" "x1 < m' / 2^n'" "m' / 2^n' < x2" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1761 | and "0 < m''" "m'' < 2 ^ n''" "x1 < m'' / 2^n''" "m'' / 2^n'' < x2" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1762 | and neq: "h (real m'' / 2^n'') \<noteq> h (real m' / 2^n')" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1763 | proof | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1764 | show "0 < Suc (2*m)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1765 | by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1766 | show m21: "Suc (2*m) < 2 ^ Suc n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1767 | using m by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1768 | show "x1 < real (Suc (2 * m)) / 2 ^ Suc n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1769 | using clo by (simp add: field_simps abs_if split: if_split_asm) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1770 | show "real (Suc (2 * m)) / 2 ^ Suc n < x2" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1771 | using n clo by (simp add: field_simps abs_if split: if_split_asm) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1772 | show "0 < 4*m + 3" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1773 | by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1774 | have "m+1 \<le> 2 ^ n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1775 | using m by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1776 | then have "4 * (m+1) \<le> 4 * (2 ^ n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1777 | by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1778 | then show m43: "4*m + 3 < 2 ^ (n+2)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1779 | by (simp add: algebra_simps) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1780 | show "x1 < real (4 * m + 3) / 2 ^ (n + 2)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1781 | using clo by (simp add: field_simps abs_if split: if_split_asm) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1782 | show "real (4 * m + 3) / 2 ^ (n + 2) < x2" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1783 | using n clo by (simp add: field_simps abs_if split: if_split_asm) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1784 | have c_fold: "midpoint (a ((2 * real m + 1) / 2 ^ Suc n)) (b ((2 * real m + 1) / 2 ^ Suc n)) = c ((2 * real m + 1) / 2 ^ Suc n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1785 | by (simp add: c_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1786 | define R where "R \<equiv> rightcut (a ((2 * real m + 1) / 2 ^ Suc n)) (b ((2 * real m + 1) / 2 ^ Suc n)) (c ((2 * real m + 1) / 2 ^ Suc n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1787 | have "R < b ((2 * real m + 1) / 2 ^ Suc n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1788 | unfolding R_def using a_less_b [of "4*m + 3" "n+2"] a43 [of "Suc n" m] b43 [of "Suc n" m] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1789 | by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1790 | then have Rless: "R < midpoint R (b ((2 * real m + 1) / 2 ^ Suc n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1791 | by (simp add: midpoint_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1792 | have midR_le: "midpoint R (b ((2 * real m + 1) / 2 ^ Suc n)) \<le> b ((2 * real m + 1) / (2 * 2 ^ n))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1793 | using \<open>R < b ((2 * real m + 1) / 2 ^ Suc n)\<close> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1794 | by (simp add: midpoint_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1795 | have "(real (Suc (2 * m)) / 2 ^ Suc n) \<in> D01" "real (4 * m + 3) / 2 ^ (n + 2) \<in> D01" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1796 | by (simp_all add: D01_def m21 m43 del: power_Suc of_nat_Suc of_nat_add add_2_eq_Suc') blast+ | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1797 | then show "h (real (4 * m + 3) / 2 ^ (n + 2)) \<noteq> h (real (Suc (2 * m)) / 2 ^ Suc n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1798 | using a_less_b [of "4*m + 3" "n+2", THEN conjunct1] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1799 | using a43 [of "Suc n" m] b43 [of "Suc n" m] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1800 | using alec [of "2*m+1" "Suc n"] cleb [of "2*m+1" "Suc n"] a_ge_0 [of "2*m+1" "Suc n"] b_le_1 [of "2*m+1" "Suc n"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1801 | apply (simp add: fc_eq [symmetric] c_def del: power_Suc) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1802 | apply (simp only: add.commute [of 1] c_fold R_def [symmetric]) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1803 | apply (rule right_neq) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1804 | using Rless apply (simp add: R_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1805 | apply (rule midR_le, auto) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1806 | done | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1807 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1808 | then show ?thesis by (metis that) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1809 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1810 | have m_div: "0 < m / 2^n" "m / 2^n < 1" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70136diff
changeset | 1811 | using m by (auto simp: field_split_simps) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1812 |       have closure0m: "{0..m / 2^n} = closure ({0<..< m / 2^n} \<inter> (\<Union>k m. {real m / 2 ^ k}))"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1813 | by (subst closure_dyadic_rationals_in_convex_set_pos_1, simp_all add: not_le m) | 
| 72531 | 1814 | have "2^n > m" | 
| 1815 | by (simp add: m(2) not_le) | |
| 1816 |       then have closurem1: "{m / 2^n .. 1} = closure ({m / 2^n <..< 1} \<inter> (\<Union>k m. {real m / 2 ^ k}))"
 | |
| 1817 | using closure_dyadic_rationals_in_convex_set_pos_1 m_div(1) by fastforce | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1818 |       have cont_h': "continuous_on (closure ({u<..<v} \<inter> (\<Union>k m. {real m / 2 ^ k}))) h"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1819 | if "0 \<le> u" "v \<le> 1" for u v | 
| 72531 | 1820 | using that by (intro continuous_on_subset [OF cont_h] closure_minimal [OF subsetI]) auto | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1821 |       have closed_f': "closed (f ` {u..v})" if "0 \<le> u" "v \<le> 1" for u v
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1822 | by (metis compact_continuous_image cont_f compact_interval atLeastatMost_subset_iff | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1823 | compact_imp_closed continuous_on_subset that) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1824 | have less_2I: "\<And>k i. real i / 2 ^ k < 1 \<Longrightarrow> i < 2 ^ k" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1825 | by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1826 |       have "h ` ({0<..<m / 2 ^ n} \<inter> (\<Union>q p. {real p / 2 ^ q})) \<subseteq> f ` {0..c (m / 2 ^ n)}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1827 | proof clarsimp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1828 | fix p q | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1829 | assume p: "0 < real p / 2 ^ q" "real p / 2 ^ q < real m / 2 ^ n" | 
| 72531 | 1830 | then have [simp]: "0 < p" | 
| 1831 | by (simp add: field_split_simps) | |
| 1832 | have [simp]: "p < 2 ^ q" | |
| 1833 | by (blast intro: p less_2I m_div less_trans) | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1834 |         have "f (c (real p / 2 ^ q)) \<in> f ` {0..c (real m / 2 ^ n)}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1835 | by (auto simp: clec p m) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1836 |         then show "h (real p / 2 ^ q) \<in> f ` {0..c (real m / 2 ^ n)}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1837 | by (simp add: h_eq) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1838 | qed | 
| 72531 | 1839 |       with m_div have "h ` {0 .. m / 2^n} \<subseteq> f ` {0 .. c(m / 2^n)}"
 | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1840 | apply (subst closure0m) | 
| 72531 | 1841 | by (rule image_closure_subset [OF cont_h' closed_f']) auto | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1842 |       then have hx1: "h x1 \<in> f ` {0 .. c(m / 2^n)}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1843 | using x12 less.prems(1) by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1844 | then obtain t1 where t1: "h x1 = f t1" "0 \<le> t1" "t1 \<le> c (m / 2 ^ n)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1845 | by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1846 |       have "h ` ({m / 2 ^ n<..<1} \<inter> (\<Union>q p. {real p / 2 ^ q})) \<subseteq> f ` {c (m / 2 ^ n)..1}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1847 | proof clarsimp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1848 | fix p q | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1849 | assume p: "real m / 2 ^ n < real p / 2 ^ q" and [simp]: "p < 2 ^ q" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1850 | then have [simp]: "0 < p" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1851 | using gr_zeroI m_div by fastforce | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1852 |         have "f (c (real p / 2 ^ q)) \<in> f ` {c (m / 2 ^ n)..1}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1853 | by (auto simp: clec p m) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1854 |         then show "h (real p / 2 ^ q) \<in> f ` {c (real m / 2 ^ n)..1}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1855 | by (simp add: h_eq) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1856 | qed | 
| 72531 | 1857 |       with m have "h ` {m / 2^n .. 1} \<subseteq> f ` {c(m / 2^n) .. 1}"
 | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1858 | apply (subst closurem1) | 
| 72531 | 1859 | by (rule image_closure_subset [OF cont_h' closed_f']) auto | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1860 |       then have hx2: "h x2 \<in> f ` {c(m / 2^n)..1}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1861 | using x12 less.prems by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1862 | then obtain t2 where t2: "h x2 = f t2" "c (m / 2 ^ n) \<le> t2" "t2 \<le> 1" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1863 | by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1864 | with t1 less neq have False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1865 | using conn [of "h x2", unfolded is_interval_connected_1 [symmetric] is_interval_1, rule_format, of t1 t2 "c(m / 2^n)"] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1866 | by (simp add: h_eq m) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1867 | then show ?case by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1868 | qed auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1869 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1870 | by (auto simp: inj_on_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1871 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1872 |   ultimately have "{0..1::real} homeomorphic f ` {0..1}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1873 | using homeomorphic_compact [OF _ cont_h] by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1874 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1875 | using homeomorphic_sym by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1876 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1877 | |
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1878 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 1879 | theorem path_contains_arc: | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1880 |   fixes p :: "real \<Rightarrow> 'a::{complete_space,real_normed_vector}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1881 | assumes "path p" and a: "pathstart p = a" and b: "pathfinish p = b" and "a \<noteq> b" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1882 | obtains q where "arc q" "path_image q \<subseteq> path_image p" "pathstart q = a" "pathfinish q = b" | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 1883 | proof - | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1884 |   have ucont_p: "uniformly_continuous_on {0..1} p"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1885 | using \<open>path p\<close> unfolding path_def | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1886 | by (metis compact_Icc compact_uniformly_continuous) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1887 |   define \<phi> where "\<phi> \<equiv> \<lambda>S. S \<subseteq> {0..1} \<and> 0 \<in> S \<and> 1 \<in> S \<and>
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1888 |                            (\<forall>x \<in> S. \<forall>y \<in> S. open_segment x y \<inter> S = {} \<longrightarrow> p x = p y)"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1889 | obtain T where "closed T" "\<phi> T" and T: "\<And>U. \<lbrakk>closed U; \<phi> U\<rbrakk> \<Longrightarrow> \<not> (U \<subset> T)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1890 |   proof (rule Brouwer_reduction_theorem_gen [of "{0..1}" \<phi>])
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1891 |     have *: "{x<..<y} \<inter> {0..1} = {x<..<y}" if "0 \<le> x" "y \<le> 1" "x \<le> y" for x y::real
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1892 | using that by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1893 |     show "\<phi> {0..1}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1894 | by (auto simp: \<phi>_def open_segment_eq_real_ivl *) | 
| 69313 | 1895 | show "\<phi> (\<Inter>(F ` UNIV))" | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1896 | if "\<And>n. closed (F n)" and \<phi>: "\<And>n. \<phi> (F n)" and Fsub: "\<And>n. F (Suc n) \<subseteq> F n" for F | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1897 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1898 |       have F01: "\<And>n. F n \<subseteq> {0..1} \<and> 0 \<in> F n \<and> 1 \<in> F n"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1899 |         and peq: "\<And>n x y. \<lbrakk>x \<in> F n; y \<in> F n; open_segment x y \<inter> F n = {}\<rbrakk> \<Longrightarrow> p x = p y"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1900 | by (metis \<phi> \<phi>_def)+ | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1901 |       have pqF: False if "\<forall>u. x \<in> F u" "\<forall>x. y \<in> F x" "open_segment x y \<inter> (\<Inter>x. F x) = {}" and neg: "p x \<noteq> p y"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1902 | for x y | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1903 | using that | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1904 | proof (induction x y rule: linorder_class.linorder_less_wlog) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1905 | case (less x y) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1906 |         have xy: "x \<in> {0..1}" "y \<in> {0..1}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1907 | by (metis less.prems subsetCE F01)+ | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1908 | have "norm(p x - p y) / 2 > 0" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1909 | using less by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1910 | then obtain e where "e > 0" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1911 |           and e: "\<And>u v. \<lbrakk>u \<in> {0..1}; v \<in> {0..1}; dist v u < e\<rbrakk> \<Longrightarrow> dist (p v) (p u) < norm(p x - p y) / 2"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1912 | by (metis uniformly_continuous_onE [OF ucont_p]) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1913 | have minxy: "min e (y - x) < (y - x) * (3 / 2)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1914 | by (subst min_less_iff_disj) (simp add: less) | 
| 72531 | 1915 | define w where "w \<equiv> x + (min e (y - x) / 3)" | 
| 1916 | define z where "z \<equiv>y - (min e (y - x) / 3)" | |
| 1917 |         have "w < z" and w: "w \<in> {x<..<y}" and z: "z \<in> {x<..<y}"
 | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1918 | and wxe: "norm(w - x) < e" and zye: "norm(z - y) < e" | 
| 72531 | 1919 | using minxy \<open>0 < e\<close> less unfolding w_def z_def by auto | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1920 | have Fclo: "\<And>T. T \<in> range F \<Longrightarrow> closed T" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1921 | by (metis \<open>\<And>n. closed (F n)\<close> image_iff) | 
| 69313 | 1922 |         have eq: "{w..z} \<inter> \<Inter>(F ` UNIV) = {}"
 | 
| 72531 | 1923 | using less w z by (simp add: open_segment_eq_real_ivl disjoint_iff) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1924 |         then obtain K where "finite K" and K: "{w..z} \<inter> (\<Inter> (F ` K)) = {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1925 | by (metis finite_subset_image compact_imp_fip [OF compact_interval Fclo]) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1926 |         then have "K \<noteq> {}"
 | 
| 69313 | 1927 |           using \<open>w < z\<close> \<open>{w..z} \<inter> \<Inter>(F ` K) = {}\<close> by auto
 | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1928 | define n where "n \<equiv> Max K" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1929 |         have "n \<in> K" unfolding n_def by (metis \<open>K \<noteq> {}\<close> \<open>finite K\<close> Max_in)
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1930 | have "F n \<subseteq> \<Inter> (F ` K)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1931 |           unfolding n_def by (metis Fsub Max_ge \<open>K \<noteq> {}\<close> \<open>finite K\<close> cINF_greatest lift_Suc_antimono_le)
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1932 |         with K have wzF_null: "{w..z} \<inter> F n = {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1933 | by (metis disjoint_iff_not_equal subset_eq) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1934 |         obtain u where u: "u \<in> F n" "u \<in> {x..w}" "({u..w} - {u}) \<inter> F n = {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1935 | proof (cases "w \<in> F n") | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1936 | case True | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1937 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1938 | by (metis wzF_null \<open>w < z\<close> atLeastAtMost_iff disjoint_iff_not_equal less_eq_real_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1939 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1940 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1941 |           obtain u where "u \<in> F n" "u \<in> {x..w}" "{u<..<w} \<inter> F n = {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1942 |           proof (rule segment_to_point_exists [of "F n \<inter> {x..w}" w])
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1943 |             show "closed (F n \<inter> {x..w})"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1944 | by (metis \<open>\<And>n. closed (F n)\<close> closed_Int closed_real_atLeastAtMost) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1945 |             show "F n \<inter> {x..w} \<noteq> {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1946 | by (metis atLeastAtMost_iff disjoint_iff_not_equal greaterThanLessThan_iff less.prems(1) less_eq_real_def w) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1947 | qed (auto simp: open_segment_eq_real_ivl intro!: that) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1948 | with False show thesis | 
| 72531 | 1949 | by (auto simp add: disjoint_iff less_eq_real_def intro!: that) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1950 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1951 |         obtain v where v: "v \<in> F n" "v \<in> {z..y}" "({z..v} - {v}) \<inter> F n = {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1952 | proof (cases "z \<in> F n") | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1953 | case True | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1954 |           have "z \<in> {w..z}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1955 | using \<open>w < z\<close> by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1956 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1957 | by (metis wzF_null Int_iff True empty_iff) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1958 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1959 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1960 | show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1961 |           proof (rule segment_to_point_exists [of "F n \<inter> {z..y}" z])
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1962 |             show "closed (F n \<inter> {z..y})"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1963 | by (metis \<open>\<And>n. closed (F n)\<close> closed_Int closed_atLeastAtMost) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1964 |             show "F n \<inter> {z..y} \<noteq> {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1965 | by (metis atLeastAtMost_iff disjoint_iff_not_equal greaterThanLessThan_iff less.prems(2) less_eq_real_def z) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1966 |             show "\<And>b. \<lbrakk>b \<in> F n \<inter> {z..y}; open_segment z b \<inter> (F n \<inter> {z..y}) = {}\<rbrakk> \<Longrightarrow> thesis"
 | 
| 72531 | 1967 | proof | 
| 1968 |               show "\<And>b. \<lbrakk>b \<in> F n \<inter> {z..y}; open_segment z b \<inter> (F n \<inter> {z..y}) = {}\<rbrakk> \<Longrightarrow> ({z..b} - {b}) \<inter> F n = {}"
 | |
| 1969 | using False by (auto simp: open_segment_eq_real_ivl less_eq_real_def) | |
| 1970 | qed auto | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1971 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1972 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1973 |         obtain u v where "u \<in> {0..1}" "v \<in> {0..1}" "norm(u - x) < e" "norm(v - y) < e" "p u = p v"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1974 | proof | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1975 |           show "u \<in> {0..1}" "v \<in> {0..1}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1976 | by (metis F01 \<open>u \<in> F n\<close> \<open>v \<in> F n\<close> subsetD)+ | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1977 | show "norm(u - x) < e" "norm (v - y) < e" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1978 |             using \<open>u \<in> {x..w}\<close> \<open>v \<in> {z..y}\<close> atLeastAtMost_iff real_norm_def wxe zye by auto
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1979 | show "p u = p v" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1980 | proof (rule peq) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1981 | show "u \<in> F n" "v \<in> F n" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1982 | by (auto simp: u v) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1983 | have "False" if "\<xi> \<in> F n" "u < \<xi>" "\<xi> < v" for \<xi> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1984 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1985 |               have "\<xi> \<notin> {z..v}"
 | 
| 71633 | 1986 | by (metis DiffI disjoint_iff_not_equal less_irrefl singletonD that(1,3) v(3)) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1987 |               moreover have "\<xi> \<notin> {w..z} \<inter> F n"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1988 | by (metis equals0D wzF_null) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1989 |               ultimately have "\<xi> \<in> {u..w}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1990 | using that by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1991 | then show ?thesis | 
| 71633 | 1992 | by (metis DiffI disjoint_iff_not_equal less_eq_real_def not_le singletonD that(1,2) u(3)) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1993 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1994 | moreover | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1995 | have "\<lbrakk>\<xi> \<in> F n; v < \<xi>; \<xi> < u\<rbrakk> \<Longrightarrow> False" for \<xi> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1996 |               using \<open>u \<in> {x..w}\<close> \<open>v \<in> {z..y}\<close> \<open>w < z\<close> by simp
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1997 | ultimately | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1998 |             show "open_segment u v \<inter> F n = {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1999 | by (force simp: open_segment_eq_real_ivl) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2000 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2001 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2002 | then show ?case | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2003 | using e [of x u] e [of y v] xy | 
| 72531 | 2004 | by (metis dist_norm dist_triangle_half_r order_less_irrefl) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2005 | qed (auto simp: open_segment_commute) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2006 | show ?thesis | 
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
72531diff
changeset | 2007 | unfolding \<phi>_def by (metis (no_types, opaque_lifting) INT_I Inf_lower2 rangeI that(3) F01 subsetCE pqF) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2008 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2009 |     show "closed {0..1::real}" by auto
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2010 | qed (meson \<phi>_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2011 |   then have "T \<subseteq> {0..1}" "0 \<in> T" "1 \<in> T"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2012 |     and peq: "\<And>x y. \<lbrakk>x \<in> T; y \<in> T; open_segment x y \<inter> T = {}\<rbrakk> \<Longrightarrow> p x = p y"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2013 | unfolding \<phi>_def by metis+ | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2014 |   then have "T \<noteq> {}" by auto
 | 
| 67613 | 2015 |   define h where "h \<equiv> \<lambda>x. p(SOME y. y \<in> T \<and> open_segment x y \<inter> T = {})"
 | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2016 |   have "p y = p z" if "y \<in> T" "z \<in> T" and xyT: "open_segment x y \<inter> T = {}" and xzT: "open_segment x z \<inter> T = {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2017 | for x y z | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2018 | proof (cases "x \<in> T") | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2019 | case True | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2020 | with that show ?thesis by (metis \<open>\<phi> T\<close> \<phi>_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2021 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2022 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2023 |     have "insert x (open_segment x y \<union> open_segment x z) \<inter> T = {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2024 | by (metis False Int_Un_distrib2 Int_insert_left Un_empty_right xyT xzT) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2025 | moreover have "open_segment y z \<inter> T \<subseteq> insert x (open_segment x y \<union> open_segment x z) \<inter> T" | 
| 72531 | 2026 | by (auto simp: open_segment_eq_real_ivl) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2027 |     ultimately have "open_segment y z \<inter> T = {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2028 | by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2029 | with that peq show ?thesis by metis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2030 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2031 |   then have h_eq_p_gen: "h x = p y" if "y \<in> T" "open_segment x y \<inter> T = {}" for x y
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2032 | using that unfolding h_def | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2033 | by (metis (mono_tags, lifting) some_eq_ex) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2034 | then have h_eq_p: "\<And>x. x \<in> T \<Longrightarrow> h x = p x" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2035 | by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2036 |   have disjoint: "\<And>x. \<exists>y. y \<in> T \<and> open_segment x y \<inter> T = {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2037 |     by (meson \<open>T \<noteq> {}\<close> \<open>closed T\<close> segment_to_point_exists)
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2038 |   have heq: "h x = h x'" if "open_segment x x' \<inter> T = {}" for x x'
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2039 | proof (cases "x \<in> T \<or> x' \<in> T") | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2040 | case True | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2041 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2042 | by (metis h_eq_p h_eq_p_gen open_segment_commute that) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2043 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2044 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2045 |     obtain y y' where "y \<in> T" "open_segment x y \<inter> T = {}" "h x = p y"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2046 |       "y' \<in> T" "open_segment x' y' \<inter> T = {}" "h x' = p y'"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2047 | by (meson disjoint h_eq_p_gen) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2048 | moreover have "open_segment y y' \<subseteq> (insert x (insert x' (open_segment x y \<union> open_segment x' y' \<union> open_segment x x')))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2049 | by (auto simp: open_segment_eq_real_ivl) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2050 | ultimately show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2051 | using False that by (fastforce simp add: h_eq_p intro!: peq) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2052 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2053 |   have "h ` {0..1} homeomorphic {0..1::real}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2054 | proof (rule homeomorphic_monotone_image_interval) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2055 |     show "continuous_on {0..1} h"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2056 | proof (clarsimp simp add: continuous_on_iff) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2057 | fix u \<epsilon>::real | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2058 | assume "0 < \<epsilon>" "0 \<le> u" "u \<le> 1" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2059 |       then obtain \<delta> where "\<delta> > 0" and \<delta>: "\<And>v. v \<in> {0..1} \<Longrightarrow> dist v u < \<delta> \<longrightarrow> dist (p v) (p u) < \<epsilon> / 2"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2060 | using ucont_p [unfolded uniformly_continuous_on_def] | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2061 | by (metis atLeastAtMost_iff half_gt_zero_iff) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2062 |       then have "dist (h v) (h u) < \<epsilon>" if "v \<in> {0..1}" "dist v u < \<delta>" for v
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2063 |       proof (cases "open_segment u v \<inter> T = {}")
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2064 | case True | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2065 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2066 | using \<open>0 < \<epsilon>\<close> heq by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2067 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2068 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2069 |         have uvT: "closed (closed_segment u v \<inter> T)" "closed_segment u v \<inter> T \<noteq> {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2070 | using False open_closed_segment by (auto simp: \<open>closed T\<close> closed_Int) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2071 |         obtain w where "w \<in> T" and w: "w \<in> closed_segment u v" "open_segment u w \<inter> T = {}"
 | 
| 72531 | 2072 | proof (rule segment_to_point_exists [OF uvT]) | 
| 2073 | fix b | |
| 2074 |           assume "b \<in> closed_segment u v \<inter> T" "open_segment u b \<inter> (closed_segment u v \<inter> T) = {}"
 | |
| 2075 | then show thesis | |
| 2076 | by (metis IntD1 IntD2 ends_in_segment(1) inf.orderE inf_assoc subset_oc_segment that) | |
| 2077 | qed | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2078 | then have puw: "dist (p u) (p w) < \<epsilon> / 2" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2079 |           by (metis (no_types) \<open>T \<subseteq> {0..1}\<close> \<open>dist v u < \<delta>\<close> \<delta> dist_commute dist_in_closed_segment le_less_trans subsetCE)
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2080 |         obtain z where "z \<in> T" and z: "z \<in> closed_segment u v" "open_segment v z \<inter> T = {}"
 | 
| 72531 | 2081 | proof (rule segment_to_point_exists [OF uvT]) | 
| 2082 | fix b | |
| 2083 |           assume "b \<in> closed_segment u v \<inter> T" "open_segment v b \<inter> (closed_segment u v \<inter> T) = {}"
 | |
| 2084 | then show thesis | |
| 2085 | by (metis IntD1 IntD2 ends_in_segment(2) inf.orderE inf_assoc subset_oc_segment that) | |
| 2086 | qed | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2087 | then have "dist (p u) (p z) < \<epsilon> / 2" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2088 |           by (metis \<open>T \<subseteq> {0..1}\<close> \<open>dist v u < \<delta>\<close> \<delta> dist_commute dist_in_closed_segment le_less_trans subsetCE)
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2089 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2090 | using puw by (metis (no_types) \<open>w \<in> T\<close> \<open>z \<in> T\<close> dist_commute dist_triangle_half_l h_eq_p_gen w(2) z(2)) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2091 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2092 |       with \<open>0 < \<delta>\<close> show "\<exists>\<delta>>0. \<forall>v\<in>{0..1}. dist v u < \<delta> \<longrightarrow> dist (h v) (h u) < \<epsilon>" by blast
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2093 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2094 |     show "connected ({0..1} \<inter> h -` {z})" for z
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2095 | proof (clarsimp simp add: connected_iff_connected_component) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2096 | fix u v | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2097 | assume huv_eq: "h v = h u" and uv: "0 \<le> u" "u \<le> 1" "0 \<le> v" "v \<le> 1" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2098 |       have "\<exists>T. connected T \<and> T \<subseteq> {0..1} \<and> T \<subseteq> h -` {h u} \<and> u \<in> T \<and> v \<in> T"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2099 | proof (intro exI conjI) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2100 | show "connected (closed_segment u v)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2101 | by simp | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2102 |         show "closed_segment u v \<subseteq> {0..1}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2103 | by (simp add: uv closed_segment_eq_real_ivl) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2104 | have pxy: "p x = p y" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2105 |           if "T \<subseteq> {0..1}" "0 \<in> T" "1 \<in> T" "x \<in> T" "y \<in> T"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2106 |           and disjT: "open_segment x y \<inter> (T - open_segment u v) = {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2107 | and xynot: "x \<notin> open_segment u v" "y \<notin> open_segment u v" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2108 | for x y | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2109 |         proof (cases "open_segment x y \<inter> open_segment u v = {}")
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2110 | case True | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2111 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2112 | by (metis Diff_Int_distrib Diff_empty peq disjT \<open>x \<in> T\<close> \<open>y \<in> T\<close>) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2113 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2114 | case False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2115 | then have "open_segment x u \<union> open_segment y v \<subseteq> open_segment x y - open_segment u v \<or> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2116 | open_segment y u \<union> open_segment x v \<subseteq> open_segment x y - open_segment u v" (is "?xuyv \<or> ?yuxv") | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2117 | using xynot by (fastforce simp add: open_segment_eq_real_ivl not_le not_less split: if_split_asm) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2118 | then show "p x = p y" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2119 | proof | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2120 | assume "?xuyv" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2121 |             then have "open_segment x u \<inter> T = {}" "open_segment y v \<inter> T = {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2122 | using disjT by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2123 | then have "h x = h y" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2124 | using heq huv_eq by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2125 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2126 | using h_eq_p \<open>x \<in> T\<close> \<open>y \<in> T\<close> by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2127 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2128 | assume "?yuxv" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2129 |             then have "open_segment y u \<inter> T = {}" "open_segment x v \<inter> T = {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2130 | using disjT by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2131 | then have "h x = h y" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2132 | using heq [of y u] heq [of x v] huv_eq by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2133 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2134 | using h_eq_p \<open>x \<in> T\<close> \<open>y \<in> T\<close> by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2135 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2136 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2137 | have "\<not> T - open_segment u v \<subset> T" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2138 | proof (rule T) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2139 | show "closed (T - open_segment u v)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2140 | by (simp add: closed_Diff [OF \<open>closed T\<close>] open_segment_eq_real_ivl) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2141 | have "0 \<notin> open_segment u v" "1 \<notin> open_segment u v" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2142 | using open_segment_eq_real_ivl uv by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2143 | then show "\<phi> (T - open_segment u v)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2144 |             using \<open>T \<subseteq> {0..1}\<close> \<open>0 \<in> T\<close> \<open>1 \<in> T\<close>
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2145 | by (auto simp: \<phi>_def) (meson peq pxy) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2146 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2147 |         then have "open_segment u v \<inter> T = {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2148 | by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2149 |         then show "closed_segment u v \<subseteq> h -` {h u}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2150 | by (force intro: heq simp: open_segment_eq_real_ivl closed_segment_eq_real_ivl split: if_split_asm)+ | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2151 | qed auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2152 |       then show "connected_component ({0..1} \<inter> h -` {h u}) u v"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2153 | by (simp add: connected_component_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2154 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2155 | show "h 1 \<noteq> h 0" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2156 | by (metis \<open>\<phi> T\<close> \<phi>_def a \<open>a \<noteq> b\<close> b h_eq_p pathfinish_def pathstart_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2157 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2158 | then obtain f and g :: "real \<Rightarrow> 'a" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2159 |     where gfeq: "(\<forall>x\<in>h ` {0..1}. (g(f x) = x))" and fhim: "f ` h ` {0..1} = {0..1}" and contf: "continuous_on (h ` {0..1}) f"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2160 |       and fgeq: "(\<forall>y\<in>{0..1}. (f(g y) = y))" and pag: "path_image g = h ` {0..1}" and contg: "continuous_on {0..1} g"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2161 | by (auto simp: homeomorphic_def homeomorphism_def path_image_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2162 | then have "arc g" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2163 | by (metis arc_def path_def inj_on_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2164 |   obtain u v where "u \<in> {0..1}" "a = g u" "v \<in> {0..1}" "b = g v"
 | 
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
72531diff
changeset | 2165 | by (metis (mono_tags, opaque_lifting) \<open>\<phi> T\<close> \<phi>_def a b fhim gfeq h_eq_p imageI path_image_def pathfinish_def pathfinish_in_path_image pathstart_def pathstart_in_path_image) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2166 | then have "a \<in> path_image g" "b \<in> path_image g" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2167 | using path_image_def by blast+ | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2168 | have ph: "path_image h \<subseteq> path_image p" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2169 |     by (metis image_mono image_subset_iff path_image_def disjoint h_eq_p_gen \<open>T \<subseteq> {0..1}\<close>)
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2170 | show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2171 | proof | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2172 | show "pathstart (subpath u v g) = a" "pathfinish (subpath u v g) = b" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2173 | by (simp_all add: \<open>a = g u\<close> \<open>b = g v\<close>) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2174 | show "path_image (subpath u v g) \<subseteq> path_image p" | 
| 71633 | 2175 |       by (metis \<open>u \<in> {0..1}\<close> \<open>v \<in> {0..1}\<close> order_trans pag path_image_def path_image_subpath_subset ph)
 | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2176 | show "arc (subpath u v g)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2177 |       using \<open>arc g\<close> \<open>a = g u\<close> \<open>b = g v\<close> \<open>u \<in> {0..1}\<close> \<open>v \<in> {0..1}\<close> arc_subpath_arc \<open>a \<noteq> b\<close> by blast
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2178 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2179 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2180 | |
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2181 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 2182 | corollary path_connected_arcwise: | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2183 |   fixes S :: "'a::{complete_space,real_normed_vector} set"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2184 | shows "path_connected S \<longleftrightarrow> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2185 | (\<forall>x \<in> S. \<forall>y \<in> S. x \<noteq> y \<longrightarrow> (\<exists>g. arc g \<and> path_image g \<subseteq> S \<and> pathstart g = x \<and> pathfinish g = y))" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2186 | (is "?lhs = ?rhs") | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 2187 | proof (intro iffI impI ballI) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2188 | fix x y | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2189 | assume "path_connected S" "x \<in> S" "y \<in> S" "x \<noteq> y" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2190 | then obtain p where p: "path p" "path_image p \<subseteq> S" "pathstart p = x" "pathfinish p = y" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2191 | by (force simp: path_connected_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2192 | then show "\<exists>g. arc g \<and> path_image g \<subseteq> S \<and> pathstart g = x \<and> pathfinish g = y" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2193 | by (metis \<open>x \<noteq> y\<close> order_trans path_contains_arc) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2194 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2195 | assume R [rule_format]: ?rhs | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2196 | show ?lhs | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2197 | unfolding path_connected_def | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2198 | proof (intro ballI) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2199 | fix x y | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2200 | assume "x \<in> S" "y \<in> S" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2201 | show "\<exists>g. path g \<and> path_image g \<subseteq> S \<and> pathstart g = x \<and> pathfinish g = y" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2202 | proof (cases "x = y") | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2203 | case True with \<open>x \<in> S\<close> path_component_def path_component_refl show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2204 | by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2205 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2206 | case False with R [OF \<open>x \<in> S\<close> \<open>y \<in> S\<close>] show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2207 | by (auto intro: arc_imp_path) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2208 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2209 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2210 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2211 | |
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2212 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 2213 | corollary arc_connected_trans: | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2214 |   fixes g :: "real \<Rightarrow> 'a::{complete_space,real_normed_vector}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2215 | assumes "arc g" "arc h" "pathfinish g = pathstart h" "pathstart g \<noteq> pathfinish h" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2216 | obtains i where "arc i" "path_image i \<subseteq> path_image g \<union> path_image h" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2217 | "pathstart i = pathstart g" "pathfinish i = pathfinish h" | 
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
72531diff
changeset | 2218 | by (metis (no_types, opaque_lifting) arc_imp_path assms path_contains_arc path_image_join path_join pathfinish_join pathstart_join) | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2219 | |
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2220 | |
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2221 | |
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2222 | |
| 69683 | 2223 | subsection\<open>Accessibility of frontier points\<close> | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2224 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 2225 | lemma dense_accessible_frontier_points: | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2226 |   fixes S :: "'a::{complete_space,real_normed_vector} set"
 | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69745diff
changeset | 2227 |   assumes "open S" and opeSV: "openin (top_of_set (frontier S)) V" and "V \<noteq> {}"
 | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2228 |   obtains g where "arc g" "g ` {0..<1} \<subseteq> S" "pathstart g \<in> S" "pathfinish g \<in> V"
 | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 2229 | proof - | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2230 | obtain z where "z \<in> V" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2231 |     using \<open>V \<noteq> {}\<close> by auto
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2232 | then obtain r where "r > 0" and r: "ball z r \<inter> frontier S \<subseteq> V" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2233 | by (metis openin_contains_ball opeSV) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2234 | then have "z \<in> frontier S" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2235 | using \<open>z \<in> V\<close> opeSV openin_contains_ball by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2236 | then have "z \<in> closure S" "z \<notin> S" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2237 | by (simp_all add: frontier_def assms interior_open) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2238 | with \<open>r > 0\<close> have "infinite (S \<inter> ball z r)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2239 | by (auto simp: closure_def islimpt_eq_infinite_ball) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2240 | then obtain y where "y \<in> S" and y: "y \<in> ball z r" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2241 | using infinite_imp_nonempty by force | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2242 | then have "y \<notin> frontier S" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2243 | by (meson \<open>open S\<close> disjoint_iff_not_equal frontier_disjoint_eq) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2244 | have "y \<noteq> z" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2245 | using \<open>y \<in> S\<close> \<open>z \<notin> S\<close> by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2246 | have "path_connected(ball z r)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2247 | by (simp add: convex_imp_path_connected) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2248 | with y \<open>r > 0\<close> obtain g where "arc g" and pig: "path_image g \<subseteq> ball z r" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2249 | and g: "pathstart g = y" "pathfinish g = z" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2250 | using \<open>y \<noteq> z\<close> by (force simp: path_connected_arcwise) | 
| 72531 | 2251 |   have "continuous_on {0..1} g"
 | 
| 2252 | using \<open>arc g\<close> arc_imp_path path_def by blast | |
| 2253 |   then have "compact (g -` frontier S \<inter> {0..1})"
 | |
| 2254 | by (simp add: bounded_Int closed_Diff closed_vimage_Int compact_eq_bounded_closed) | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2255 |   moreover have "g -` frontier S \<inter> {0..1} \<noteq> {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2256 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2257 |     have "\<exists>r. r \<in> g -` frontier S \<and> r \<in> {0..1}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2258 | by (metis \<open>z \<in> frontier S\<close> g(2) imageE path_image_def pathfinish_in_path_image vimageI2) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2259 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2260 | by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2261 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2262 | ultimately obtain t where gt: "g t \<in> frontier S" and "0 \<le> t" "t \<le> 1" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2263 | and t: "\<And>u. \<lbrakk>g u \<in> frontier S; 0 \<le> u; u \<le> 1\<rbrakk> \<Longrightarrow> t \<le> u" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2264 | by (force simp: dest!: compact_attains_inf) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2265 | moreover have "t \<noteq> 0" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2266 | by (metis \<open>y \<notin> frontier S\<close> g(1) gt pathstart_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2267 | ultimately have t01: "0 < t" "t \<le> 1" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2268 | by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2269 | have "V \<subseteq> frontier S" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2270 | using opeSV openin_contains_ball by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2271 | show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2272 | proof | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2273 | show "arc (subpath 0 t g)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2274 | by (simp add: \<open>0 \<le> t\<close> \<open>t \<le> 1\<close> \<open>arc g\<close> \<open>t \<noteq> 0\<close> arc_subpath_arc) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2275 | have "g 0 \<in> S" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2276 | by (metis \<open>y \<in> S\<close> g(1) pathstart_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2277 | then show "pathstart (subpath 0 t g) \<in> S" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2278 | by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2279 | have "g t \<in> V" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2280 | by (metis IntI atLeastAtMost_iff gt image_eqI path_image_def pig r subsetCE \<open>0 \<le> t\<close> \<open>t \<le> 1\<close>) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2281 | then show "pathfinish (subpath 0 t g) \<in> V" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2282 | by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2283 |     then have "inj_on (subpath 0 t g) {0..1}"
 | 
| 72531 | 2284 | using t01 \<open>arc (subpath 0 t g)\<close> arc_imp_inj_on by blast | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2285 |     then have "subpath 0 t g ` {0..<1} \<subseteq> subpath 0 t g ` {0..1} - {subpath 0 t g 1}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2286 | by (force simp: dest: inj_onD) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2287 |     moreover have False if "subpath 0 t g ` ({0..<1}) - S \<noteq> {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2288 | proof - | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2289 |       have contg: "continuous_on {0..1} g"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2290 | using \<open>arc g\<close> by (auto simp: arc_def path_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2291 |       have "subpath 0 t g ` {0..<1} \<inter> frontier S \<noteq> {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2292 | proof (rule connected_Int_frontier [OF _ _ that]) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2293 |         show "connected (subpath 0 t g ` {0..<1})"
 | 
| 72531 | 2294 | proof (rule connected_continuous_image) | 
| 2295 |           show "continuous_on {0..<1} (subpath 0 t g)"
 | |
| 2296 | by (meson \<open>arc (subpath 0 t g)\<close> arc_def atLeastLessThan_subseteq_atLeastAtMost_iff continuous_on_subset order_refl path_def) | |
| 2297 | qed auto | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2298 |         show "subpath 0 t g ` {0..<1} \<inter> S \<noteq> {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2299 | using \<open>y \<in> S\<close> g(1) by (force simp: subpath_def image_def pathstart_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2300 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2301 |       then obtain x where "x \<in> subpath 0 t g ` {0..<1}" "x \<in> frontier S"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2302 | by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2303 | with t01 \<open>0 \<le> t\<close> mult_le_one t show False | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2304 | by (fastforce simp: subpath_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2305 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2306 |     then have "subpath 0 t g ` {0..1} - {subpath 0 t g 1} \<subseteq> S"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2307 | using subsetD by fastforce | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2308 |     ultimately  show "subpath 0 t g ` {0..<1} \<subseteq> S"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2309 | by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2310 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2311 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2312 | |
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2313 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 2314 | lemma dense_accessible_frontier_points_connected: | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2315 |   fixes S :: "'a::{complete_space,real_normed_vector} set"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2316 |   assumes "open S" "connected S" "x \<in> S" "V \<noteq> {}"
 | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69745diff
changeset | 2317 | and ope: "openin (top_of_set (frontier S)) V" | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2318 |   obtains g where "arc g" "g ` {0..<1} \<subseteq> S" "pathstart g = x" "pathfinish g \<in> V"
 | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 2319 | proof - | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2320 | have "V \<subseteq> frontier S" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2321 | using ope openin_imp_subset by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2322 | with \<open>open S\<close> \<open>x \<in> S\<close> have "x \<notin> V" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2323 | using interior_open by (auto simp: frontier_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2324 |   obtain g where "arc g" and g: "g ` {0..<1} \<subseteq> S" "pathstart g \<in> S" "pathfinish g \<in> V"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2325 |     by (metis dense_accessible_frontier_points [OF \<open>open S\<close> ope \<open>V \<noteq> {}\<close>])
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2326 | then have "path_connected S" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2327 | by (simp add: assms connected_open_path_connected) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2328 | with \<open>pathstart g \<in> S\<close> \<open>x \<in> S\<close> have "path_component S x (pathstart g)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2329 | by (simp add: path_connected_component) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2330 | then obtain f where "path f" and f: "path_image f \<subseteq> S" "pathstart f = x" "pathfinish f = pathstart g" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2331 | by (auto simp: path_component_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2332 | then have "path (f +++ g)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2333 | by (simp add: \<open>arc g\<close> arc_imp_path) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2334 | then obtain h where "arc h" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2335 | and h: "path_image h \<subseteq> path_image (f +++ g)" "pathstart h = x" "pathfinish h = pathfinish g" | 
| 72531 | 2336 | using path_contains_arc [of "f +++ g" x "pathfinish g"] \<open>x \<notin> V\<close> \<open>pathfinish g \<in> V\<close> f | 
| 2337 | by (metis pathfinish_join pathstart_join) | |
| 2338 | have "path_image h \<subseteq> path_image f \<union> path_image g" | |
| 2339 | using h(1) path_image_join_subset by auto | |
| 2340 |   then have "h ` {0..1} - {h 1} \<subseteq> S"
 | |
| 2341 | using f g h | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2342 | apply (simp add: path_image_def pathfinish_def subset_iff image_def Bex_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2343 | by (metis le_less) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2344 |   then have "h ` {0..<1} \<subseteq> S"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2345 | using \<open>arc h\<close> by (force simp: arc_def dest: inj_onD) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2346 | then show thesis | 
| 72531 | 2347 | using \<open>arc h\<close> g(3) h that by presburger | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2348 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2349 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 2350 | lemma dense_access_fp_aux: | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2351 |   fixes S :: "'a::{complete_space,real_normed_vector} set"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2352 | assumes S: "open S" "connected S" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69745diff
changeset | 2353 | and opeSU: "openin (top_of_set (frontier S)) U" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69745diff
changeset | 2354 | and opeSV: "openin (top_of_set (frontier S)) V" | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2355 |       and "V \<noteq> {}" "\<not> U \<subseteq> V"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2356 |   obtains g where "arc g" "pathstart g \<in> U" "pathfinish g \<in> V" "g ` {0<..<1} \<subseteq> S"
 | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 2357 | proof - | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2358 |   have "S \<noteq> {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2359 |     using opeSV \<open>V \<noteq> {}\<close> by (metis frontier_empty openin_subtopology_empty)
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2360 | then obtain x where "x \<in> S" by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2361 |   obtain g where "arc g" and g: "g ` {0..<1} \<subseteq> S" "pathstart g = x" "pathfinish g \<in> V"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2362 |     using dense_accessible_frontier_points_connected [OF S \<open>x \<in> S\<close> \<open>V \<noteq> {}\<close> opeSV] by blast
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2363 |   obtain h where "arc h" and h: "h ` {0..<1} \<subseteq> S" "pathstart h = x" "pathfinish h \<in> U - {pathfinish g}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2364 | proof (rule dense_accessible_frontier_points_connected [OF S \<open>x \<in> S\<close>]) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2365 |     show "U - {pathfinish g} \<noteq> {}"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2366 | using \<open>pathfinish g \<in> V\<close> \<open>\<not> U \<subseteq> V\<close> by blast | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69745diff
changeset | 2367 |     show "openin (top_of_set (frontier S)) (U - {pathfinish g})"
 | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2368 | by (simp add: opeSU openin_delete) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2369 | qed auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2370 | obtain \<gamma> where "arc \<gamma>" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2371 | and \<gamma>: "path_image \<gamma> \<subseteq> path_image (reversepath h +++ g)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2372 | "pathstart \<gamma> = pathfinish h" "pathfinish \<gamma> = pathfinish g" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2373 | proof (rule path_contains_arc [of "(reversepath h +++ g)" "pathfinish h" "pathfinish g"]) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2374 | show "path (reversepath h +++ g)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2375 | by (simp add: \<open>arc g\<close> \<open>arc h\<close> \<open>pathstart g = x\<close> \<open>pathstart h = x\<close> arc_imp_path) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2376 | show "pathstart (reversepath h +++ g) = pathfinish h" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2377 | "pathfinish (reversepath h +++ g) = pathfinish g" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2378 | by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2379 | show "pathfinish h \<noteq> pathfinish g" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2380 |       using \<open>pathfinish h \<in> U - {pathfinish g}\<close> by auto
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2381 | qed auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2382 | show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2383 | proof | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2384 | show "arc \<gamma>" "pathstart \<gamma> \<in> U" "pathfinish \<gamma> \<in> V" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2385 |       using \<gamma> \<open>arc \<gamma>\<close> \<open>pathfinish h \<in> U - {pathfinish g}\<close>  \<open>pathfinish g \<in> V\<close> by auto
 | 
| 72531 | 2386 | have "path_image \<gamma> \<subseteq> path_image h \<union> path_image g" | 
| 2387 | by (metis \<gamma>(1) g(2) h(2) path_image_join path_image_reversepath pathfinish_reversepath) | |
| 2388 |     then have "\<gamma> ` {0..1} - {\<gamma> 0, \<gamma> 1} \<subseteq> S"
 | |
| 2389 | using \<gamma> g h | |
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2390 | apply (simp add: path_image_def pathstart_def pathfinish_def subset_iff image_def Bex_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2391 | by (metis linorder_neqE_linordered_idom not_less) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2392 |     then show "\<gamma> ` {0<..<1} \<subseteq> S"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2393 | using \<open>arc h\<close> \<open>arc \<gamma>\<close> | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2394 | by (metis arc_imp_simple_path path_image_def pathfinish_def pathstart_def simple_path_endless) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2395 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2396 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2397 | |
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 2398 | lemma dense_accessible_frontier_point_pairs: | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2399 |   fixes S :: "'a::{complete_space,real_normed_vector} set"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2400 | assumes S: "open S" "connected S" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69745diff
changeset | 2401 | and opeSU: "openin (top_of_set (frontier S)) U" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69745diff
changeset | 2402 | and opeSV: "openin (top_of_set (frontier S)) V" | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2403 |       and "U \<noteq> {}" "V \<noteq> {}" "U \<noteq> V"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2404 |     obtains g where "arc g" "pathstart g \<in> U" "pathfinish g \<in> V" "g ` {0<..<1} \<subseteq> S"
 | 
| 69652 
3417a8f154eb
updated tagging first 5
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69517diff
changeset | 2405 | proof - | 
| 64790 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2406 | consider "\<not> U \<subseteq> V" | "\<not> V \<subseteq> U" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2407 | using \<open>U \<noteq> V\<close> by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2408 | then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2409 | proof cases | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2410 | case 1 then show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2411 | using assms dense_access_fp_aux [OF S opeSU opeSV] that by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2412 | next | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2413 | case 2 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2414 |     obtain g where "arc g" and g: "pathstart g \<in> V" "pathfinish g \<in> U" "g ` {0<..<1} \<subseteq> S"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2415 | using assms dense_access_fp_aux [OF S opeSV opeSU] "2" by blast | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2416 | show ?thesis | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2417 | proof | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2418 | show "arc (reversepath g)" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2419 | by (simp add: \<open>arc g\<close> arc_reversepath) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2420 | show "pathstart (reversepath g) \<in> U" "pathfinish (reversepath g) \<in> V" | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2421 | using g by auto | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2422 |       show "reversepath g ` {0<..<1} \<subseteq> S"
 | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2423 | using g by (auto simp: reversepath_def) | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2424 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2425 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2426 | qed | 
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2427 | |
| 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2428 | end |