| author | immler | 
| Sat, 29 Sep 2018 16:30:44 -0400 | |
| changeset 69096 | 62a0d10386c1 | 
| parent 68833 | fde093888c16 | 
| child 69313 | b021008c5397 | 
| 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  | 
|
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
5  | 
section \<open>Arcwise-connected sets\<close>  | 
| 
 
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  | 
| 
66453
 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 
wenzelm 
parents: 
65417 
diff
changeset
 | 
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  | 
|
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
10  | 
begin  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
11  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
12  | 
subsection%important \<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
 | 
13  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
14  | 
theorem%important Brouwer_reduction_theorem_gen:  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
15  | 
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
 | 
16  | 
assumes "closed S" "\<phi> S"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
17  | 
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)"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
18  | 
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)"  | 
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
19  | 
proof%unimportant -  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
20  | 
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
 | 
21  | 
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
 | 
22  | 
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
 | 
23  | 
  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 | 24  | 
                                        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
 | 
25  | 
else a)"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
26  | 
have [simp]: "A 0 = S"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
27  | 
by (simp add: A_def)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
28  | 
  have ASuc: "A(Suc n) = (if \<exists>U. U \<subseteq> A n \<and> closed U \<and> \<phi> U \<and> U \<inter> (B n) = {}
 | 
| 67613 | 29  | 
                          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
 | 
30  | 
else A n)" for n  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
31  | 
by (auto simp: A_def)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
32  | 
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
 | 
33  | 
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
 | 
34  | 
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
 | 
35  | 
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
 | 
36  | 
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
 | 
37  | 
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
 | 
38  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
39  | 
proof  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
40  | 
show "\<Inter>range A \<subseteq> S"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
41  | 
using \<open>\<And>n. A n \<subseteq> S\<close> by blast  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
42  | 
show "closed (INTER UNIV A)"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
43  | 
using clo by blast  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
44  | 
show "\<phi> (INTER UNIV A)"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
45  | 
by (simp add: clo \<phi> sub)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
46  | 
show "\<not> U \<subset> INTER UNIV A" if "U \<subseteq> S" "closed U" "\<phi> U" for U  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
47  | 
proof -  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
48  | 
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
 | 
49  | 
proof -  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
50  | 
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
 | 
51  | 
using \<open>closed U\<close> \<open>x \<notin> U\<close> openE [of "-U"] by blast  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
52  | 
moreover obtain K where K: "ball x e = UNION K B"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
53  | 
using open_cov [of "ball x e"] by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
54  | 
ultimately have "UNION K B \<subseteq> -U"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
55  | 
by blast  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
56  | 
        have "K \<noteq> {}"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
57  | 
using \<open>0 < e\<close> \<open>ball x e = UNION K B\<close> by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
58  | 
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
 | 
59  | 
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
 | 
60  | 
        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
 | 
61  | 
using K e by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
62  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
63  | 
        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
 | 
64  | 
case True  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
65  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
66  | 
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
 | 
67  | 
apply (simp add: ASuc)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
68  | 
apply (erule someI2_ex)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
69  | 
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
 | 
70  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
71  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
72  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
73  | 
            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
 | 
74  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
75  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
76  | 
with that show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
77  | 
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
 | 
78  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
79  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
80  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
81  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
82  | 
corollary%important Brouwer_reduction_theorem:  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
83  | 
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
 | 
84  | 
  assumes "compact S" "\<phi> S" "S \<noteq> {}"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
85  | 
      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)"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
86  | 
  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
 | 
87  | 
                  "\<And>U. \<lbrakk>U \<subseteq> S; closed U; U \<noteq> {}; \<phi> U\<rbrakk> \<Longrightarrow> \<not> (U \<subset> T)"
 | 
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
88  | 
proof%unimportant (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
 | 
89  | 
fix F  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
90  | 
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
 | 
91  | 
     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"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
92  | 
  show "INTER UNIV F \<noteq> {} \<and> INTER UNIV F \<subseteq> S \<and> \<phi> (INTER UNIV F)"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
93  | 
proof (intro conjI)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
94  | 
    show "INTER UNIV F \<noteq> {}"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
95  | 
apply (rule compact_nest)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
96  | 
apply (meson F cloF \<open>compact S\<close> 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
 | 
97  | 
apply (simp add: F)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
98  | 
by (meson Fsub lift_Suc_antimono_le)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
99  | 
show " INTER UNIV F \<subseteq> S"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
100  | 
using F by blast  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
101  | 
show "\<phi> (INTER UNIV F)"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
102  | 
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
 | 
103  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
104  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
105  | 
  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
 | 
106  | 
by (simp add: assms)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
107  | 
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
 | 
108  | 
|
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
109  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
110  | 
subsection%important\<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
 | 
111  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
112  | 
subsection%important\<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
 | 
113  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
114  | 
proposition%important closure_dyadic_rationals:  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
115  | 
"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
 | 
116  | 
                   { \<Sum>i :: 'a :: euclidean_space \<in> Basis. (f i / 2^k) *\<^sub>R i }) = UNIV"
 | 
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
117  | 
proof%unimportant -  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
118  | 
  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
 | 
119  | 
proof (clarsimp simp: closure_approachable)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
120  | 
fix e::real  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
121  | 
assume "e > 0"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
122  | 
    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
 | 
123  | 
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
 | 
124  | 
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
 | 
125  | 
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
 | 
126  | 
by (simp add: euclidean_representation)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
127  | 
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
 | 
128  | 
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
 | 
129  | 
    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
 | 
130  | 
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
 | 
131  | 
fix i::'a  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
132  | 
assume "i \<in> Basis"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
133  | 
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
 | 
134  | 
\<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
 | 
135  | 
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
 | 
136  | 
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
 | 
137  | 
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
 | 
138  | 
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
 | 
139  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
140  | 
    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
 | 
141  | 
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
 | 
142  | 
also have "... = e"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
143  | 
by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
144  | 
finally have "dist (\<Sum>i\<in>Basis. (\<lfloor>2^k*(x \<bullet> i)\<rfloor> / 2^k) *\<^sub>R i) x < e" .  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
145  | 
then  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
146  | 
show "\<exists>k. \<exists>f \<in> Basis \<rightarrow> \<int>. dist (\<Sum>b\<in>Basis. (f b / 2^k) *\<^sub>R b) x < e"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
147  | 
apply (rule_tac x=k in exI)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
148  | 
apply (rule_tac x="\<lambda>i. of_int (floor (2^k*(x \<bullet> i)))" in bexI)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
149  | 
apply auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
150  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
151  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
152  | 
then show ?thesis by auto  | 
| 
 
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  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
155  | 
corollary%important closure_rational_coordinates:  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
156  | 
    "closure (\<Union>f \<in> Basis \<rightarrow> \<rat>. { \<Sum>i :: 'a :: euclidean_space \<in> Basis. f i *\<^sub>R i }) = UNIV"
 | 
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
157  | 
proof%unimportant -  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
158  | 
  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
 | 
159  | 
           \<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
 | 
160  | 
proof clarsimp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
161  | 
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
 | 
162  | 
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
 | 
163  | 
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
 | 
164  | 
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
 | 
165  | 
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
 | 
166  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
167  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
168  | 
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
 | 
169  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
170  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
171  | 
lemma%unimportant 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
 | 
172  | 
   "\<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
 | 
173  | 
\<Longrightarrow> closure(S \<inter>  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
174  | 
(\<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
 | 
175  | 
                   { \<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
 | 
176  | 
closure S"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
177  | 
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
 | 
178  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
179  | 
lemma%unimportant 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
 | 
180  | 
   "\<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
 | 
181  | 
    \<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
 | 
182  | 
closure S"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
183  | 
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
 | 
184  | 
|
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
185  | 
|
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
186  | 
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
 | 
187  | 
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
 | 
188  | 
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
 | 
189  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
190  | 
lemma%important 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
 | 
191  | 
fixes S :: "real set"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
192  | 
  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
 | 
193  | 
    shows "closure(S \<inter> (\<Union>k m. {of_nat m / 2^k})) = closure S"
 | 
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
194  | 
proof%unimportant -  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
195  | 
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
 | 
196  | 
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
 | 
197  | 
  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
 | 
198  | 
             (\<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
 | 
199  | 
by force  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
200  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
201  | 
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
 | 
202  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
203  | 
|
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
204  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
205  | 
definition%important 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
 | 
206  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
207  | 
lemma%unimportant real_in_dyadics [simp]: "real m \<in> dyadics"  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
208  | 
apply (simp add: dyadics_def)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
209  | 
by (metis divide_numeral_1 numeral_One power_0)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
210  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
211  | 
lemma%unimportant 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
 | 
212  | 
proof  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
213  | 
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
 | 
214  | 
then have "of_nat (m * (2 * 2^n)) = (of_nat (Suc (4 * k)) :: 'a)"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
215  | 
by (simp add: divide_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
216  | 
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
 | 
217  | 
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
 | 
218  | 
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
 | 
219  | 
by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
220  | 
then show False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
221  | 
by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
222  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
223  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
224  | 
lemma%important nat_neq_4k3: "of_nat m \<noteq> (4 * of_nat k + 3) / (2 * 2^n :: 'a::field_char_0)"  | 
| 
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
225  | 
proof%unimportant  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
226  | 
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
 | 
227  | 
then have "of_nat (m * (2 * 2^n)) = (of_nat (4 * k + 3) :: 'a)"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
228  | 
by (simp add: divide_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
229  | 
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
 | 
230  | 
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
 | 
231  | 
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
 | 
232  | 
by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
233  | 
then show False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
234  | 
by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
235  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
236  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
237  | 
lemma%important iff_4k:  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
238  | 
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
 | 
239  | 
shows "(4 * real m + r) / (2 * 2^n) = (4 * real m' + r) / (2 * 2 ^ n') \<longleftrightarrow> m=m' \<and> n=n'"  | 
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
240  | 
proof%unimportant -  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
241  | 
  { 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
 | 
242  | 
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
 | 
243  | 
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
 | 
244  | 
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
 | 
245  | 
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
 | 
246  | 
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
 | 
247  | 
by linarith  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
248  | 
then obtain "4*m + k = 4*m' + k" "n=n'"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
249  | 
apply (rule prime_power_cancel2 [OF two_is_prime_nat])  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
250  | 
using assms by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
251  | 
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
 | 
252  | 
by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
253  | 
}  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
254  | 
then show ?thesis by blast  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
255  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
256  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
257  | 
lemma%important neq_4k1_k43: "(4 * real m + 1) / (2 * 2^n) \<noteq> (4 * real m' + 3) / (2 * 2 ^ n')"  | 
| 
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
258  | 
proof%unimportant  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
259  | 
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
 | 
260  | 
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
 | 
261  | 
by (auto simp: field_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
262  | 
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
 | 
263  | 
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
 | 
264  | 
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
 | 
265  | 
by linarith  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
266  | 
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
 | 
267  | 
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
 | 
268  | 
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
 | 
269  | 
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
 | 
270  | 
then show False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
271  | 
using even_Suc by presburger  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
272  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
273  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
274  | 
lemma%important dyadic_413_cases:  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
275  | 
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
 | 
276  | 
| 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
 | 
277  | 
| m' k' where "k' < k" "(of_nat m:: 'a) / 2^k = of_nat (4*m' + 3) / 2^Suc k'"  | 
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
278  | 
proof%unimportant (cases "m>0")  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
279  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
280  | 
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
 | 
281  | 
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
 | 
282  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
283  | 
case True  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
284  | 
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
 | 
285  | 
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
 | 
286  | 
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
 | 
287  | 
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
 | 
288  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
289  | 
proof (cases "k \<le> k'")  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
290  | 
case True  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
291  | 
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
 | 
292  | 
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
 | 
293  | 
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
 | 
294  | 
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
 | 
295  | 
also have "... \<in> \<nat>"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
296  | 
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
 | 
297  | 
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
 | 
298  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
299  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
300  | 
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
 | 
301  | 
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
 | 
302  | 
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
 | 
303  | 
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
 | 
304  | 
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
 | 
305  | 
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
 | 
306  | 
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
 | 
307  | 
using q by force  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
308  | 
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
 | 
309  | 
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
 | 
310  | 
using q m' by presburger+  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
311  | 
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
 | 
312  | 
by linarith  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
313  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
314  | 
proof cases  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
315  | 
assume "r = 1"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
316  | 
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
 | 
317  | 
by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
318  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
319  | 
assume "r = 3"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
320  | 
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
 | 
321  | 
by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
322  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
323  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
324  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
325  | 
|
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
326  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
327  | 
lemma%important dyadics_iff:  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
328  | 
"(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
 | 
329  | 
    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
 | 
330  | 
(is "_ = ?rhs")  | 
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
331  | 
proof%unimportant  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
332  | 
show "dyadics \<subseteq> ?rhs"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
333  | 
unfolding dyadics_def  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
334  | 
apply clarify  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
335  | 
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
 | 
336  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
337  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
338  | 
show "?rhs \<subseteq> dyadics"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
339  | 
apply (clarsimp simp: dyadics_def Nats_def simp del: power_Suc)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
340  | 
apply (intro conjI subsetI)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
341  | 
apply (auto simp del: power_Suc)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
342  | 
apply (metis divide_numeral_1 numeral_One power_0)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
343  | 
apply (metis 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
 | 
344  | 
by (metis of_nat_add 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
 | 
345  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
346  | 
|
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
347  | 
|
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
348  | 
function (domintros) dyad_rec :: "[nat \<Rightarrow> 'a, 'a\<Rightarrow>'a, 'a\<Rightarrow>'a, real] \<Rightarrow> 'a" where  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
349  | 
"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
 | 
350  | 
| "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
 | 
351  | 
| "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
 | 
352  | 
| "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
 | 
353  | 
using iff_4k [of _ 1] iff_4k [of _ 3]  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
354  | 
apply (simp_all add: nat_neq_4k1 nat_neq_4k3 neq_4k1_k43, atomize_elim)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
355  | 
apply (fastforce simp add: dyadics_iff Nats_def field_simps)+  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
356  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
357  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
358  | 
lemma%unimportant 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
 | 
359  | 
unfolding dyadics_def by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
360  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
361  | 
lemma%important dyad_rec_level_termination:  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
362  | 
assumes "k < K"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
363  | 
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
 | 
364  | 
using assms  | 
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
365  | 
proof%unimportant (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
 | 
366  | 
case 0  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
367  | 
then show ?case by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
368  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
369  | 
case (Suc K)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
370  | 
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
 | 
371  | 
using less_antisym by blast  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
372  | 
then show ?case  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
373  | 
proof cases  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
374  | 
assume "k = K"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
375  | 
show ?case  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
376  | 
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
 | 
377  | 
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
 | 
378  | 
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
 | 
379  | 
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
 | 
380  | 
proof -  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
381  | 
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
 | 
382  | 
proof (rule dyad_rec.domintros)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
383  | 
fix m n  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
384  | 
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
 | 
385  | 
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
 | 
386  | 
by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
387  | 
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
 | 
388  | 
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
 | 
389  | 
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
 | 
390  | 
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
 | 
391  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
392  | 
fix m n  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
393  | 
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
 | 
394  | 
then have "False"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
395  | 
by (metis neq_4k1_k43)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
396  | 
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
 | 
397  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
398  | 
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
 | 
399  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
400  | 
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
 | 
401  | 
proof -  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
402  | 
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
 | 
403  | 
proof (rule dyad_rec.domintros)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
404  | 
fix m n  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
405  | 
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
 | 
406  | 
then have "False"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
407  | 
by (metis neq_4k1_k43)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
408  | 
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
 | 
409  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
410  | 
fix m n  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
411  | 
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
 | 
412  | 
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
 | 
413  | 
by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
414  | 
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
 | 
415  | 
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
 | 
416  | 
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
 | 
417  | 
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
 | 
418  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
419  | 
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
 | 
420  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
421  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
422  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
423  | 
assume "k < K"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
424  | 
then show ?case  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
425  | 
using Suc.IH by blast  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
426  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
427  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
428  | 
|
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
429  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
430  | 
lemma%unimportant 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
 | 
431  | 
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
 | 
432  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
433  | 
lemma%unimportant 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
 | 
434  | 
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
 | 
435  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
436  | 
lemma%unimportant 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))"  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
437  | 
apply (rule dyad_rec.psimps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
438  | 
by (metis dyad_rec_level_termination lessI add.commute 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
 | 
439  | 
|
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
440  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
441  | 
lemma%unimportant 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))"  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
442  | 
apply (rule dyad_rec.psimps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
443  | 
by (metis dyad_rec_level_termination lessI of_nat_add 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
 | 
444  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
445  | 
lemma%unimportant dyad_rec_41_times2:  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
446  | 
assumes "n > 0"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
447  | 
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
 | 
448  | 
proof -  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
449  | 
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
 | 
450  | 
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
 | 
451  | 
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
 | 
452  | 
by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
453  | 
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
 | 
454  | 
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
 | 
455  | 
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
 | 
456  | 
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
 | 
457  | 
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
 | 
458  | 
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
 | 
459  | 
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
 | 
460  | 
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
 | 
461  | 
finally show ?thesis .  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
462  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
463  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
464  | 
lemma%important dyad_rec_43_times2:  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
465  | 
assumes "n > 0"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
466  | 
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))"  | 
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
467  | 
proof%unimportant -  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
468  | 
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
 | 
469  | 
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
 | 
470  | 
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
 | 
471  | 
by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
472  | 
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
 | 
473  | 
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
 | 
474  | 
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
 | 
475  | 
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
 | 
476  | 
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
 | 
477  | 
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
 | 
478  | 
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
 | 
479  | 
by (simp add: n')  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
480  | 
finally show ?thesis .  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
481  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
482  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
483  | 
definition%important dyad_rec2  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
484  | 
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
 | 
485  | 
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
 | 
486  | 
|
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
487  | 
abbreviation leftrec where "leftrec u v lc rc x \<equiv> fst (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
 | 
488  | 
abbreviation rightrec where "rightrec u v lc rc x \<equiv> snd (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
 | 
489  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
490  | 
lemma%unimportant 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
 | 
491  | 
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
 | 
492  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
493  | 
lemma%unimportant 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)"  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
494  | 
apply (simp only: dyad_rec2_def dyad_rec_41_times2)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
495  | 
apply (simp add: case_prod_beta)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
496  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
497  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
498  | 
lemma%unimportant 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
 | 
499  | 
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
 | 
500  | 
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
 | 
501  | 
(midpoint (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
 | 
502  | 
apply (simp only: dyad_rec2_def dyad_rec_43_times2)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
503  | 
apply (simp add: case_prod_beta)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
504  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
505  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
506  | 
lemma%unimportant 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
 | 
507  | 
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
 | 
508  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
509  | 
lemma%unimportant 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
 | 
510  | 
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
 | 
511  | 
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
 | 
512  | 
(midpoint (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  | 
apply (simp only: dyad_rec2_def dyad_rec_41_times2)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
514  | 
apply (simp add: case_prod_beta)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
515  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
516  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
517  | 
lemma%unimportant 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)"  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
518  | 
apply (simp only: dyad_rec2_def dyad_rec_43_times2)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
519  | 
apply (simp add: case_prod_beta)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
520  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
521  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
522  | 
lemma%unimportant 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
 | 
523  | 
   "{0<..<1} \<inter> (\<Union>k m. {real m / 2^k}) = (\<Union>k. \<Union>m \<in> {0<..<2^k}. {real m / 2^k})"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
524  | 
by (auto simp: divide_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
525  | 
|
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
526  | 
|
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
527  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
528  | 
theorem%important homeomorphic_monotone_image_interval:  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
529  | 
  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
 | 
530  | 
  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
 | 
531  | 
      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
 | 
532  | 
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
 | 
533  | 
    shows "(f ` {0..1}) homeomorphic {0..1::real}"
 | 
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
534  | 
proof%unimportant -  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
535  | 
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
 | 
536  | 
              (\<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
 | 
537  | 
              (\<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
 | 
538  | 
              (\<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
 | 
539  | 
              (\<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
 | 
540  | 
    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
 | 
541  | 
proof -  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
542  | 
    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
 | 
543  | 
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
 | 
544  | 
    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
 | 
545  | 
      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
 | 
546  | 
by (metis Int_commute)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
547  | 
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
 | 
548  | 
by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
549  | 
    obtain c d where cd: "{a..b} \<inter> f -` {f m} = {c..d}"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
550  | 
apply (rule_tac c="max a c0" and d="min b d0" in that)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
551  | 
using ab01 cd0 by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
552  | 
    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
 | 
553  | 
by blast  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
554  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
555  | 
proof (intro exI conjI ballI)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
556  | 
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
 | 
557  | 
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
 | 
558  | 
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
 | 
559  | 
using cd m by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
560  | 
      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
 | 
561  | 
using cd by blast  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
562  | 
      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
 | 
563  | 
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
 | 
564  | 
      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
 | 
565  | 
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
 | 
566  | 
      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
 | 
567  | 
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
 | 
568  | 
case True  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
569  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
570  | 
          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
 | 
571  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
572  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
573  | 
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
 | 
574  | 
proof -  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
575  | 
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
 | 
576  | 
            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
 | 
577  | 
          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
 | 
578  | 
            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
 | 
579  | 
          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
 | 
580  | 
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
 | 
581  | 
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
 | 
582  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
583  | 
then show ?thesis by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
584  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
585  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
586  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
587  | 
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
 | 
588  | 
    "\<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
 | 
589  | 
(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
 | 
590  | 
            (\<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
 | 
591  | 
            (\<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
 | 
592  | 
            (\<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
 | 
593  | 
            (\<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
 | 
594  | 
apply atomize  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
595  | 
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
 | 
596  | 
apply (rule that, blast)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
597  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
598  | 
  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"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
599  | 
    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"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
600  | 
by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
601  | 
  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
 | 
602  | 
    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
 | 
603  | 
    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
 | 
604  | 
    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
 | 
605  | 
\<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
 | 
606  | 
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
 | 
607  | 
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
 | 
608  | 
                             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
 | 
609  | 
by (metis feqm)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
610  | 
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
 | 
611  | 
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
 | 
612  | 
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
 | 
613  | 
  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
 | 
614  | 
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
 | 
615  | 
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
 | 
616  | 
  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
 | 
617  | 
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
 | 
618  | 
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
 | 
619  | 
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
 | 
620  | 
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
 | 
621  | 
  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
 | 
622  | 
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
 | 
623  | 
  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
 | 
624  | 
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
 | 
625  | 
  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
 | 
626  | 
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
 | 
627  | 
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
 | 
628  | 
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
 | 
629  | 
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
 | 
630  | 
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
 | 
631  | 
using a_def leftrec_base  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
632  | 
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
 | 
633  | 
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
 | 
634  | 
using b_def rightrec_base  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
635  | 
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
 | 
636  | 
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
 | 
637  | 
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
 | 
638  | 
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
 | 
639  | 
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
 | 
640  | 
(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
 | 
641  | 
(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
 | 
642  | 
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
 | 
643  | 
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
 | 
644  | 
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
 | 
645  | 
(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
 | 
646  | 
(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
 | 
647  | 
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
 | 
648  | 
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
 | 
649  | 
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
 | 
650  | 
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
 | 
651  | 
proof (induction n arbitrary: m)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
652  | 
case 0  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
653  | 
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
 | 
654  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
655  | 
case (Suc n p)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
656  | 
show ?case  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
657  | 
proof (cases "even p")  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
658  | 
case True  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
659  | 
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
 | 
660  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
661  | 
by (simp add: Suc.IH)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
662  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
663  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
664  | 
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
 | 
665  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
666  | 
proof (cases n)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
667  | 
case 0  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
668  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
669  | 
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
 | 
670  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
671  | 
case (Suc n')  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
672  | 
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
 | 
673  | 
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
 | 
674  | 
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
 | 
675  | 
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
 | 
676  | 
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
 | 
677  | 
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
 | 
678  | 
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
 | 
679  | 
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
 | 
680  | 
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
 | 
681  | 
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
 | 
682  | 
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
 | 
683  | 
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
 | 
684  | 
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
 | 
685  | 
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
 | 
686  | 
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
 | 
687  | 
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
 | 
688  | 
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
 | 
689  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
690  | 
proof (cases "even m")  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
691  | 
case True  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
692  | 
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
 | 
693  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
694  | 
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"]  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
695  | 
Suc.IH [of "m+1"]  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
696  | 
apply (simp add: r m add.commute [of 1] \<open>n > 0\<close> a41 b41 del: power_Suc)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
697  | 
apply (auto simp: left_right [THEN conjunct1])  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
698  | 
using order_trans [OF left_le c_le_v]  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
699  | 
by (metis (no_types, hide_lams) add.commute mult_2 of_nat_Suc of_nat_add)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
700  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
701  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
702  | 
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
 | 
703  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
704  | 
using a_le_c [of "m"] c_le_b [of "m"] a_ge_0 [of "m"] b_le_1 [of "m"]  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
705  | 
Suc.IH [of "m+1"]  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
706  | 
apply (simp add: r m add.commute [of 3] \<open>n > 0\<close> a43 b43 del: power_Suc)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
707  | 
apply (auto simp: add.commute left_right [THEN conjunct2])  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
708  | 
using order_trans [OF c_ge_u right_ge]  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
709  | 
apply (metis (no_types, hide_lams) mult_2 numeral_One of_nat_add of_nat_numeral)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
710  | 
apply (metis Suc.IH mult_2 of_nat_1 of_nat_add)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
711  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
712  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
713  | 
qed  | 
| 
 
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  | 
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
 | 
717  | 
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
 | 
718  | 
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
 | 
719  | 
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
 | 
720  | 
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
 | 
721  | 
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
 | 
722  | 
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  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
723  | 
using a_ge_0 alec order_trans apply blast  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
724  | 
by (meson b_le_1 cleb order_trans)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
725  | 
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
 | 
726  | 
\<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
 | 
727  | 
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
 | 
728  | 
case (less d j n)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
729  | 
show ?case  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
730  | 
proof (cases "m \<le> n")  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
731  | 
case True  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
732  | 
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
 | 
733  | 
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
 | 
734  | 
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
 | 
735  | 
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
 | 
736  | 
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
 | 
737  | 
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
 | 
738  | 
also have "... \<in> \<int>"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
739  | 
by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
740  | 
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
 | 
741  | 
with True Ints_abs show "\<bar>2^n\<bar> * \<bar>real i / 2^m - real j / 2^n\<bar> \<in> \<int>"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
742  | 
by (fastforce simp: divide_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
743  | 
show "\<bar>\<bar>2^n\<bar> * \<bar>real i / 2^m - real j / 2^n\<bar>\<bar> < 1"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
744  | 
using less.prems by (auto simp: divide_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
745  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
746  | 
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
 | 
747  | 
by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
748  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
749  | 
by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
750  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
751  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
752  | 
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
 | 
753  | 
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
 | 
754  | 
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
 | 
755  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
756  | 
proof (cases "n > 0")  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
757  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
758  | 
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
 | 
759  | 
by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
760  | 
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
 | 
761  | 
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
 | 
762  | 
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
 | 
763  | 
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
 | 
764  | 
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
 | 
765  | 
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
 | 
766  | 
using False by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
767  | 
finally show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
768  | 
by (auto simp: ac)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
769  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
770  | 
case True show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
771  | 
proof (cases "real i / 2^m" "real j / 2^n" rule: linorder_cases)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
772  | 
case less  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
773  | 
moreover have "real (4 * k + 1) / 2 ^ Suc n + 1 / (2 ^ Suc n) = real j / 2 ^ n"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
774  | 
using k by (force simp: divide_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
775  | 
moreover have "\<bar>real i / 2 ^ m - real j / 2 ^ n\<bar> < 2 / (2 ^ Suc n)"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
776  | 
using less.prems by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
777  | 
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
 | 
778  | 
using less.prems by linarith  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
779  | 
have *: "a (real (4 * k + 1) / 2 ^ Suc n) \<le> c (real i / 2 ^ m) \<and>  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
780  | 
c (real i / 2 ^ m) \<le> b (real (4 * k + 1) / 2 ^ Suc n)"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
781  | 
apply (rule less.IH [OF _ refl])  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
782  | 
using closer \<open>n < m\<close> \<open>d = m - n\<close> apply (auto simp: divide_simps \<open>n < m\<close> diff_less_mono2)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
783  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
784  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
785  | 
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
 | 
786  | 
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"]  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
787  | 
using k a41 b41 * \<open>0 < n\<close>  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
788  | 
apply (simp add: add.commute)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
789  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
790  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
791  | 
case 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
 | 
792  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
793  | 
case greater  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
794  | 
moreover have "real (4 * k + 3) / 2 ^ Suc n - 1 / (2 ^ Suc n) = real j / 2 ^ n"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
795  | 
using k by (force simp: divide_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
796  | 
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
 | 
797  | 
using less.prems by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
798  | 
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
 | 
799  | 
using less.prems by linarith  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
800  | 
have *: "a (real (4 * k + 3) / 2 ^ Suc n) \<le> c (real i / 2 ^ m) \<and>  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
801  | 
c (real i / 2 ^ m) \<le> b (real (4 * k + 3) / 2 ^ Suc n)"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
802  | 
apply (rule less.IH [OF _ refl])  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
803  | 
using closer \<open>n < m\<close> \<open>d = m - n\<close> apply (auto simp: divide_simps \<open>n < m\<close> diff_less_mono2)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
804  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
805  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
806  | 
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
 | 
807  | 
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"]  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
808  | 
using k a43 b43 * \<open>0 < n\<close>  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
809  | 
apply (simp add: add.commute)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
810  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
811  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
812  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
813  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
814  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
815  | 
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
 | 
816  | 
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
 | 
817  | 
using that by blast+  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
818  | 
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
 | 
819  | 
proof (induction n arbitrary: m)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
820  | 
case 0  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
821  | 
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
 | 
822  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
823  | 
case (Suc n m)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
824  | 
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
 | 
825  | 
show ?case  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
826  | 
proof (cases "n > 0")  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
827  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
828  | 
with u01 v01 show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
829  | 
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
 | 
830  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
831  | 
case True  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
832  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
833  | 
proof (cases "even k")  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
834  | 
case True  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
835  | 
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
 | 
836  | 
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
 | 
837  | 
proof -  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
838  | 
have "odd (Suc k)"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
839  | 
using True by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
840  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
841  | 
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
 | 
842  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
843  | 
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
 | 
844  | 
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
 | 
845  | 
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
 | 
846  | 
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
 | 
847  | 
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
 | 
848  | 
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
 | 
849  | 
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
 | 
850  | 
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
 | 
851  | 
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
 | 
852  | 
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
 | 
853  | 
\<le> 2/2 ^ Suc n"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
854  | 
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
 | 
855  | 
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
 | 
856  | 
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
 | 
857  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
858  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
859  | 
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
 | 
860  | 
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
 | 
861  | 
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
 | 
862  | 
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
 | 
863  | 
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
 | 
864  | 
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
 | 
865  | 
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
 | 
866  | 
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
 | 
867  | 
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
 | 
868  | 
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
 | 
869  | 
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
 | 
870  | 
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
 | 
871  | 
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
 | 
872  | 
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
 | 
873  | 
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
 | 
874  | 
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
 | 
875  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
876  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
877  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
878  | 
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
 | 
879  | 
using that by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
880  | 
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
 | 
881  | 
proof (induction n arbitrary: j)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
882  | 
case 0  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
883  | 
then show ?case by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
884  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
885  | 
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
 | 
886  | 
proof (cases "n > 0")  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
887  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
888  | 
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
 | 
889  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
890  | 
case True  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
891  | 
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
 | 
892  | 
case True  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
893  | 
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
 | 
894  | 
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
 | 
895  | 
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
 | 
896  | 
with k \<open>0 < n\<close> Suc.IH [of k] show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
897  | 
apply (simp add: m1_to_3 a41 b43 del: power_Suc)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
898  | 
apply (subst of_nat_diff, auto)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
899  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
900  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
901  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
902  | 
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
 | 
903  | 
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
 | 
904  | 
= 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
 | 
905  | 
"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
 | 
906  | 
= 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
 | 
907  | 
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
 | 
908  | 
using left_right_m [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
 | 
909  | 
apply (auto simp: add.commute feqm [OF order_refl] feqm [OF _ order_refl, symmetric])  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
910  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
911  | 
then  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
912  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
913  | 
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
 | 
914  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
915  | 
qed  | 
| 
 
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  | 
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
 | 
918  | 
\<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
 | 
919  | 
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
 | 
920  | 
proof (induction n arbitrary: j)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
921  | 
case 0  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
922  | 
then show ?case by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
923  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
924  | 
case (Suc n)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
925  | 
show ?case  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
926  | 
proof (cases "even j")  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
927  | 
case True  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
928  | 
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
 | 
929  | 
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
 | 
930  | 
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
 | 
931  | 
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
 | 
932  | 
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
 | 
933  | 
by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
934  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
935  | 
using Suc.IH [of k] k \<open>0 < k\<close>  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
936  | 
apply (simp add: less2n add.commute [of 1] m1_to_3 a41 b43 del: power_Suc)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
937  | 
apply (auto simp: of_nat_diff)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
938  | 
done  | 
| 
 
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"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
967  | 
by (force simp: divide_simps)  | 
| 
 
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"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1042  | 
apply (subst one_less_power)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1043  | 
using \<open>n > 0\<close> by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1044  | 
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
 | 
1045  | 
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
 | 
1046  | 
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
 | 
1047  | 
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
 | 
1048  | 
qed simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1049  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1050  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1051  | 
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
 | 
1052  | 
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
 | 
1053  | 
for i k m p  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1054  | 
proof -  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1055  | 
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
 | 
1056  | 
using j by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1057  | 
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
 | 
1058  | 
using k by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1059  | 
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
 | 
1060  | 
by (simp only: mult_ac)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1061  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1062  | 
by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1063  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1064  | 
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
 | 
1065  | 
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
 | 
1066  | 
for i k m p  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1067  | 
proof -  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1068  | 
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
 | 
1069  | 
using j by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1070  | 
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
 | 
1071  | 
by (rule k)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1072  | 
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
 | 
1073  | 
by (simp add: algebra_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1074  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1075  | 
by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1076  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1077  | 
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
 | 
1078  | 
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
 | 
1079  | 
proof -  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1080  | 
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
 | 
1081  | 
using j by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1082  | 
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
 | 
1083  | 
using i by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1084  | 
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
 | 
1085  | 
by (simp only: mult_ac)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1086  | 
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
 | 
1087  | 
by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1088  | 
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
 | 
1089  | 
by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1090  | 
finally show ?thesis by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1091  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1092  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1093  | 
proof  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1094  | 
have "real j < 2 ^ n"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1095  | 
using j_le i k  | 
| 
66912
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66453 
diff
changeset
 | 
1096  | 
apply (auto simp: le_max_iff_disj simp del: of_nat_less_numeral_power_cancel_iff  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
66453 
diff
changeset
 | 
1097  | 
elim!: le_less_trans)  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1098  | 
apply (auto simp: field_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1099  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1100  | 
then show "j < 2 ^ n"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1101  | 
by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1102  | 
show "\<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
 | 
1103  | 
using clo less_j1 j_le  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1104  | 
apply (auto simp: le_max_iff_disj divide_simps dist_norm)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1105  | 
apply (auto simp: algebra_simps abs_if split: if_split_asm dest: 1 2)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1106  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1107  | 
show "\<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
 | 
1108  | 
using clo less_j1 j_le  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1109  | 
apply (auto simp: le_max_iff_disj divide_simps dist_norm)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1110  | 
apply (auto simp: algebra_simps not_less abs_if split: if_split_asm dest: 3 2)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1111  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1112  | 
qed (use False in simp)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1113  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1114  | 
qed  | 
| 
 
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 (real i / 2 ^ m))) < e"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1116  | 
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
 | 
1117  | 
show "dist (f (c (real k / 2 ^ p))) (f(c(j / 2^n))) < e/2"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1118  | 
apply (rule dist_fc_close)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1119  | 
using \<open>0 < j\<close> \<open>j < 2 ^ n\<close> k clo_kj by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1120  | 
show "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
 | 
1121  | 
apply (rule dist_fc_close)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1122  | 
using \<open>0 < j\<close> \<open>j < 2 ^ n\<close> i clo_ij by auto  | 
| 
 
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  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1126  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1127  | 
  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
 | 
1128  | 
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
 | 
1129  | 
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
 | 
1130  | 
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
 | 
1131  | 
  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
 | 
1132  | 
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
 | 
1133  | 
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
 | 
1134  | 
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
 | 
1135  | 
  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
 | 
1136  | 
proof  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1137  | 
    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
 | 
1138  | 
proof (rule image_closure_subset)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1139  | 
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
 | 
1140  | 
using cont_h by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1141  | 
      show "closed (f ` {0..1})"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1142  | 
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
 | 
1143  | 
      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
 | 
1144  | 
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
 | 
1145  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1146  | 
    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
 | 
1147  | 
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
 | 
1148  | 
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
 | 
1149  | 
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
 | 
1150  | 
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
 | 
1151  | 
    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
 | 
1152  | 
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
 | 
1153  | 
fix x e::real  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1154  | 
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
 | 
1155  | 
      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
 | 
1156  | 
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
 | 
1157  | 
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
 | 
1158  | 
        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
 | 
1159  | 
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
 | 
1160  | 
      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
 | 
1161  | 
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
 | 
1162  | 
using that  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1163  | 
proof (induction n)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1164  | 
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
 | 
1165  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1166  | 
case (Suc n)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1167  | 
show ?case  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1168  | 
proof (cases "n=0")  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1169  | 
case True  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1170  | 
          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
 | 
1171  | 
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
 | 
1172  | 
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
 | 
1173  | 
proof cases  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1174  | 
case 1  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1175  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1176  | 
apply (rule_tac x=u in exI)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1177  | 
using uabv [of 1 1] f0u [of u] f0u [of x] by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1178  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1179  | 
case 2  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1180  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1181  | 
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
 | 
1182  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1183  | 
case 3  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1184  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1185  | 
apply (rule_tac x=v in exI)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1186  | 
using uabv [of 1 1] fv1 [of v] fv1 [of x] by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1187  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1188  | 
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
 | 
1189  | 
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
 | 
1190  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1191  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1192  | 
with Suc obtain m y  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1193  | 
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
 | 
1194  | 
              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
 | 
1195  | 
by metis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1196  | 
then obtain j where j: "m = 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
 | 
1197  | 
          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
 | 
1198  | 
            | "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
 | 
1199  | 
            | "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
 | 
1200  | 
using y j by force  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1201  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1202  | 
proof cases  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1203  | 
case 1  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1204  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1205  | 
apply (rule_tac x="4*j + 1" in exI)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1206  | 
apply (rule_tac x=y in exI)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1207  | 
using mless j \<open>n \<noteq> 0\<close>  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1208  | 
apply (simp add: feq a41 b41 add.commute [of 1] del: power_Suc)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1209  | 
apply (simp add: algebra_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1210  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1211  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1212  | 
case 2  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1213  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1214  | 
apply (rule_tac x="4*j + 1" in exI)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1215  | 
apply (rule_tac x="b((4*j + 1) / 2 ^ (Suc n))" in exI)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1216  | 
using mless \<open>n \<noteq> 0\<close> 2 j  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1217  | 
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]  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1218  | 
using left_right [of "c((2*j + 1) / 2^n)" "a((2*j + 1) / 2^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
 | 
1219  | 
apply (simp add: a41 b41 a43 b43 add.commute [of 1] add.commute [of 3] del: power_Suc)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1220  | 
apply (auto simp: feq [symmetric] intro: f_eqI)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1221  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1222  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1223  | 
case 3  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1224  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1225  | 
apply (rule_tac x="4*j + 3" in exI)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1226  | 
apply (rule_tac x=y in exI)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1227  | 
using mless j \<open>n \<noteq> 0\<close>  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1228  | 
apply (simp add: feq a43 b43 del: power_Suc)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1229  | 
apply (simp add: algebra_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1230  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1231  | 
qed  | 
| 
 
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  | 
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
 | 
1235  | 
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
 | 
1236  | 
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
 | 
1237  | 
by fastforce  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1238  | 
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
 | 
1239  | 
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
 | 
1240  | 
          and y: "y \<in> {a(m / 2^n) .. b(m / 2^n)}" and feq: "f x = f y"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1241  | 
by metis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1242  | 
then have "0 \<le> y" "y \<le> 1"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1243  | 
by (metis atLeastAtMost_iff a_ge_0 b_le_1 order.trans)+  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1244  | 
moreover have "y < \<delta> + c (real m / 2 ^ n)" "c (real m / 2 ^ n) < \<delta> + y"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1245  | 
using y apply simp_all  | 
| 
68527
 
2f4e2aab190a
Generalising and renaming some basic results
 
paulson <lp15@cam.ac.uk> 
parents: 
67968 
diff
changeset
 | 
1246  | 
using 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>  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1249  | 
      ultimately show "\<exists>k. \<exists>m\<in>{0<..<2 ^ k}. dist (h (real m / 2 ^ k)) (f x) < e"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1250  | 
apply (rule_tac x=n in exI)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1251  | 
apply (rule_tac x=m in bexI)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1252  | 
apply (auto simp: dist_norm h_eq feq \<delta>)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1253  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1254  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1255  | 
    also have "... \<subseteq> h ` {0..1}"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1256  | 
apply (rule closure_minimal)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1257  | 
using compact_continuous_image [OF cont_h] compact_imp_closed by (auto simp: D01_def)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1258  | 
    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
 | 
1259  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1260  | 
  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
 | 
1261  | 
proof -  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1262  | 
have "u < v"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1263  | 
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
 | 
1264  | 
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
 | 
1265  | 
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
 | 
1266  | 
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
 | 
1267  | 
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
 | 
1268  | 
have a_less_b:  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1269  | 
"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
 | 
1270  | 
(\<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
 | 
1271  | 
(\<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
 | 
1272  | 
proof (induction n arbitrary: j)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1273  | 
case 0 then show ?case  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1274  | 
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
 | 
1275  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1276  | 
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
 | 
1277  | 
proof (cases "n > 0")  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1278  | 
case False then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1279  | 
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
 | 
1280  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1281  | 
case True show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1282  | 
proof (cases "even j")  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1283  | 
case True  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1284  | 
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
 | 
1285  | 
by (auto elim!: evenE)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1286  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1287  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1288  | 
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
 | 
1289  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1290  | 
proof (cases "even k")  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1291  | 
case True  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1292  | 
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
 | 
1293  | 
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
 | 
1294  | 
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
 | 
1295  | 
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
 | 
1296  | 
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
 | 
1297  | 
by (auto intro: f_eqI)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1298  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1299  | 
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
 | 
1300  | 
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
 | 
1301  | 
proof -  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1302  | 
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
 | 
1303  | 
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
 | 
1304  | 
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
 | 
1305  | 
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
 | 
1306  | 
by (auto simp: algebra_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1307  | 
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
 | 
1308  | 
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
 | 
1309  | 
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
 | 
1310  | 
using cleb by blast  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1311  | 
ultimately show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1312  | 
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
 | 
1313  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1314  | 
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
 | 
1315  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1316  | 
fix x  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1317  | 
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
 | 
1318  | 
then show False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1319  | 
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
 | 
1320  | 
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
 | 
1321  | 
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
 | 
1322  | 
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
 | 
1323  | 
by (auto simp: algebra_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1324  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1325  | 
fix x  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1326  | 
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
 | 
1327  | 
then show False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1328  | 
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
 | 
1329  | 
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
 | 
1330  | 
by (auto simp: algebra_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1331  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1332  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1333  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1334  | 
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
 | 
1335  | 
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
 | 
1336  | 
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
 | 
1337  | 
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
 | 
1338  | 
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
 | 
1339  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1340  | 
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
 | 
1341  | 
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
 | 
1342  | 
proof -  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1343  | 
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
 | 
1344  | 
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
 | 
1345  | 
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
 | 
1346  | 
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
 | 
1347  | 
by (auto simp: algebra_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1348  | 
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
 | 
1349  | 
using alec by blast  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1350  | 
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
 | 
1351  | 
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
 | 
1352  | 
ultimately show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1353  | 
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
 | 
1354  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1355  | 
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
 | 
1356  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1357  | 
fix x  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1358  | 
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
 | 
1359  | 
then show False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1360  | 
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
 | 
1361  | 
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
 | 
1362  | 
by (auto simp: algebra_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1363  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1364  | 
fix x  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1365  | 
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
 | 
1366  | 
then show False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1367  | 
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
 | 
1368  | 
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
 | 
1369  | 
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
 | 
1370  | 
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
 | 
1371  | 
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
 | 
1372  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1373  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1374  | 
qed  | 
| 
 
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  | 
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
 | 
1378  | 
using a_less_b [of m n] apply (simp_all add: c_def midpoint_def)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1379  | 
using a_ge_0 [of m n] b_le_1 [of m n] apply linarith+  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1380  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1381  | 
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
 | 
1382  | 
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
 | 
1383  | 
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
 | 
1384  | 
\<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
 | 
1385  | 
\<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
 | 
1386  | 
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
 | 
1387  | 
using that  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1388  | 
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
 | 
1389  | 
case (less N)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1390  | 
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
 | 
1391  | 
by linarith  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1392  | 
then show ?case  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1393  | 
proof cases  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1394  | 
case 1  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1395  | 
with less.prems show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1396  | 
by (rule_tac x=1 in exI)+ (fastforce simp: divide_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1397  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1398  | 
case 2 show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1399  | 
proof (cases m)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1400  | 
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
 | 
1401  | 
by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1402  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1403  | 
case (Suc m') show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1404  | 
proof (cases p)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1405  | 
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
 | 
1406  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1407  | 
case (Suc p')  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1408  | 
have False if "real i * 2 ^ p' < real k * 2 ^ m'" "k < 2 ^ p'" "2 ^ m' \<le> i"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1409  | 
proof -  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1410  | 
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
 | 
1411  | 
using that by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1412  | 
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
 | 
1413  | 
using that by linarith  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1414  | 
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
 | 
1415  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1416  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1417  | 
using less.IH [of "m'+p'" i m' k p'] less.prems \<open>m = Suc m'\<close> 2 Suc  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1418  | 
apply atomize  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1419  | 
apply (force simp: divide_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1420  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1421  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1422  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1423  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1424  | 
case 3 show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1425  | 
proof (cases m)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1426  | 
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
 | 
1427  | 
by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1428  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1429  | 
case (Suc m') show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1430  | 
proof (cases p)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1431  | 
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
 | 
1432  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1433  | 
case (Suc p')  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1434  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1435  | 
using less.IH [of "m'+p'" "i - 2^m'" m' "k - 2 ^ p'" p'] less.prems \<open>m = Suc m'\<close> Suc 3  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1436  | 
apply atomize  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1437  | 
apply (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
 | 
1438  | 
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
 | 
1439  | 
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
 | 
1440  | 
apply (auto simp: field_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1441  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1442  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1443  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1444  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1445  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1446  | 
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
 | 
1447  | 
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
 | 
1448  | 
proof -  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1449  | 
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
 | 
1450  | 
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
 | 
1451  | 
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
 | 
1452  | 
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
 | 
1453  | 
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
 | 
1454  | 
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
 | 
1455  | 
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
 | 
1456  | 
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
 | 
1457  | 
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
 | 
1458  | 
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
 | 
1459  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1460  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1461  | 
with i_le_j q have less: "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
 | 
1462  | 
by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1463  | 
have *: "\<lbrakk>i < q; abs(i - q) < s*2; q = r + 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
 | 
1464  | 
by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1465  | 
have "c(i / 2^m) \<le> b(real(4 * q + 1) / 2 ^ (Suc n'))"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1466  | 
apply (rule ci_le_bj, force)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1467  | 
apply (rule * [OF less])  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1468  | 
using i_le_j clo_ij q apply (auto simp: divide_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1469  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1470  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1471  | 
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
 | 
1472  | 
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
 | 
1473  | 
by (auto simp: algebra_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1474  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1475  | 
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
 | 
1476  | 
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
 | 
1477  | 
case True  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1478  | 
then show ?thesis by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1479  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1480  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1481  | 
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
 | 
1482  | 
by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1483  | 
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
 | 
1484  | 
by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1485  | 
have "a(real(4*q + 3) / 2 ^ (Suc n')) \<le> c(j / 2^n)"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1486  | 
apply (rule aj_le_ci, force)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1487  | 
apply (rule * [OF less])  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1488  | 
using j_le_j clo_jj q apply (auto simp: divide_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1489  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1490  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1491  | 
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
 | 
1492  | 
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
 | 
1493  | 
by (auto simp: algebra_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1494  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1495  | 
finally show ?thesis .  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1496  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1497  | 
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
 | 
1498  | 
using that  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1499  | 
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
 | 
1500  | 
case (less x1 x2)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1501  | 
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
 | 
1502  | 
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
 | 
1503  | 
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
 | 
1504  | 
proof -  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1505  | 
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
 | 
1506  | 
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
 | 
1507  | 
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
 | 
1508  | 
unfolding closure_approachable  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1509  | 
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
 | 
1510  | 
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
 | 
1511  | 
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
 | 
1512  | 
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
 | 
1513  | 
proof -  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1514  | 
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
 | 
1515  | 
using less by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1516  | 
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
 | 
1517  | 
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
 | 
1518  | 
then have "N > 0"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1519  | 
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
 | 
1520  | 
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
 | 
1521  | 
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
 | 
1522  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1523  | 
proof  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1524  | 
show "0 < 2^N * p"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1525  | 
using p by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1526  | 
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
 | 
1527  | 
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
 | 
1528  | 
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
 | 
1529  | 
by (simp add: power_add)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1530  | 
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
 | 
1531  | 
by (simp add: yeq)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1532  | 
also have "... < (x2 - x1) / 64"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1533  | 
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
 | 
1534  | 
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
 | 
1535  | 
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
 | 
1536  | 
by (simp add: field_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1537  | 
also have "... < (x2 - x1) / 128"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1538  | 
using N by force  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1539  | 
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
 | 
1540  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1541  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1542  | 
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
 | 
1543  | 
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
 | 
1544  | 
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
 | 
1545  | 
proof  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1546  | 
show "0 < Suc (2*m)"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1547  | 
by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1548  | 
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
 | 
1549  | 
using m by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1550  | 
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
 | 
1551  | 
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
 | 
1552  | 
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
 | 
1553  | 
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
 | 
1554  | 
show "0 < 4*m + 3"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1555  | 
by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1556  | 
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
 | 
1557  | 
using m by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1558  | 
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
 | 
1559  | 
by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1560  | 
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
 | 
1561  | 
by (simp add: algebra_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1562  | 
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
 | 
1563  | 
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
 | 
1564  | 
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
 | 
1565  | 
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
 | 
1566  | 
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
 | 
1567  | 
by (simp add: c_def)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1568  | 
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
 | 
1569  | 
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
 | 
1570  | 
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
 | 
1571  | 
by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1572  | 
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
 | 
1573  | 
by (simp add: midpoint_def)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1574  | 
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
 | 
1575  | 
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
 | 
1576  | 
by (simp add: midpoint_def)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1577  | 
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
 | 
1578  | 
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
 | 
1579  | 
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
 | 
1580  | 
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
 | 
1581  | 
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
 | 
1582  | 
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
 | 
1583  | 
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
 | 
1584  | 
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
 | 
1585  | 
apply (rule right_neq)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1586  | 
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
 | 
1587  | 
apply (rule midR_le, auto)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1588  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1589  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1590  | 
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
 | 
1591  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1592  | 
have m_div: "0 < m / 2^n" "m / 2^n < 1"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1593  | 
using m by (auto simp: divide_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1594  | 
      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
 | 
1595  | 
by (subst closure_dyadic_rationals_in_convex_set_pos_1, simp_all add: not_le m)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1596  | 
      have closurem1: "{m / 2^n .. 1} = closure ({m / 2^n <..< 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
 | 
1597  | 
apply (subst closure_dyadic_rationals_in_convex_set_pos_1; simp add: not_le m)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1598  | 
using \<open>0 < real m / 2 ^ n\<close> by linarith  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1599  | 
      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
 | 
1600  | 
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
 | 
1601  | 
apply (rule continuous_on_subset [OF cont_h])  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1602  | 
apply (rule closure_minimal [OF subsetI])  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1603  | 
using that apply auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1604  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1605  | 
      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
 | 
1606  | 
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
 | 
1607  | 
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
 | 
1608  | 
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
 | 
1609  | 
by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1610  | 
      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
 | 
1611  | 
proof clarsimp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1612  | 
fix p q  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1613  | 
assume p: "0 < real p / 2 ^ q" "real p / 2 ^ q < real m / 2 ^ n"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1614  | 
then have [simp]: "0 < p" "p < 2 ^ q"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1615  | 
apply (simp add: divide_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1616  | 
apply (blast intro: p less_2I m_div less_trans)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1617  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1618  | 
        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
 | 
1619  | 
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
 | 
1620  | 
        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
 | 
1621  | 
by (simp add: h_eq)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1622  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1623  | 
      then have "h ` {0 .. m / 2^n} \<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
 | 
1624  | 
apply (subst closure0m)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1625  | 
apply (rule image_closure_subset [OF cont_h' closed_f'])  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1626  | 
using m_div apply auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1627  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1628  | 
      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
 | 
1629  | 
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
 | 
1630  | 
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
 | 
1631  | 
by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1632  | 
      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
 | 
1633  | 
proof clarsimp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1634  | 
fix p q  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1635  | 
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
 | 
1636  | 
then have [simp]: "0 < p"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1637  | 
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
 | 
1638  | 
        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
 | 
1639  | 
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
 | 
1640  | 
        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
 | 
1641  | 
by (simp add: h_eq)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1642  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1643  | 
      then have "h ` {m / 2^n .. 1} \<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
 | 
1644  | 
apply (subst closurem1)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1645  | 
apply (rule image_closure_subset [OF cont_h' closed_f'])  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1646  | 
using m apply auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1647  | 
done  | 
| 
 
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  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
1667  | 
theorem%important 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"  | 
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
1671  | 
proof%unimportant -  | 
| 
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 *)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1683  | 
show "\<phi> (INTER UNIV F)"  | 
| 
 
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)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1703  | 
        obtain w z where "w < z" and w: "w \<in> {x<..<y}" and z: "z \<in> {x<..<y}"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1704  | 
and wxe: "norm(w - x) < e" and zye: "norm(z - y) < e"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1705  | 
apply (rule_tac w = "x + (min e (y - x) / 3)" and z = "y - (min e (y - x) / 3)" in that)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1706  | 
using minxy \<open>0 < e\<close> less by simp_all  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1707  | 
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
 | 
1708  | 
by (metis \<open>\<And>n. closed (F n)\<close> image_iff)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1709  | 
        have eq: "{w..z} \<inter> INTER UNIV F = {}"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1710  | 
using less w z apply (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
 | 
1711  | 
by (metis (no_types, hide_lams) INT_I IntI empty_iff greaterThanLessThan_iff not_le order.trans)  | 
| 
 
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> {}"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1715  | 
          using \<open>w < z\<close> \<open>{w..z} \<inter> INTER K F = {}\<close> by auto
 | 
| 
 
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  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1737  | 
apply (auto simp: disjoint_iff_not_equal intro!: that)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1738  | 
by (metis greaterThanLessThan_iff less_eq_real_def)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1739  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1740  | 
        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
 | 
1741  | 
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
 | 
1742  | 
case True  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1743  | 
          have "z \<in> {w..z}"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1744  | 
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
 | 
1745  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1746  | 
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
 | 
1747  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1748  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1749  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1750  | 
          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
 | 
1751  | 
            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
 | 
1752  | 
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
 | 
1753  | 
            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
 | 
1754  | 
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
 | 
1755  | 
            show "\<And>b. \<lbrakk>b \<in> F n \<inter> {z..y}; open_segment z b \<inter> (F n \<inter> {z..y}) = {}\<rbrakk> \<Longrightarrow> thesis"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1756  | 
apply (rule that)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1757  | 
apply (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
 | 
1758  | 
by (metis DiffI Int_iff atLeastAtMost_diff_ends atLeastAtMost_iff atLeastatMost_empty_iff empty_iff insert_iff False)  | 
| 
 
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}"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1774  | 
by (metis DiffI disjoint_iff_not_equal less_irrefl singletonD that v(3))  | 
| 
 
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  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1780  | 
by (metis DiffI disjoint_iff_not_equal less_eq_real_def not_le singletonD that u(3))  | 
| 
 
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  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1792  | 
apply (simp add: open_segment_eq_real_ivl dist_norm del: divide_const_simps)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1793  | 
by (metis dist_norm dist_triangle_half_r less_irrefl)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1794  | 
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
 | 
1795  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1796  | 
unfolding \<phi>_def by (metis (no_types, hide_lams) INT_I Inf_lower2 rangeI that F01 subsetCE pqF)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1797  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1798  | 
    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
 | 
1799  | 
qed (meson \<phi>_def)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1800  | 
  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
 | 
1801  | 
    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
 | 
1802  | 
unfolding \<phi>_def by metis+  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1803  | 
  then have "T \<noteq> {}" by auto
 | 
| 67613 | 1804  | 
  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
 | 
1805  | 
  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
 | 
1806  | 
for x y z  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1807  | 
proof (cases "x \<in> T")  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1808  | 
case True  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1809  | 
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
 | 
1810  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1811  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1812  | 
    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
 | 
1813  | 
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
 | 
1814  | 
moreover have "open_segment y z \<inter> T \<subseteq> 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
 | 
1815  | 
apply auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1816  | 
by (metis greaterThanLessThan_iff less_eq_real_def less_le_trans linorder_neqE_linordered_idom open_segment_eq_real_ivl)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1817  | 
    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
 | 
1818  | 
by blast  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1819  | 
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
 | 
1820  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1821  | 
  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
 | 
1822  | 
using that unfolding h_def  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1823  | 
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
 | 
1824  | 
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
 | 
1825  | 
by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1826  | 
  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
 | 
1827  | 
    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
 | 
1828  | 
  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
 | 
1829  | 
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
 | 
1830  | 
case True  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1831  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1832  | 
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
 | 
1833  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1834  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1835  | 
    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
 | 
1836  | 
      "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
 | 
1837  | 
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
 | 
1838  | 
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
 | 
1839  | 
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
 | 
1840  | 
ultimately show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1841  | 
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
 | 
1842  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1843  | 
  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
 | 
1844  | 
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
 | 
1845  | 
    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
 | 
1846  | 
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
 | 
1847  | 
fix u \<epsilon>::real  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1848  | 
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
 | 
1849  | 
      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
 | 
1850  | 
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
 | 
1851  | 
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
 | 
1852  | 
      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
 | 
1853  | 
      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
 | 
1854  | 
case True  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1855  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1856  | 
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
 | 
1857  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1858  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1859  | 
        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
 | 
1860  | 
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
 | 
1861  | 
        obtain w where "w \<in> T" and w: "w \<in> closed_segment u v" "open_segment u w \<inter> T = {}"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1862  | 
apply (rule segment_to_point_exists [OF uvT, of u])  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1863  | 
by (metis IntD1 Int_commute Int_left_commute ends_in_segment(1) inf.orderE subset_oc_segment)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1864  | 
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
 | 
1865  | 
          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
 | 
1866  | 
        obtain z where "z \<in> T" and z: "z \<in> closed_segment u v" "open_segment v z \<inter> T = {}"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1867  | 
apply (rule segment_to_point_exists [OF uvT, of v])  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1868  | 
by (metis IntD2 Int_commute Int_left_commute ends_in_segment(2) inf.orderE subset_oc_segment)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1869  | 
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
 | 
1870  | 
          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
 | 
1871  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1872  | 
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
 | 
1873  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1874  | 
      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
 | 
1875  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1876  | 
    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
 | 
1877  | 
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
 | 
1878  | 
fix u v  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1879  | 
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
 | 
1880  | 
      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
 | 
1881  | 
proof (intro exI conjI)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1882  | 
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
 | 
1883  | 
by simp  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1884  | 
        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
 | 
1885  | 
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
 | 
1886  | 
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
 | 
1887  | 
          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
 | 
1888  | 
          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
 | 
1889  | 
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
 | 
1890  | 
for x y  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1891  | 
        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
 | 
1892  | 
case True  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1893  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1894  | 
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
 | 
1895  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1896  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1897  | 
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
 | 
1898  | 
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
 | 
1899  | 
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
 | 
1900  | 
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
 | 
1901  | 
proof  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1902  | 
assume "?xuyv"  | 
| 
 
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 \<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
 | 
1904  | 
using disjT by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1905  | 
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
 | 
1906  | 
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
 | 
1907  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1908  | 
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
 | 
1909  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1910  | 
assume "?yuxv"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1911  | 
            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
 | 
1912  | 
using disjT by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1913  | 
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
 | 
1914  | 
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
 | 
1915  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1916  | 
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
 | 
1917  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1918  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1919  | 
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
 | 
1920  | 
proof (rule T)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1921  | 
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
 | 
1922  | 
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
 | 
1923  | 
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
 | 
1924  | 
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
 | 
1925  | 
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
 | 
1926  | 
            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
 | 
1927  | 
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
 | 
1928  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1929  | 
        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
 | 
1930  | 
by blast  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1931  | 
        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
 | 
1932  | 
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
 | 
1933  | 
qed auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1934  | 
      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
 | 
1935  | 
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
 | 
1936  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1937  | 
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
 | 
1938  | 
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
 | 
1939  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1940  | 
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
 | 
1941  | 
    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
 | 
1942  | 
      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
 | 
1943  | 
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
 | 
1944  | 
then have "arc g"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1945  | 
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
 | 
1946  | 
  obtain u v where "u \<in> {0..1}" "a = g u" "v \<in> {0..1}" "b = g v"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1947  | 
by (metis (mono_tags, hide_lams) \<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)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1948  | 
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
 | 
1949  | 
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
 | 
1950  | 
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
 | 
1951  | 
    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
 | 
1952  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1953  | 
proof  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1954  | 
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
 | 
1955  | 
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
 | 
1956  | 
show "path_image (subpath u v g) \<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 \<open>arc g\<close> \<open>u \<in> {0..1}\<close> \<open>v \<in> {0..1}\<close> arc_imp_path order_trans pag path_image_def path_image_subpath_subset ph)
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1958  | 
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
 | 
1959  | 
      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
 | 
1960  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1961  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1962  | 
|
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1963  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
1964  | 
corollary%important path_connected_arcwise:  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1965  | 
  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
 | 
1966  | 
shows "path_connected S \<longleftrightarrow>  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1967  | 
(\<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
 | 
1968  | 
(is "?lhs = ?rhs")  | 
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
1969  | 
proof%unimportant (intro iffI impI ballI)  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1970  | 
fix x y  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1971  | 
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
 | 
1972  | 
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
 | 
1973  | 
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
 | 
1974  | 
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
 | 
1975  | 
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
 | 
1976  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1977  | 
assume R [rule_format]: ?rhs  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1978  | 
show ?lhs  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1979  | 
unfolding path_connected_def  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1980  | 
proof (intro ballI)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1981  | 
fix x y  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1982  | 
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
 | 
1983  | 
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
 | 
1984  | 
proof (cases "x = y")  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1985  | 
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
 | 
1986  | 
by blast  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1987  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1988  | 
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
 | 
1989  | 
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
 | 
1990  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1991  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1992  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1993  | 
|
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1994  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
1995  | 
corollary%important arc_connected_trans:  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1996  | 
  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
 | 
1997  | 
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
 | 
1998  | 
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
 | 
1999  | 
"pathstart i = pathstart g" "pathfinish i = pathfinish h"  | 
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
2000  | 
by%unimportant (metis (no_types, hide_lams) 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
 | 
2001  | 
|
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2002  | 
|
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2003  | 
|
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2004  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
2005  | 
subsection%important\<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
 | 
2006  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
2007  | 
lemma%important dense_accessible_frontier_points:  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2008  | 
  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
 | 
2009  | 
  assumes "open S" and opeSV: "openin (subtopology euclidean (frontier S)) V" and "V \<noteq> {}"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2010  | 
  obtains g where "arc g" "g ` {0..<1} \<subseteq> S" "pathstart g \<in> S" "pathfinish g \<in> V"
 | 
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
2011  | 
proof%unimportant -  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2012  | 
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
 | 
2013  | 
    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
 | 
2014  | 
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
 | 
2015  | 
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
 | 
2016  | 
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
 | 
2017  | 
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
 | 
2018  | 
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
 | 
2019  | 
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
 | 
2020  | 
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
 | 
2021  | 
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
 | 
2022  | 
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
 | 
2023  | 
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
 | 
2024  | 
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
 | 
2025  | 
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
 | 
2026  | 
have "y \<noteq> z"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2027  | 
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
 | 
2028  | 
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
 | 
2029  | 
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
 | 
2030  | 
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
 | 
2031  | 
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
 | 
2032  | 
using \<open>y \<noteq> z\<close> by (force simp: path_connected_arcwise)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2033  | 
  have "compact (g -` frontier S \<inter> {0..1})"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2034  | 
apply (simp add: compact_eq_bounded_closed bounded_Int bounded_closed_interval)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2035  | 
apply (rule closed_vimage_Int)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2036  | 
using \<open>arc g\<close> apply (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
 | 
2037  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2038  | 
  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
 | 
2039  | 
proof -  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2040  | 
    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
 | 
2041  | 
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
 | 
2042  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2043  | 
by blast  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2044  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2045  | 
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
 | 
2046  | 
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
 | 
2047  | 
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
 | 
2048  | 
moreover have "t \<noteq> 0"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2049  | 
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
 | 
2050  | 
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
 | 
2051  | 
by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2052  | 
have "V \<subseteq> frontier S"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2053  | 
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
 | 
2054  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2055  | 
proof  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2056  | 
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
 | 
2057  | 
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
 | 
2058  | 
have "g 0 \<in> S"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2059  | 
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
 | 
2060  | 
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
 | 
2061  | 
by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2062  | 
have "g t \<in> V"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2063  | 
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
 | 
2064  | 
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
 | 
2065  | 
by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2066  | 
    then have "inj_on (subpath 0 t g) {0..1}"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2067  | 
using t01  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2068  | 
apply (clarsimp simp: inj_on_def subpath_def)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2069  | 
apply (drule inj_onD [OF arc_imp_inj_on [OF \<open>arc g\<close>]])  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2070  | 
using mult_le_one apply auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2071  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2072  | 
    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
 | 
2073  | 
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
 | 
2074  | 
    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
 | 
2075  | 
proof -  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2076  | 
      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
 | 
2077  | 
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
 | 
2078  | 
      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
 | 
2079  | 
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
 | 
2080  | 
        show "connected (subpath 0 t g ` {0..<1})"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2081  | 
apply (rule connected_continuous_image)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2082  | 
apply (simp add: subpath_def)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2083  | 
apply (intro continuous_intros continuous_on_compose2 [OF contg])  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2084  | 
apply (auto simp: \<open>0 \<le> t\<close> \<open>t \<le> 1\<close> mult_le_one)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2085  | 
done  | 
| 
 
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  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
2102  | 
lemma%important 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> {}"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2105  | 
and ope: "openin (subtopology euclidean (frontier S)) V"  | 
| 
 
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"
 | 
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
2107  | 
proof%unimportant -  | 
| 
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"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2124  | 
apply (rule path_contains_arc [of "f +++ g" x "pathfinish g"])  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2125  | 
using f \<open>x \<notin> V\<close> \<open>pathfinish g \<in> V\<close> by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2126  | 
  have "h ` {0..1} - {h 1} \<subseteq> S"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2127  | 
using f g h apply (clarsimp simp: path_image_join)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2128  | 
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
 | 
2129  | 
by (metis le_less)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2130  | 
  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
 | 
2131  | 
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
 | 
2132  | 
then show thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2133  | 
apply (rule that [OF \<open>arc h\<close>])  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2134  | 
using h \<open>pathfinish g \<in> V\<close> by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2135  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2136  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
2137  | 
lemma%important dense_access_fp_aux:  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2138  | 
  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
 | 
2139  | 
assumes S: "open S" "connected S"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2140  | 
and opeSU: "openin (subtopology euclidean (frontier S)) U"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2141  | 
and opeSV: "openin (subtopology euclidean (frontier S)) V"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2142  | 
      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
 | 
2143  | 
  obtains g where "arc g" "pathstart g \<in> U" "pathfinish g \<in> V" "g ` {0<..<1} \<subseteq> S"
 | 
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
2144  | 
proof%unimportant -  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2145  | 
  have "S \<noteq> {}"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2146  | 
    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
 | 
2147  | 
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
 | 
2148  | 
  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
 | 
2149  | 
    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
 | 
2150  | 
  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
 | 
2151  | 
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
 | 
2152  | 
    show "U - {pathfinish g} \<noteq> {}"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2153  | 
using \<open>pathfinish g \<in> V\<close> \<open>\<not> U \<subseteq> V\<close> by blast  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2154  | 
    show "openin (subtopology euclidean (frontier S)) (U - {pathfinish g})"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2155  | 
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
 | 
2156  | 
qed auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2157  | 
obtain \<gamma> where "arc \<gamma>"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2158  | 
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
 | 
2159  | 
"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
 | 
2160  | 
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
 | 
2161  | 
show "path (reversepath h +++ g)"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2162  | 
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
 | 
2163  | 
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
 | 
2164  | 
"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
 | 
2165  | 
by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2166  | 
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
 | 
2167  | 
      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
 | 
2168  | 
qed auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2169  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2170  | 
proof  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2171  | 
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
 | 
2172  | 
      using \<gamma> \<open>arc \<gamma>\<close> \<open>pathfinish h \<in> U - {pathfinish g}\<close>  \<open>pathfinish g \<in> V\<close> by auto
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2173  | 
    have "\<gamma> ` {0..1} - {\<gamma> 0, \<gamma> 1} \<subseteq> S"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2174  | 
using \<gamma> g h  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2175  | 
apply (simp add: path_image_join)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2176  | 
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
 | 
2177  | 
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
 | 
2178  | 
    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
 | 
2179  | 
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
 | 
2180  | 
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
 | 
2181  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2182  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2183  | 
|
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
2184  | 
lemma%important 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
 | 
2185  | 
  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
 | 
2186  | 
assumes S: "open S" "connected S"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2187  | 
and opeSU: "openin (subtopology euclidean (frontier S)) U"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2188  | 
and opeSV: "openin (subtopology euclidean (frontier S)) V"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2189  | 
      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
 | 
2190  | 
    obtains g where "arc g" "pathstart g \<in> U" "pathfinish g \<in> V" "g ` {0<..<1} \<subseteq> S"
 | 
| 
68833
 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
68527 
diff
changeset
 | 
2191  | 
proof%unimportant -  | 
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2192  | 
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
 | 
2193  | 
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
 | 
2194  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2195  | 
proof cases  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2196  | 
case 1 then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2197  | 
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
 | 
2198  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2199  | 
case 2  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2200  | 
    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
 | 
2201  | 
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
 | 
2202  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2203  | 
proof  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2204  | 
show "arc (reversepath g)"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2205  | 
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
 | 
2206  | 
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
 | 
2207  | 
using g by auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2208  | 
      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
 | 
2209  | 
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
 | 
2210  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2211  | 
qed  | 
| 
 
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  | 
|
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2214  | 
end  |