author | wenzelm |
Sat, 13 Aug 2022 18:06:30 +0200 | |
changeset 75848 | 9e4c0aaa30aa |
parent 73932 | fd21b4a93043 |
child 77943 | ffdad62bc235 |
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:
70817
diff
changeset
|
11 |
lemma path_connected_interval [simp]: |
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
immler
parents:
70817
diff
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:
70817
diff
changeset
|
13 |
shows "path_connected {a..b}" |
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
immler
parents:
70817
diff
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:
70817
diff
changeset
|
15 |
|
71028
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
immler
parents:
71025
diff
changeset
|
16 |
lemma segment_to_closest_point: |
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
immler
parents:
71025
diff
changeset
|
17 |
fixes S :: "'a :: euclidean_space set" |
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
immler
parents:
71025
diff
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:
71025
diff
changeset
|
21 |
|
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
immler
parents:
71025
diff
changeset
|
22 |
lemma segment_to_point_exists: |
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
immler
parents:
71025
diff
changeset
|
23 |
fixes S :: "'a :: euclidean_space set" |
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
immler
parents:
71025
diff
changeset
|
24 |
assumes "closed S" "S \<noteq> {}" |
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
immler
parents:
71025
diff
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:
71025
diff
changeset
|
26 |
by (metis assms segment_to_closest_point closest_point_exists that) |
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
immler
parents:
71025
diff
changeset
|
27 |
|
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
immler
parents:
71025
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
70136
diff
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:
69517
diff
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:
69517
diff
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:
70136
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
69517
diff
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:
70136
diff
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 |
|
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
536 |
|
69652
3417a8f154eb
updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69517
diff
changeset
|
537 |
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
|
538 |
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
|
539 |
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
|
540 |
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
|
541 |
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
|
542 |
shows "(f ` {0..1}) homeomorphic {0..1::real}" |
69652
3417a8f154eb
updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69517
diff
changeset
|
543 |
proof - |
64790
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
544 |
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
|
545 |
(\<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
|
546 |
(\<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
|
547 |
(\<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
|
548 |
(\<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
|
549 |
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
|
550 |
proof - |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
551 |
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
|
552 |
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
|
553 |
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
|
554 |
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
|
555 |
by (metis Int_commute) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
556 |
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
|
557 |
by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
558 |
obtain c d where cd: "{a..b} \<inter> f -` {f m} = {c..d}" |
72531 | 559 |
using ab01 cd0 |
560 |
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
|
561 |
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
|
562 |
by blast |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
563 |
show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
564 |
proof (intro exI conjI ballI) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
565 |
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
|
566 |
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
|
567 |
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
|
568 |
using cd m by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
569 |
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
|
570 |
using cd by blast |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
571 |
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
|
572 |
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
|
573 |
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
|
574 |
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
|
575 |
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
|
576 |
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
|
577 |
case True |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
578 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
579 |
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
|
580 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
581 |
case False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
582 |
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
|
583 |
proof - |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
584 |
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
|
585 |
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
|
586 |
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
|
587 |
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
|
588 |
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
|
589 |
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
|
590 |
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
|
591 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
592 |
then show ?thesis by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
593 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
594 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
595 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
596 |
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
|
597 |
"\<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
|
598 |
(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
|
599 |
(\<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
|
600 |
(\<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
|
601 |
(\<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
|
602 |
(\<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
|
603 |
apply atomize |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
604 |
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
|
605 |
apply (rule that, blast) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
606 |
done |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
607 |
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 | 608 |
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
|
609 |
by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
610 |
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
|
611 |
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
|
612 |
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
|
613 |
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
|
614 |
\<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
|
615 |
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
|
616 |
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
|
617 |
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
|
618 |
by (metis feqm) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
619 |
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
|
620 |
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
|
621 |
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
|
622 |
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
|
623 |
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
|
624 |
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
|
625 |
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
|
626 |
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
|
627 |
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
|
628 |
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
|
629 |
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
|
630 |
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
|
631 |
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
|
632 |
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
|
633 |
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
|
634 |
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
|
635 |
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
|
636 |
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
|
637 |
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
|
638 |
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
|
639 |
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
|
640 |
using a_def leftrec_base |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
641 |
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
|
642 |
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
|
643 |
using b_def rightrec_base |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
644 |
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
|
645 |
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
|
646 |
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
|
647 |
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
|
648 |
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
|
649 |
(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
|
650 |
(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
|
651 |
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
|
652 |
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
|
653 |
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
|
654 |
(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
|
655 |
(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
|
656 |
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
|
657 |
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
|
658 |
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
|
659 |
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
|
660 |
proof (induction n arbitrary: m) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
661 |
case 0 |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
662 |
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
|
663 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
664 |
case (Suc n p) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
665 |
show ?case |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
666 |
proof (cases "even p") |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
667 |
case True |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
668 |
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
|
669 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
670 |
by (simp add: Suc.IH) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
671 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
672 |
case False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
673 |
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
|
674 |
show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
675 |
proof (cases n) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
676 |
case 0 |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
677 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
678 |
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
|
679 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
680 |
case (Suc n') |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
681 |
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
|
682 |
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
|
683 |
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
|
684 |
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
|
685 |
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
|
686 |
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
|
687 |
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
|
688 |
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
|
689 |
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
|
690 |
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
|
691 |
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
|
692 |
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
|
693 |
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
|
694 |
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
|
695 |
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
|
696 |
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
|
697 |
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
|
698 |
show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
699 |
proof (cases "even m") |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
700 |
case True |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
701 |
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
|
702 |
show ?thesis |
72531 | 703 |
using order_trans [OF left_le c_le_v, of "1+2*r"] Suc.IH [of "m+1"] |
704 |
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> |
|
705 |
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
|
706 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
707 |
case False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
708 |
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
|
709 |
show ?thesis |
72531 | 710 |
using order_trans [OF c_ge_u right_ge, of "1+2*r"] Suc.IH [of "m"] |
711 |
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> |
|
712 |
apply (simp_all add: r m add.commute [of 3] a43 b43 del: power_Suc) |
|
713 |
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
|
714 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
715 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
716 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
717 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
718 |
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
|
719 |
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
|
720 |
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
|
721 |
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
|
722 |
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
|
723 |
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
|
724 |
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 | 725 |
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
|
726 |
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
|
727 |
\<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
|
728 |
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
|
729 |
case (less d j n) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
730 |
show ?case |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
731 |
proof (cases "m \<le> n") |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
732 |
case True |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
733 |
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
|
734 |
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
|
735 |
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
|
736 |
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
|
737 |
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
|
738 |
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
|
739 |
also have "... \<in> \<int>" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
740 |
by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
741 |
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
|
742 |
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:
70136
diff
changeset
|
743 |
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
|
744 |
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:
70136
diff
changeset
|
745 |
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
|
746 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
747 |
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
|
748 |
by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
749 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
750 |
by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
751 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
752 |
case False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
753 |
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
|
754 |
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
|
755 |
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
|
756 |
show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
757 |
proof (cases "n > 0") |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
758 |
case False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
759 |
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
|
760 |
by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
761 |
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
|
762 |
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
|
763 |
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
|
764 |
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
|
765 |
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
|
766 |
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
|
767 |
using False by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
768 |
finally show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
769 |
by (auto simp: ac) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
770 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
771 |
case True show ?thesis |
72531 | 772 |
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
|
773 |
case less |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
774 |
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:
70136
diff
changeset
|
775 |
using k by (force simp: field_split_simps) |
72531 | 776 |
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
|
777 |
using less.prems by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
778 |
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
|
779 |
using less.prems by linarith |
72531 | 780 |
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
|
781 |
c (real i / 2 ^ m) \<le> b (real (4 * k + 1) / 2 ^ Suc n)" |
72531 | 782 |
proof (rule less.IH [OF _ refl]) |
783 |
show "m - Suc n < d" |
|
784 |
using \<open>n < m\<close> diff_less_mono2 less.prems(1) lessI by presburger |
|
785 |
show "\<bar>real i / 2 ^ m - real (4 * k + 1) / 2 ^ Suc n\<bar> < 1 / 2 ^ Suc n" |
|
786 |
using closer \<open>n < m\<close> \<open>d = m - n\<close> by (auto simp: field_split_simps \<open>n < m\<close> diff_less_mono2) |
|
787 |
qed auto |
|
788 |
then show ?thesis |
|
64790
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
789 |
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
|
790 |
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 | 791 |
using k a41 b41 \<open>0 < n\<close> |
792 |
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
|
793 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
794 |
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
|
795 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
796 |
case greater |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
797 |
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:
70136
diff
changeset
|
798 |
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
|
799 |
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
|
800 |
using less.prems by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
801 |
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
|
802 |
using less.prems by linarith |
72531 | 803 |
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
|
804 |
c (real i / 2 ^ m) \<le> b (real (4 * k + 3) / 2 ^ Suc n)" |
72531 | 805 |
proof (rule less.IH [OF _ refl]) |
806 |
show "m - Suc n < d" |
|
807 |
using \<open>n < m\<close> diff_less_mono2 less.prems(1) by blast |
|
808 |
show "\<bar>real i / 2 ^ m - real (4 * k + 3) / 2 ^ Suc n\<bar> < 1 / 2 ^ Suc n" |
|
809 |
using closer \<open>n < m\<close> \<open>d = m - n\<close> by (auto simp: field_split_simps \<open>n < m\<close> diff_less_mono2) |
|
810 |
qed auto |
|
811 |
then show ?thesis |
|
64790
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
812 |
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
|
813 |
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 | 814 |
using k a43 b43 \<open>0 < n\<close> |
815 |
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
|
816 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
817 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
818 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
819 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
820 |
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
|
821 |
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
|
822 |
using that by blast+ |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
823 |
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
|
824 |
proof (induction n arbitrary: m) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
825 |
case 0 |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
826 |
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
|
827 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
828 |
case (Suc n m) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
829 |
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
|
830 |
show ?case |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
831 |
proof (cases "n > 0") |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
832 |
case False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
833 |
with u01 v01 show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
834 |
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
|
835 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
836 |
case True |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
837 |
show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
838 |
proof (cases "even k") |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
839 |
case True |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
840 |
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
|
841 |
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
|
842 |
proof - |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
843 |
have "odd (Suc k)" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
844 |
using True by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
845 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
846 |
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
|
847 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
848 |
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
|
849 |
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
|
850 |
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
|
851 |
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
|
852 |
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
|
853 |
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
|
854 |
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
|
855 |
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
|
856 |
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
|
857 |
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
|
858 |
\<le> 2/2 ^ Suc n" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
859 |
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
|
860 |
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
|
861 |
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
|
862 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
863 |
case False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
864 |
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
|
865 |
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
|
866 |
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
|
867 |
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
|
868 |
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
|
869 |
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
|
870 |
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
|
871 |
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
|
872 |
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
|
873 |
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
|
874 |
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
|
875 |
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
|
876 |
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
|
877 |
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
|
878 |
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
|
879 |
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
|
880 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
881 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
882 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
883 |
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
|
884 |
using that by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
885 |
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
|
886 |
proof (induction n arbitrary: j) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
887 |
case 0 |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
888 |
then show ?case by auto |
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 j) show ?case |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
891 |
proof (cases "n > 0") |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
892 |
case False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
893 |
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
|
894 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
895 |
case True |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
896 |
show ?thesis proof (cases "even j") |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
897 |
case True |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
898 |
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
|
899 |
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
|
900 |
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
|
901 |
with k \<open>0 < n\<close> Suc.IH [of k] show ?thesis |
72531 | 902 |
by (simp add: m1_to_3 a41 b43 del: power_Suc) (auto simp: of_nat_diff) |
64790
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
903 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
904 |
case False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
905 |
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
|
906 |
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
|
907 |
= 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
|
908 |
"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
|
909 |
= 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
|
910 |
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
|
911 |
using left_right_m [of "c((2*k + 1) / 2^n)" "a((2*k + 1) / 2^n)" "b((2*k + 1) / 2^n)"] |
72531 | 912 |
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
|
913 |
then |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
914 |
show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
915 |
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
|
916 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
917 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
918 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
919 |
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
|
920 |
\<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
|
921 |
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
|
922 |
proof (induction n arbitrary: j) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
923 |
case 0 |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
924 |
then show ?case by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
925 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
926 |
case (Suc n) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
927 |
show ?case |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
928 |
proof (cases "even j") |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
929 |
case True |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
930 |
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
|
931 |
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
|
932 |
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
|
933 |
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
|
934 |
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
|
935 |
by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
936 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
937 |
using Suc.IH [of k] k \<open>0 < k\<close> |
72531 | 938 |
by (simp add: less2n add.commute [of 1] m1_to_3 a41 b43 del: power_Suc) (auto simp: of_nat_diff) |
64790
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
939 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
940 |
case False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
941 |
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
|
942 |
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
|
943 |
show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
944 |
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
|
945 |
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
|
946 |
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
|
947 |
apply (force intro: feqm) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
948 |
done |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
949 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
950 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
951 |
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
|
952 |
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
|
953 |
unfolding D01_def |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
954 |
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
|
955 |
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
|
956 |
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
|
957 |
fix e::real |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
958 |
assume "0 < e" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
959 |
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
|
960 |
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
|
961 |
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
|
962 |
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
|
963 |
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
|
964 |
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
|
965 |
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
|
966 |
with gr0I have "n > 0" |
70817
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents:
70136
diff
changeset
|
967 |
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
|
968 |
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
|
969 |
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
|
970 |
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
|
971 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
972 |
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
|
973 |
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
|
974 |
proof - |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
975 |
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
|
976 |
by linarith |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
977 |
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
|
978 |
| "\<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
|
979 |
| "\<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
|
980 |
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
|
981 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
982 |
proof cases |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
983 |
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
|
984 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
985 |
case 2 |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
986 |
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
|
987 |
by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
988 |
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
|
989 |
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
|
990 |
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
|
991 |
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
|
992 |
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
|
993 |
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
|
994 |
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
|
995 |
done |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
996 |
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
|
997 |
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
|
998 |
ultimately show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
999 |
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
|
1000 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1001 |
case 3 |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1002 |
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
|
1003 |
by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1004 |
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
|
1005 |
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
|
1006 |
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
|
1007 |
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
|
1008 |
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
|
1009 |
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
|
1010 |
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
|
1011 |
done |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1012 |
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
|
1013 |
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
|
1014 |
ultimately show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1015 |
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
|
1016 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1017 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1018 |
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
|
1019 |
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
|
1020 |
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
|
1021 |
proof clarsimp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1022 |
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
|
1023 |
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
|
1024 |
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
|
1025 |
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
|
1026 |
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
|
1027 |
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
|
1028 |
proof - |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1029 |
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
|
1030 |
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
|
1031 |
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
|
1032 |
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
|
1033 |
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
|
1034 |
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
|
1035 |
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
|
1036 |
show thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1037 |
proof (cases "j = 0") |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1038 |
case True |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1039 |
show thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1040 |
proof |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1041 |
show "(1::nat) < 2 ^ n" |
72531 | 1042 |
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
|
1043 |
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
|
1044 |
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
|
1045 |
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
|
1046 |
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
|
1047 |
qed simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1048 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1049 |
case False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1050 |
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
|
1051 |
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
|
1052 |
for i k m p |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1053 |
proof - |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1054 |
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
|
1055 |
using j by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1056 |
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
|
1057 |
using k by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1058 |
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
|
1059 |
by (simp only: mult_ac) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1060 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1061 |
by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1062 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1063 |
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
|
1064 |
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
|
1065 |
for i k m p |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1066 |
proof - |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1067 |
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
|
1068 |
using j by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1069 |
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
|
1070 |
by (rule k) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1071 |
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
|
1072 |
by (simp add: algebra_simps) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1073 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1074 |
by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1075 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1076 |
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
|
1077 |
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
|
1078 |
proof - |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1079 |
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
|
1080 |
using j by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1081 |
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
|
1082 |
using i by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1083 |
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
|
1084 |
by (simp only: mult_ac) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1085 |
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
|
1086 |
by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1087 |
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
|
1088 |
by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1089 |
finally show ?thesis by simp |
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 |
show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1092 |
proof |
72531 | 1093 |
have "2 ^ n * real i / 2 ^ m < 2 ^ n" "2 ^ n * real k / 2 ^ p < 2 ^ n" |
1094 |
using i k by (auto simp: field_simps) |
|
1095 |
then have "max (2^n * i / 2^m) (2^n * k / 2^p) < 2^n" |
|
1096 |
by simp |
|
1097 |
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
|
1098 |
then show "j < 2 ^ n" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1099 |
by auto |
72531 | 1100 |
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
|
1101 |
using clo less_j1 j_le |
72531 | 1102 |
by (auto simp: le_max_iff_disj field_split_simps dist_norm abs_if split: if_split_asm dest: 1 2) |
1103 |
then show "\<bar>real i / 2 ^ m - real j / 2 ^ n\<bar> < 1/2 ^ n" |
|
1104 |
by (auto simp: field_split_simps) |
|
1105 |
have "\<bar>real k * 2 ^ n - real j * 2 ^ p\<bar> < 2 ^ p" |
|
1106 |
using clo less_j1 j_le |
|
1107 |
by (auto simp: le_max_iff_disj field_split_simps dist_norm abs_if split: if_split_asm dest: 3 2) |
|
1108 |
then show "\<bar>real k / 2 ^ p - real j / 2 ^ n\<bar> < 1/2 ^ n" |
|
1109 |
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
|
1110 |
qed (use False in simp) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1111 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1112 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1113 |
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
|
1114 |
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
|
1115 |
show "dist (f (c (real k / 2 ^ p))) (f(c(j / 2^n))) < e/2" |
72531 | 1116 |
using \<open>0 < j\<close> \<open>j < 2 ^ n\<close> k clo_kj |
1117 |
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
|
1118 |
show "dist (f (c (real i / 2 ^ m))) (f (c (real j / 2 ^ n))) < e/2" |
72531 | 1119 |
using \<open>0 < j\<close> \<open>j < 2 ^ n\<close> i clo_ij |
1120 |
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
|
1121 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1122 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1123 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1124 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1125 |
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
|
1126 |
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
|
1127 |
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
|
1128 |
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
|
1129 |
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
|
1130 |
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
|
1131 |
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
|
1132 |
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
|
1133 |
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
|
1134 |
proof |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1135 |
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
|
1136 |
proof (rule image_closure_subset) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1137 |
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
|
1138 |
using cont_h by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1139 |
show "closed (f ` {0..1})" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1140 |
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
|
1141 |
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
|
1142 |
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
|
1143 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1144 |
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
|
1145 |
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
|
1146 |
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
|
1147 |
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
|
1148 |
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
|
1149 |
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
|
1150 |
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
|
1151 |
fix x e::real |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1152 |
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
|
1153 |
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
|
1154 |
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
|
1155 |
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
|
1156 |
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
|
1157 |
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
|
1158 |
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
|
1159 |
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
|
1160 |
using that |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1161 |
proof (induction n) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1162 |
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
|
1163 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1164 |
case (Suc n) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1165 |
show ?case |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1166 |
proof (cases "n=0") |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1167 |
case True |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1168 |
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
|
1169 |
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
|
1170 |
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
|
1171 |
proof cases |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1172 |
case 1 |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1173 |
then show ?thesis |
72531 | 1174 |
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
|
1175 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1176 |
case 2 |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1177 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1178 |
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
|
1179 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1180 |
case 3 |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1181 |
then show ?thesis |
72531 | 1182 |
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
|
1183 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1184 |
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
|
1185 |
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
|
1186 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1187 |
case False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1188 |
with Suc obtain m y |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1189 |
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
|
1190 |
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
|
1191 |
by metis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1192 |
then obtain j where j: "m = 2*j + 1" by (metis oddE) |
72531 | 1193 |
have j4: "4 * j + 1 < 2 ^ Suc n" |
1194 |
using mless j by (simp add: algebra_simps) |
|
1195 |
||
64790
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1196 |
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
|
1197 |
| "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
|
1198 |
| "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
|
1199 |
using y j by force |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1200 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1201 |
proof cases |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1202 |
case 1 |
72531 | 1203 |
show ?thesis |
1204 |
proof (intro exI conjI) |
|
1205 |
show "y \<in> {a (real (4 * j + 1) / 2 ^ Suc n)..b (real (4 * j + 1) / 2 ^ Suc n)}" |
|
1206 |
using mless j \<open>n \<noteq> 0\<close> 1 by (simp add: a41 b41 add.commute [of 1] del: power_Suc) |
|
1207 |
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
|
1208 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1209 |
case 2 |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1210 |
show ?thesis |
72531 | 1211 |
proof (intro exI conjI) |
1212 |
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)}" |
|
1213 |
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] |
|
1214 |
using left_right [of "c((2*j + 1) / 2^n)" "a((2*j + 1) / 2^n)" "b((2*j + 1) / 2^n)"] |
|
1215 |
by (simp add: a41 b41 add.commute [of 1] del: power_Suc) |
|
1216 |
show "f (b (real (4 * j + 1) / 2 ^ Suc n)) = f x" |
|
1217 |
using \<open>n \<noteq> 0\<close> 2 |
|
1218 |
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] |
|
1219 |
by (force simp add: b41 a43 add.commute [of 1] feq [symmetric] simp del: power_Suc intro: f_eqI) |
|
1220 |
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
|
1221 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1222 |
case 3 |
72531 | 1223 |
show ?thesis |
1224 |
proof (intro exI conjI) |
|
1225 |
show "4 * j + 3 < 2 ^ Suc n" |
|
1226 |
using mless j by simp |
|
1227 |
show "f y = f x" |
|
1228 |
by fact |
|
1229 |
show "y \<in> {a (real (4 * j + 3) / 2 ^ Suc n) .. b (real (4 * j + 3) / 2 ^ Suc n)}" |
|
1230 |
using 3 False b43 [of n j] by (simp add: add.commute) |
|
1231 |
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
|
1232 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1233 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1234 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1235 |
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
|
1236 |
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
|
1237 |
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
|
1238 |
by fastforce |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1239 |
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
|
1240 |
where "odd m" "0 < m" and mless: "m < 2 ^ n" |
72531 | 1241 |
and y: "a(m / 2^n) \<le> y \<and> y \<le> b(m / 2^n)" and feq: "f x = f y" |
1242 |
by (metis atLeastAtMost_iff) |
|
64790
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1243 |
then have "0 \<le> y" "y \<le> 1" |
72531 | 1244 |
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
|
1245 |
moreover have "y < \<delta> + c (real m / 2 ^ n)" "c (real m / 2 ^ n) < \<delta> + y" |
72531 | 1246 |
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
|
1247 |
by linarith+ |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1248 |
moreover note \<open>0 < m\<close> mless \<open>0 \<le> x\<close> \<open>x \<le> 1\<close> |
72531 | 1249 |
ultimately have "dist (h (real m / 2 ^ n)) (f x) < e" |
1250 |
by (auto simp: dist_norm h_eq feq \<delta>) |
|
1251 |
then show "\<exists>k. \<exists>m\<in>{0<..<2 ^ k}. dist (h (real m / 2 ^ k)) (f x) < e" |
|
1252 |
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
|
1253 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1254 |
also have "... \<subseteq> h ` {0..1}" |
72531 | 1255 |
proof (rule closure_minimal) |
1256 |
show "h ` D01 \<subseteq> h ` {0..1}" |
|
1257 |
using cloD01 closure_subset by blast |
|
1258 |
show "closed (h ` {0..1})" |
|
1259 |
using compact_continuous_image [OF cont_h] compact_imp_closed by auto |
|
1260 |
qed |
|
64790
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1261 |
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
|
1262 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1263 |
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
|
1264 |
proof - |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1265 |
have "u < v" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1266 |
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
|
1267 |
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
|
1268 |
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
|
1269 |
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
|
1270 |
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
|
1271 |
have a_less_b: |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1272 |
"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
|
1273 |
(\<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
|
1274 |
(\<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
|
1275 |
proof (induction n arbitrary: j) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1276 |
case 0 then show ?case |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1277 |
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
|
1278 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1279 |
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
|
1280 |
proof (cases "n > 0") |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1281 |
case False then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1282 |
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
|
1283 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1284 |
case True show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1285 |
proof (cases "even j") |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1286 |
case True |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1287 |
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
|
1288 |
by (auto elim!: evenE) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1289 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1290 |
case False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1291 |
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
|
1292 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1293 |
proof (cases "even k") |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1294 |
case True |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1295 |
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
|
1296 |
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
|
1297 |
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
|
1298 |
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
|
1299 |
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
|
1300 |
by (auto intro: f_eqI) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1301 |
show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1302 |
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
|
1303 |
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
|
1304 |
proof - |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1305 |
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
|
1306 |
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
|
1307 |
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
|
1308 |
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
|
1309 |
by (auto simp: algebra_simps) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1310 |
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
|
1311 |
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
|
1312 |
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
|
1313 |
using cleb by blast |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1314 |
ultimately show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1315 |
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
|
1316 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1317 |
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
|
1318 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1319 |
fix x |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1320 |
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
|
1321 |
then show False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1322 |
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
|
1323 |
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
|
1324 |
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
|
1325 |
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
|
1326 |
by (auto simp: algebra_simps) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1327 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1328 |
fix x |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1329 |
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
|
1330 |
then show False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1331 |
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
|
1332 |
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
|
1333 |
by (auto simp: algebra_simps) |
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 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1336 |
case False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1337 |
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
|
1338 |
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
|
1339 |
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
|
1340 |
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
|
1341 |
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
|
1342 |
show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1343 |
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
|
1344 |
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
|
1345 |
proof - |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1346 |
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
|
1347 |
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
|
1348 |
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
|
1349 |
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
|
1350 |
by (auto simp: algebra_simps) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1351 |
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
|
1352 |
using alec by blast |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1353 |
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
|
1354 |
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
|
1355 |
ultimately show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1356 |
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
|
1357 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1358 |
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
|
1359 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1360 |
fix x |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1361 |
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
|
1362 |
then show False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1363 |
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
|
1364 |
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
|
1365 |
by (auto simp: algebra_simps) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1366 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1367 |
fix x |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1368 |
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
|
1369 |
then show False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1370 |
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
|
1371 |
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
|
1372 |
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
|
1373 |
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
|
1374 |
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
|
1375 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1376 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1377 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1378 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1379 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1380 |
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
|
1381 |
using a_less_b [of m n] apply (simp_all add: c_def midpoint_def) |
72531 | 1382 |
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
|
1383 |
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
|
1384 |
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
|
1385 |
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
|
1386 |
\<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
|
1387 |
\<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
|
1388 |
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
|
1389 |
using that |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1390 |
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
|
1391 |
case (less N) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1392 |
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
|
1393 |
by linarith |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1394 |
then show ?case |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1395 |
proof cases |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1396 |
case 1 |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1397 |
with less.prems show ?thesis |
70817
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents:
70136
diff
changeset
|
1398 |
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
|
1399 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1400 |
case 2 show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1401 |
proof (cases m) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1402 |
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
|
1403 |
by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1404 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1405 |
case (Suc m') show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1406 |
proof (cases p) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1407 |
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
|
1408 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1409 |
case (Suc p') |
72531 | 1410 |
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
|
1411 |
proof - |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1412 |
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
|
1413 |
using that by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1414 |
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
|
1415 |
using that by linarith |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1416 |
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
|
1417 |
qed |
72531 | 1418 |
moreover have *: "real i / 2 ^ m' < real k / 2^p'" "k < 2 ^ p'" |
1419 |
using less.prems \<open>m = Suc m'\<close> 2 Suc by (force simp: field_split_simps)+ |
|
1420 |
moreover have "i < 2 ^ m' " |
|
1421 |
using \<section> * by (clarsimp simp: divide_simps linorder_not_le) (meson linorder_not_le) |
|
1422 |
ultimately show ?thesis |
|
64790
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1423 |
using less.IH [of "m'+p'" i m' k p'] less.prems \<open>m = Suc m'\<close> 2 Suc |
72531 | 1424 |
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
|
1425 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1426 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1427 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1428 |
case 3 show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1429 |
proof (cases m) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1430 |
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
|
1431 |
by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1432 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1433 |
case (Suc m') show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1434 |
proof (cases p) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1435 |
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
|
1436 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1437 |
case (Suc p') |
72531 | 1438 |
have "real (i - 2 ^ m') / 2 ^ m' < real (k - 2 ^ p') / 2 ^ p'" |
1439 |
using less.prems \<open>m = Suc m'\<close> Suc 3 by (auto simp: field_simps of_nat_diff) |
|
1440 |
moreover have "k - 2 ^ p' < 2 ^ p'" "i - 2 ^ m' < 2 ^ m'" |
|
1441 |
using less.prems Suc \<open>m = Suc m'\<close> by auto |
|
1442 |
moreover |
|
1443 |
have "2 ^ p' \<le> k" "2 ^ p' \<noteq> k" |
|
1444 |
using less.prems \<open>m = Suc m'\<close> Suc 3 by auto |
|
1445 |
then have "2 ^ p' < k" |
|
1446 |
by linarith |
|
1447 |
ultimately show ?thesis |
|
64790
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1448 |
using less.IH [of "m'+p'" "i - 2^m'" m' "k - 2 ^ p'" p'] less.prems \<open>m = Suc m'\<close> Suc 3 |
72531 | 1449 |
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
|
1450 |
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
|
1451 |
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
|
1452 |
apply (auto simp: field_simps) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1453 |
done |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1454 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1455 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1456 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1457 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1458 |
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
|
1459 |
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
|
1460 |
proof - |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1461 |
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
|
1462 |
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
|
1463 |
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
|
1464 |
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
|
1465 |
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
|
1466 |
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
|
1467 |
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
|
1468 |
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
|
1469 |
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
|
1470 |
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
|
1471 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1472 |
case False |
72531 | 1473 |
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'" |
1474 |
by (auto simp: field_split_simps) |
|
1475 |
then have "c(i / 2^m) \<le> b(real(4 * q + 1) / 2 ^ (Suc n'))" |
|
1476 |
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
|
1477 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1478 |
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
|
1479 |
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
|
1480 |
by (auto simp: algebra_simps) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1481 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1482 |
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
|
1483 |
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
|
1484 |
case True |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1485 |
then show ?thesis by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1486 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1487 |
case False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1488 |
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
|
1489 |
by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1490 |
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
|
1491 |
by auto |
72531 | 1492 |
have "\<bar>real j / 2 ^ n - real (4 * q + 3) / 2 ^ Suc n'\<bar> < 1 / 2 ^ Suc n'" |
1493 |
by (rule * [OF less]) (use j_le_j clo_jj q in \<open>auto simp: field_split_simps\<close>) |
|
1494 |
then have "a(real(4*q + 3) / 2 ^ (Suc n')) \<le> c(j / 2^n)" |
|
1495 |
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
|
1496 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1497 |
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
|
1498 |
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
|
1499 |
by (auto simp: algebra_simps) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1500 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1501 |
finally show ?thesis . |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1502 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1503 |
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
|
1504 |
using that |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1505 |
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
|
1506 |
case (less x1 x2) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1507 |
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
|
1508 |
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
|
1509 |
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
|
1510 |
proof - |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1511 |
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
|
1512 |
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
|
1513 |
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
|
1514 |
unfolding closure_approachable |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1515 |
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
|
1516 |
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
|
1517 |
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
|
1518 |
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
|
1519 |
proof - |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1520 |
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
|
1521 |
using less by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1522 |
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
|
1523 |
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
|
1524 |
then have "N > 0" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1525 |
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
|
1526 |
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
|
1527 |
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
|
1528 |
show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1529 |
proof |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1530 |
show "0 < 2^N * p" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1531 |
using p by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1532 |
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
|
1533 |
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
|
1534 |
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
|
1535 |
by (simp add: power_add) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1536 |
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
|
1537 |
by (simp add: yeq) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1538 |
also have "... < (x2 - x1) / 64" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1539 |
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
|
1540 |
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
|
1541 |
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
|
1542 |
by (simp add: field_simps) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1543 |
also have "... < (x2 - x1) / 128" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1544 |
using N by force |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1545 |
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
|
1546 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1547 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1548 |
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
|
1549 |
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
|
1550 |
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
|
1551 |
proof |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1552 |
show "0 < Suc (2*m)" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1553 |
by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1554 |
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
|
1555 |
using m by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1556 |
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
|
1557 |
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
|
1558 |
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
|
1559 |
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
|
1560 |
show "0 < 4*m + 3" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1561 |
by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1562 |
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
|
1563 |
using m by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1564 |
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
|
1565 |
by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1566 |
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
|
1567 |
by (simp add: algebra_simps) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1568 |
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
|
1569 |
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
|
1570 |
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
|
1571 |
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
|
1572 |
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
|
1573 |
by (simp add: c_def) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1574 |
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
|
1575 |
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
|
1576 |
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
|
1577 |
by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1578 |
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
|
1579 |
by (simp add: midpoint_def) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1580 |
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
|
1581 |
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
|
1582 |
by (simp add: midpoint_def) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1583 |
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
|
1584 |
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
|
1585 |
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
|
1586 |
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
|
1587 |
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
|
1588 |
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
|
1589 |
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
|
1590 |
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
|
1591 |
apply (rule right_neq) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1592 |
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
|
1593 |
apply (rule midR_le, auto) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1594 |
done |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1595 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1596 |
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
|
1597 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1598 |
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:
70136
diff
changeset
|
1599 |
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
|
1600 |
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
|
1601 |
by (subst closure_dyadic_rationals_in_convex_set_pos_1, simp_all add: not_le m) |
72531 | 1602 |
have "2^n > m" |
1603 |
by (simp add: m(2) not_le) |
|
1604 |
then have closurem1: "{m / 2^n .. 1} = closure ({m / 2^n <..< 1} \<inter> (\<Union>k m. {real m / 2 ^ k}))" |
|
1605 |
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
|
1606 |
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
|
1607 |
if "0 \<le> u" "v \<le> 1" for u v |
72531 | 1608 |
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
|
1609 |
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
|
1610 |
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
|
1611 |
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
|
1612 |
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
|
1613 |
by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1614 |
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
|
1615 |
proof clarsimp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1616 |
fix p q |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1617 |
assume p: "0 < real p / 2 ^ q" "real p / 2 ^ q < real m / 2 ^ n" |
72531 | 1618 |
then have [simp]: "0 < p" |
1619 |
by (simp add: field_split_simps) |
|
1620 |
have [simp]: "p < 2 ^ q" |
|
1621 |
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
|
1622 |
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
|
1623 |
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
|
1624 |
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
|
1625 |
by (simp add: h_eq) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1626 |
qed |
72531 | 1627 |
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
|
1628 |
apply (subst closure0m) |
72531 | 1629 |
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
|
1630 |
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
|
1631 |
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
|
1632 |
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
|
1633 |
by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1634 |
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
|
1635 |
proof clarsimp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1636 |
fix p q |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1637 |
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
|
1638 |
then have [simp]: "0 < p" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1639 |
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
|
1640 |
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
|
1641 |
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
|
1642 |
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
|
1643 |
by (simp add: h_eq) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1644 |
qed |
72531 | 1645 |
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
|
1646 |
apply (subst closurem1) |
72531 | 1647 |
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
|
1648 |
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
|
1649 |
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
|
1650 |
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
|
1651 |
by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1652 |
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
|
1653 |
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
|
1654 |
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
|
1655 |
then show ?case by blast |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1656 |
qed auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1657 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1658 |
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
|
1659 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1660 |
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
|
1661 |
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
|
1662 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1663 |
using homeomorphic_sym by blast |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1664 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1665 |
|
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1666 |
|
69652
3417a8f154eb
updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69517
diff
changeset
|
1667 |
theorem path_contains_arc: |
64790
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1668 |
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
|
1669 |
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
|
1670 |
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:
69517
diff
changeset
|
1671 |
proof - |
64790
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1672 |
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
|
1673 |
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
|
1674 |
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
|
1675 |
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
|
1676 |
(\<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
|
1677 |
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
|
1678 |
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
|
1679 |
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
|
1680 |
using that by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1681 |
show "\<phi> {0..1}" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1682 |
by (auto simp: \<phi>_def open_segment_eq_real_ivl *) |
69313 | 1683 |
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
|
1684 |
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
|
1685 |
proof - |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1686 |
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
|
1687 |
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
|
1688 |
by (metis \<phi> \<phi>_def)+ |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1689 |
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
|
1690 |
for x y |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1691 |
using that |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1692 |
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
|
1693 |
case (less x y) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1694 |
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
|
1695 |
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
|
1696 |
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
|
1697 |
using less by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1698 |
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
|
1699 |
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
|
1700 |
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
|
1701 |
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
|
1702 |
by (subst min_less_iff_disj) (simp add: less) |
72531 | 1703 |
define w where "w \<equiv> x + (min e (y - x) / 3)" |
1704 |
define z where "z \<equiv>y - (min e (y - x) / 3)" |
|
1705 |
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
|
1706 |
and wxe: "norm(w - x) < e" and zye: "norm(z - y) < e" |
72531 | 1707 |
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
|
1708 |
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
|
1709 |
by (metis \<open>\<And>n. closed (F n)\<close> image_iff) |
69313 | 1710 |
have eq: "{w..z} \<inter> \<Inter>(F ` UNIV) = {}" |
72531 | 1711 |
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
|
1712 |
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
|
1713 |
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
|
1714 |
then have "K \<noteq> {}" |
69313 | 1715 |
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
|
1716 |
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
|
1717 |
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
|
1718 |
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
|
1719 |
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
|
1720 |
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
|
1721 |
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
|
1722 |
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
|
1723 |
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
|
1724 |
case True |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1725 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1726 |
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
|
1727 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1728 |
case False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1729 |
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
|
1730 |
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
|
1731 |
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
|
1732 |
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
|
1733 |
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
|
1734 |
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
|
1735 |
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
|
1736 |
with False show thesis |
72531 | 1737 |
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
|
1738 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1739 |
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
|
1740 |
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
|
1741 |
case True |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1742 |
have "z \<in> {w..z}" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1743 |
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
|
1744 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1745 |
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
|
1746 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1747 |
case False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1748 |
show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1749 |
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
|
1750 |
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
|
1751 |
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
|
1752 |
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
|
1753 |
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
|
1754 |
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 | 1755 |
proof |
1756 |
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 = {}" |
|
1757 |
using False by (auto simp: open_segment_eq_real_ivl less_eq_real_def) |
|
1758 |
qed auto |
|
64790
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 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1761 |
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
|
1762 |
proof |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1763 |
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
|
1764 |
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
|
1765 |
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
|
1766 |
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
|
1767 |
show "p u = p v" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1768 |
proof (rule peq) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1769 |
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
|
1770 |
by (auto simp: u v) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1771 |
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
|
1772 |
proof - |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1773 |
have "\<xi> \<notin> {z..v}" |
71633 | 1774 |
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
|
1775 |
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
|
1776 |
by (metis equals0D wzF_null) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1777 |
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
|
1778 |
using that by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1779 |
then show ?thesis |
71633 | 1780 |
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
|
1781 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1782 |
moreover |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1783 |
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
|
1784 |
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
|
1785 |
ultimately |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1786 |
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
|
1787 |
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
|
1788 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1789 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1790 |
then show ?case |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1791 |
using e [of x u] e [of y v] xy |
72531 | 1792 |
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
|
1793 |
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
|
1794 |
show ?thesis |
73932
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents:
72531
diff
changeset
|
1795 |
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
|
1796 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1797 |
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
|
1798 |
qed (meson \<phi>_def) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1799 |
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
|
1800 |
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
|
1801 |
unfolding \<phi>_def by metis+ |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1802 |
then have "T \<noteq> {}" by auto |
67613 | 1803 |
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
|
1804 |
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
|
1805 |
for x y z |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1806 |
proof (cases "x \<in> T") |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1807 |
case True |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1808 |
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
|
1809 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1810 |
case False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1811 |
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
|
1812 |
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
|
1813 |
moreover have "open_segment y z \<inter> T \<subseteq> insert x (open_segment x y \<union> open_segment x z) \<inter> T" |
72531 | 1814 |
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
|
1815 |
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
|
1816 |
by blast |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1817 |
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
|
1818 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1819 |
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
|
1820 |
using that unfolding h_def |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1821 |
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
|
1822 |
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
|
1823 |
by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1824 |
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
|
1825 |
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
|
1826 |
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
|
1827 |
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
|
1828 |
case True |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1829 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1830 |
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
|
1831 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1832 |
case False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1833 |
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
|
1834 |
"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
|
1835 |
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
|
1836 |
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
|
1837 |
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
|
1838 |
ultimately show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1839 |
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
|
1840 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1841 |
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
|
1842 |
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
|
1843 |
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
|
1844 |
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
|
1845 |
fix u \<epsilon>::real |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1846 |
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
|
1847 |
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
|
1848 |
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
|
1849 |
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
|
1850 |
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
|
1851 |
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
|
1852 |
case True |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1853 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1854 |
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
|
1855 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1856 |
case False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1857 |
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
|
1858 |
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
|
1859 |
obtain w where "w \<in> T" and w: "w \<in> closed_segment u v" "open_segment u w \<inter> T = {}" |
72531 | 1860 |
proof (rule segment_to_point_exists [OF uvT]) |
1861 |
fix b |
|
1862 |
assume "b \<in> closed_segment u v \<inter> T" "open_segment u b \<inter> (closed_segment u v \<inter> T) = {}" |
|
1863 |
then show thesis |
|
1864 |
by (metis IntD1 IntD2 ends_in_segment(1) inf.orderE inf_assoc subset_oc_segment that) |
|
1865 |
qed |
|
64790
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1866 |
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
|
1867 |
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
|
1868 |
obtain z where "z \<in> T" and z: "z \<in> closed_segment u v" "open_segment v z \<inter> T = {}" |
72531 | 1869 |
proof (rule segment_to_point_exists [OF uvT]) |
1870 |
fix b |
|
1871 |
assume "b \<in> closed_segment u v \<inter> T" "open_segment v b \<inter> (closed_segment u v \<inter> T) = {}" |
|
1872 |
then show thesis |
|
1873 |
by (metis IntD1 IntD2 ends_in_segment(2) inf.orderE inf_assoc subset_oc_segment that) |
|
1874 |
qed |
|
64790
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1875 |
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
|
1876 |
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
|
1877 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1878 |
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
|
1879 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1880 |
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
|
1881 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1882 |
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
|
1883 |
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
|
1884 |
fix u v |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1885 |
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
|
1886 |
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
|
1887 |
proof (intro exI conjI) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1888 |
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
|
1889 |
by simp |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1890 |
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
|
1891 |
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
|
1892 |
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
|
1893 |
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
|
1894 |
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
|
1895 |
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
|
1896 |
for x y |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1897 |
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
|
1898 |
case True |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1899 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1900 |
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
|
1901 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1902 |
case False |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1903 |
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
|
1904 |
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
|
1905 |
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
|
1906 |
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
|
1907 |
proof |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1908 |
assume "?xuyv" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1909 |
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
|
1910 |
using disjT by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1911 |
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
|
1912 |
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
|
1913 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1914 |
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
|
1915 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1916 |
assume "?yuxv" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1917 |
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
|
1918 |
using disjT by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1919 |
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
|
1920 |
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
|
1921 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1922 |
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
|
1923 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1924 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1925 |
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
|
1926 |
proof (rule T) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1927 |
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
|
1928 |
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
|
1929 |
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
|
1930 |
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
|
1931 |
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
|
1932 |
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
|
1933 |
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
|
1934 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1935 |
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
|
1936 |
by blast |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1937 |
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
|
1938 |
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
|
1939 |
qed auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1940 |
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
|
1941 |
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
|
1942 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1943 |
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
|
1944 |
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
|
1945 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1946 |
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
|
1947 |
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
|
1948 |
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
|
1949 |
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
|
1950 |
then have "arc g" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1951 |
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
|
1952 |
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:
72531
diff
changeset
|
1953 |
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
|
1954 |
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
|
1955 |
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
|
1956 |
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
|
1957 |
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
|
1958 |
show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1959 |
proof |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1960 |
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
|
1961 |
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
|
1962 |
show "path_image (subpath u v g) \<subseteq> path_image p" |
71633 | 1963 |
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
|
1964 |
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
|
1965 |
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
|
1966 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1967 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1968 |
|
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1969 |
|
69652
3417a8f154eb
updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69517
diff
changeset
|
1970 |
corollary path_connected_arcwise: |
64790
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1971 |
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
|
1972 |
shows "path_connected S \<longleftrightarrow> |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1973 |
(\<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
|
1974 |
(is "?lhs = ?rhs") |
69652
3417a8f154eb
updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69517
diff
changeset
|
1975 |
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
|
1976 |
fix x y |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1977 |
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
|
1978 |
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
|
1979 |
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
|
1980 |
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
|
1981 |
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
|
1982 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1983 |
assume R [rule_format]: ?rhs |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1984 |
show ?lhs |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1985 |
unfolding path_connected_def |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1986 |
proof (intro ballI) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1987 |
fix x y |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1988 |
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
|
1989 |
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
|
1990 |
proof (cases "x = y") |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1991 |
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
|
1992 |
by blast |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1993 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1994 |
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
|
1995 |
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
|
1996 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1997 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1998 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1999 |
|
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2000 |
|
69652
3417a8f154eb
updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69517
diff
changeset
|
2001 |
corollary arc_connected_trans: |
64790
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2002 |
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
|
2003 |
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
|
2004 |
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
|
2005 |
"pathstart i = pathstart g" "pathfinish i = pathfinish h" |
73932
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents:
72531
diff
changeset
|
2006 |
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
|
2007 |
|
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2008 |
|
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2009 |
|
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2010 |
|
69683 | 2011 |
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
|
2012 |
|
69652
3417a8f154eb
updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69517
diff
changeset
|
2013 |
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
|
2014 |
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:
69745
diff
changeset
|
2015 |
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
|
2016 |
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:
69517
diff
changeset
|
2017 |
proof - |
64790
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2018 |
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
|
2019 |
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
|
2020 |
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
|
2021 |
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
|
2022 |
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
|
2023 |
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
|
2024 |
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
|
2025 |
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
|
2026 |
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
|
2027 |
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
|
2028 |
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
|
2029 |
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
|
2030 |
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
|
2031 |
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
|
2032 |
have "y \<noteq> z" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2033 |
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
|
2034 |
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
|
2035 |
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
|
2036 |
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
|
2037 |
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
|
2038 |
using \<open>y \<noteq> z\<close> by (force simp: path_connected_arcwise) |
72531 | 2039 |
have "continuous_on {0..1} g" |
2040 |
using \<open>arc g\<close> arc_imp_path path_def by blast |
|
2041 |
then have "compact (g -` frontier S \<inter> {0..1})" |
|
2042 |
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
|
2043 |
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
|
2044 |
proof - |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2045 |
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
|
2046 |
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
|
2047 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2048 |
by blast |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2049 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2050 |
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
|
2051 |
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
|
2052 |
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
|
2053 |
moreover have "t \<noteq> 0" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2054 |
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
|
2055 |
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
|
2056 |
by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2057 |
have "V \<subseteq> frontier S" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2058 |
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
|
2059 |
show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2060 |
proof |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2061 |
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
|
2062 |
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
|
2063 |
have "g 0 \<in> S" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2064 |
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
|
2065 |
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
|
2066 |
by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2067 |
have "g t \<in> V" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2068 |
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
|
2069 |
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
|
2070 |
by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2071 |
then have "inj_on (subpath 0 t g) {0..1}" |
72531 | 2072 |
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
|
2073 |
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
|
2074 |
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
|
2075 |
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
|
2076 |
proof - |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2077 |
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
|
2078 |
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
|
2079 |
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
|
2080 |
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
|
2081 |
show "connected (subpath 0 t g ` {0..<1})" |
72531 | 2082 |
proof (rule connected_continuous_image) |
2083 |
show "continuous_on {0..<1} (subpath 0 t g)" |
|
2084 |
by (meson \<open>arc (subpath 0 t g)\<close> arc_def atLeastLessThan_subseteq_atLeastAtMost_iff continuous_on_subset order_refl path_def) |
|
2085 |
qed auto |
|
64790
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2086 |
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
|
2087 |
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
|
2088 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2089 |
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
|
2090 |
by blast |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2091 |
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
|
2092 |
by (fastforce simp: subpath_def) |
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 |
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
|
2095 |
using subsetD by fastforce |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2096 |
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
|
2097 |
by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2098 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2099 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2100 |
|
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2101 |
|
69652
3417a8f154eb
updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69517
diff
changeset
|
2102 |
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
|
2103 |
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
|
2104 |
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:
69745
diff
changeset
|
2105 |
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
|
2106 |
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:
69517
diff
changeset
|
2107 |
proof - |
64790
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2108 |
have "V \<subseteq> frontier S" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2109 |
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
|
2110 |
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
|
2111 |
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
|
2112 |
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
|
2113 |
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
|
2114 |
then have "path_connected S" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2115 |
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
|
2116 |
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
|
2117 |
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
|
2118 |
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
|
2119 |
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
|
2120 |
then have "path (f +++ g)" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2121 |
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
|
2122 |
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
|
2123 |
and h: "path_image h \<subseteq> path_image (f +++ g)" "pathstart h = x" "pathfinish h = pathfinish g" |
72531 | 2124 |
using path_contains_arc [of "f +++ g" x "pathfinish g"] \<open>x \<notin> V\<close> \<open>pathfinish g \<in> V\<close> f |
2125 |
by (metis pathfinish_join pathstart_join) |
|
2126 |
have "path_image h \<subseteq> path_image f \<union> path_image g" |
|
2127 |
using h(1) path_image_join_subset by auto |
|
2128 |
then have "h ` {0..1} - {h 1} \<subseteq> S" |
|
2129 |
using f g h |
|
64790
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2130 |
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
|
2131 |
by (metis le_less) |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2132 |
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
|
2133 |
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
|
2134 |
then show thesis |
72531 | 2135 |
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
|
2136 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2137 |
|
69652
3417a8f154eb
updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69517
diff
changeset
|
2138 |
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
|
2139 |
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
|
2140 |
assumes S: "open S" "connected S" |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69745
diff
changeset
|
2141 |
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:
69745
diff
changeset
|
2142 |
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
|
2143 |
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
|
2144 |
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:
69517
diff
changeset
|
2145 |
proof - |
64790
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2146 |
have "S \<noteq> {}" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2147 |
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
|
2148 |
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
|
2149 |
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
|
2150 |
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
|
2151 |
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
|
2152 |
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
|
2153 |
show "U - {pathfinish g} \<noteq> {}" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2154 |
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:
69745
diff
changeset
|
2155 |
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
|
2156 |
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
|
2157 |
qed auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2158 |
obtain \<gamma> where "arc \<gamma>" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2159 |
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
|
2160 |
"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
|
2161 |
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
|
2162 |
show "path (reversepath h +++ g)" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2163 |
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
|
2164 |
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
|
2165 |
"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
|
2166 |
by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2167 |
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
|
2168 |
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
|
2169 |
qed auto |
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 "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
|
2173 |
using \<gamma> \<open>arc \<gamma>\<close> \<open>pathfinish h \<in> U - {pathfinish g}\<close> \<open>pathfinish g \<in> V\<close> by auto |
72531 | 2174 |
have "path_image \<gamma> \<subseteq> path_image h \<union> path_image g" |
2175 |
by (metis \<gamma>(1) g(2) h(2) path_image_join path_image_reversepath pathfinish_reversepath) |
|
2176 |
then have "\<gamma> ` {0..1} - {\<gamma> 0, \<gamma> 1} \<subseteq> S" |
|
2177 |
using \<gamma> g h |
|
64790
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2178 |
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
|
2179 |
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
|
2180 |
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
|
2181 |
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
|
2182 |
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
|
2183 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2184 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2185 |
|
69652
3417a8f154eb
updated tagging first 5
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69517
diff
changeset
|
2186 |
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
|
2187 |
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
|
2188 |
assumes S: "open S" "connected S" |
69922
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents:
69745
diff
changeset
|
2189 |
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:
69745
diff
changeset
|
2190 |
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
|
2191 |
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
|
2192 |
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:
69517
diff
changeset
|
2193 |
proof - |
64790
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2194 |
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
|
2195 |
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
|
2196 |
then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2197 |
proof cases |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2198 |
case 1 then show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2199 |
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
|
2200 |
next |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2201 |
case 2 |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2202 |
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
|
2203 |
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
|
2204 |
show ?thesis |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2205 |
proof |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2206 |
show "arc (reversepath g)" |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2207 |
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
|
2208 |
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
|
2209 |
using g by auto |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2210 |
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
|
2211 |
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
|
2212 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2213 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2214 |
qed |
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2215 |
|
ed38f9a834d8
New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2216 |
end |