| author | haftmann | 
| Sun, 09 Feb 2020 21:58:42 +0000 | |
| changeset 71425 | f2da99316b86 | 
| parent 71255 | 4258ee13f5d4 | 
| child 73932 | fd21b4a93043 | 
| permissions | -rw-r--r-- | 
| 63938 | 1 | (* Author: L C Paulson, University of Cambridge | 
| 33175 | 2 | Author: Amine Chaieb, University of Cambridge | 
| 3 | Author: Robert Himmelmann, TU Muenchen | |
| 44075 | 4 | Author: Brian Huffman, Portland State University | 
| 33175 | 5 | *) | 
| 6 | ||
| 69676 | 7 | chapter \<open>Vector Analysis\<close> | 
| 33175 | 8 | |
| 9 | theory Topology_Euclidean_Space | |
| 69516 
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
 immler parents: 
69508diff
changeset | 10 | imports | 
| 69544 
5aa5a8d6e5b5
split off theorems involving classes below metric_space and real_normed_vector
 immler parents: 
69516diff
changeset | 11 | Elementary_Normed_Spaces | 
| 69516 
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
 immler parents: 
69508diff
changeset | 12 | Linear_Algebra | 
| 
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
 immler parents: 
69508diff
changeset | 13 | Norm_Arith | 
| 51343 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 hoelzl parents: 
51342diff
changeset | 14 | begin | 
| 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 hoelzl parents: 
51342diff
changeset | 15 | |
| 69676 | 16 | section \<open>Elementary Topology in Euclidean Space\<close> | 
| 17 | ||
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50324diff
changeset | 18 | lemma euclidean_dist_l2: | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50324diff
changeset | 19 | fixes x y :: "'a :: euclidean_space" | 
| 67155 | 20 | shows "dist x y = L2_set (\<lambda>i. dist (x \<bullet> i) (y \<bullet> i)) Basis" | 
| 21 | unfolding dist_norm norm_eq_sqrt_inner L2_set_def | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50324diff
changeset | 22 | by (subst euclidean_inner) (simp add: power2_eq_square inner_diff_left) | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50324diff
changeset | 23 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 24 | lemma norm_nth_le: "norm (x \<bullet> i) \<le> norm x" if "i \<in> Basis" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 25 | proof - | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 26 |   have "(x \<bullet> i)\<^sup>2 = (\<Sum>i\<in>{i}. (x \<bullet> i)\<^sup>2)"
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 27 | by simp | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 28 | also have "\<dots> \<le> (\<Sum>i\<in>Basis. (x \<bullet> i)\<^sup>2)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 29 | by (intro sum_mono2) (auto simp: that) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 30 | finally show ?thesis | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 31 | unfolding norm_conv_dist euclidean_dist_l2[of x] L2_set_def | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 32 | by (auto intro!: real_le_rsqrt) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 33 | qed | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 34 | |
| 71033 | 35 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Continuity of the representation WRT an orthogonal basis\<close> | 
| 70065 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 36 | |
| 70086 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70065diff
changeset | 37 | lemma orthogonal_Basis: "pairwise orthogonal Basis" | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70065diff
changeset | 38 | by (simp add: inner_not_same_Basis orthogonal_def pairwise_def) | 
| 
72c52a897de2
First tranche of the Homology development: Simplices
 paulson <lp15@cam.ac.uk> parents: 
70065diff
changeset | 39 | |
| 70065 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 40 | lemma representation_bound: | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 41 | fixes B :: "'N::real_inner set" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 42 | assumes "finite B" "independent B" "b \<in> B" and orth: "pairwise orthogonal B" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 43 | obtains m where "m > 0" "\<And>x. x \<in> span B \<Longrightarrow> \<bar>representation B x b\<bar> \<le> m * norm x" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 44 | proof | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 45 | fix x | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 46 | assume x: "x \<in> span B" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 47 | have "b \<noteq> 0" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 48 | using \<open>independent B\<close> \<open>b \<in> B\<close> dependent_zero by blast | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 49 | have [simp]: "b \<bullet> b' = (if b' = b then (norm b)\<^sup>2 else 0)" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 50 | if "b \<in> B" "b' \<in> B" for b b' | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 51 | using orth by (simp add: orthogonal_def pairwise_def norm_eq_sqrt_inner that) | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 52 | have "norm x = norm (\<Sum>b\<in>B. representation B x b *\<^sub>R b)" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 53 | using real_vector.sum_representation_eq [OF \<open>independent B\<close> x \<open>finite B\<close> order_refl] | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 54 | by simp | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 55 | also have "\<dots> = sqrt ((\<Sum>b\<in>B. representation B x b *\<^sub>R b) \<bullet> (\<Sum>b\<in>B. representation B x b *\<^sub>R b))" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 56 | by (simp add: norm_eq_sqrt_inner) | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 57 | also have "\<dots> = sqrt (\<Sum>b\<in>B. (representation B x b *\<^sub>R b) \<bullet> (representation B x b *\<^sub>R b))" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 58 | using \<open>finite B\<close> | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 59 | by (simp add: inner_sum_left inner_sum_right if_distrib [of "\<lambda>x. _ * x"] cong: if_cong sum.cong_simp) | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 60 | also have "\<dots> = sqrt (\<Sum>b\<in>B. (norm (representation B x b *\<^sub>R b))\<^sup>2)" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 61 | by (simp add: mult.commute mult.left_commute power2_eq_square) | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 62 | also have "\<dots> = sqrt (\<Sum>b\<in>B. (representation B x b)\<^sup>2 * (norm b)\<^sup>2)" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 63 | by (simp add: norm_mult power_mult_distrib) | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 64 | finally have "norm x = sqrt (\<Sum>b\<in>B. (representation B x b)\<^sup>2 * (norm b)\<^sup>2)" . | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 65 | moreover | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 66 | have "sqrt ((representation B x b)\<^sup>2 * (norm b)\<^sup>2) \<le> sqrt (\<Sum>b\<in>B. (representation B x b)\<^sup>2 * (norm b)\<^sup>2)" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 67 | using \<open>b \<in> B\<close> \<open>finite B\<close> by (auto intro: member_le_sum) | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 68 | then have "\<bar>representation B x b\<bar> \<le> (1 / norm b) * sqrt (\<Sum>b\<in>B. (representation B x b)\<^sup>2 * (norm b)\<^sup>2)" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 69 | using \<open>b \<noteq> 0\<close> by (simp add: field_split_simps real_sqrt_mult del: real_sqrt_le_iff) | 
| 70065 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 70 | ultimately show "\<bar>representation B x b\<bar> \<le> (1 / norm b) * norm x" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 71 | by simp | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 72 | next | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 73 | show "0 < 1 / norm b" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 74 | using \<open>independent B\<close> \<open>b \<in> B\<close> dependent_zero by auto | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 75 | qed | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 76 | |
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 77 | lemma continuous_on_representation: | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 78 | fixes B :: "'N::euclidean_space set" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 79 | assumes "finite B" "independent B" "b \<in> B" "pairwise orthogonal B" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 80 | shows "continuous_on (span B) (\<lambda>x. representation B x b)" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 81 | proof | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 82 | show "\<exists>d>0. \<forall>x'\<in>span B. dist x' x < d \<longrightarrow> dist (representation B x' b) (representation B x b) \<le> e" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 83 | if "e > 0" "x \<in> span B" for x e | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 84 | proof - | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 85 | obtain m where "m > 0" and m: "\<And>x. x \<in> span B \<Longrightarrow> \<bar>representation B x b\<bar> \<le> m * norm x" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 86 | using assms representation_bound by blast | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 87 | show ?thesis | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 88 | unfolding dist_norm | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 89 | proof (intro exI conjI ballI impI) | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 90 | show "e/m > 0" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 91 | by (simp add: \<open>e > 0\<close> \<open>m > 0\<close>) | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 92 | show "norm (representation B x' b - representation B x b) \<le> e" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 93 | if x': "x' \<in> span B" and less: "norm (x'-x) < e/m" for x' | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 94 | proof - | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 95 | have "\<bar>representation B (x'-x) b\<bar> \<le> m * norm (x'-x)" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 96 | using m [of "x'-x"] \<open>x \<in> span B\<close> span_diff x' by blast | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 97 | also have "\<dots> < e" | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 98 | by (metis \<open>m > 0\<close> less mult.commute pos_less_divide_eq) | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 99 | finally have "\<bar>representation B (x'-x) b\<bar> \<le> e" by simp | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 100 | then show ?thesis | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 101 | by (simp add: \<open>x \<in> span B\<close> \<open>independent B\<close> representation_diff x') | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 102 | qed | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 103 | qed | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 104 | qed | 
| 
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 105 | qed | 
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 106 | |
| 70136 | 107 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Balls in Euclidean Space\<close> | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 108 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 109 | lemma cball_subset_cball_iff: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 110 | fixes a :: "'a :: euclidean_space" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 111 | shows "cball a r \<subseteq> cball a' r' \<longleftrightarrow> dist a a' + r \<le> r' \<or> r < 0" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 112 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 113 | proof | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 114 | assume ?lhs | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 115 | then show ?rhs | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 116 | proof (cases "r < 0") | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 117 | case True | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 118 | then show ?rhs by simp | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 119 | next | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 120 | case False | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 121 | then have [simp]: "r \<ge> 0" by simp | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 122 | have "norm (a - a') + r \<le> r'" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 123 | proof (cases "a = a'") | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 124 | case True | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 125 | then show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 126 | using subsetD [where c = "a + r *\<^sub>R (SOME i. i \<in> Basis)", OF \<open>?lhs\<close>] subsetD [where c = a, OF \<open>?lhs\<close>] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 127 | by (force simp: SOME_Basis dist_norm) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 128 | next | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 129 | case False | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 130 | have "norm (a' - (a + (r / norm (a - a')) *\<^sub>R (a - a'))) = norm (a' - a - (r / norm (a - a')) *\<^sub>R (a - a'))" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 131 | by (simp add: algebra_simps) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 132 | also have "... = norm ((-1 - (r / norm (a - a'))) *\<^sub>R (a - a'))" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 133 | by (simp add: algebra_simps) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 134 | also from \<open>a \<noteq> a'\<close> have "... = \<bar>- norm (a - a') - r\<bar>" | 
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70136diff
changeset | 135 | by simp (simp add: field_simps) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 136 | finally have [simp]: "norm (a' - (a + (r / norm (a - a')) *\<^sub>R (a - a'))) = \<bar>norm (a - a') + r\<bar>" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 137 | by linarith | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 138 | from \<open>a \<noteq> a'\<close> show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 139 | using subsetD [where c = "a' + (1 + r / norm(a - a')) *\<^sub>R (a - a')", OF \<open>?lhs\<close>] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 140 | by (simp add: dist_norm scaleR_add_left) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 141 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 142 | then show ?rhs | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 143 | by (simp add: dist_norm) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 144 | qed | 
| 70952 
f140135ff375
example applications of the 'metric' decision procedure, by Maximilian Schäffeler
 immler parents: 
70817diff
changeset | 145 | qed metric | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 146 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 147 | lemma cball_subset_ball_iff: "cball a r \<subseteq> ball a' r' \<longleftrightarrow> dist a a' + r < r' \<or> r < 0" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 148 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 149 | for a :: "'a::euclidean_space" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 150 | proof | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 151 | assume ?lhs | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 152 | then show ?rhs | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 153 | proof (cases "r < 0") | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 154 | case True then | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 155 | show ?rhs by simp | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 156 | next | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 157 | case False | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 158 | then have [simp]: "r \<ge> 0" by simp | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 159 | have "norm (a - a') + r < r'" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 160 | proof (cases "a = a'") | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 161 | case True | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 162 | then show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 163 | using subsetD [where c = "a + r *\<^sub>R (SOME i. i \<in> Basis)", OF \<open>?lhs\<close>] subsetD [where c = a, OF \<open>?lhs\<close>] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 164 | by (force simp: SOME_Basis dist_norm) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 165 | next | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 166 | case False | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 167 | have False if "norm (a - a') + r \<ge> r'" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 168 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 169 | from that have "\<bar>r' - norm (a - a')\<bar> \<le> r" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 170 | by (simp split: abs_split) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 171 | (metis \<open>0 \<le> r\<close> \<open>?lhs\<close> centre_in_cball dist_commute dist_norm less_asym mem_ball subset_eq) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 172 | then show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 173 | using subsetD [where c = "a + (r' / norm(a - a') - 1) *\<^sub>R (a - a')", OF \<open>?lhs\<close>] \<open>a \<noteq> a'\<close> | 
| 70802 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70136diff
changeset | 174 | apply (simp add: dist_norm) | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70136diff
changeset | 175 | apply (simp add: scaleR_left_diff_distrib) | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70136diff
changeset | 176 | apply (simp add: field_simps) | 
| 
160eaf566bcb
formally augmented corresponding rules for field_simps
 haftmann parents: 
70136diff
changeset | 177 | done | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 178 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 179 | then show ?thesis by force | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 180 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 181 | then show ?rhs by (simp add: dist_norm) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 182 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 183 | next | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 184 | assume ?rhs | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 185 | then show ?lhs | 
| 70952 
f140135ff375
example applications of the 'metric' decision procedure, by Maximilian Schäffeler
 immler parents: 
70817diff
changeset | 186 | by metric | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 187 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 188 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 189 | lemma ball_subset_cball_iff: "ball a r \<subseteq> cball a' r' \<longleftrightarrow> dist a a' + r \<le> r' \<or> r \<le> 0" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 190 | (is "?lhs = ?rhs") | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 191 | for a :: "'a::euclidean_space" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 192 | proof (cases "r \<le> 0") | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 193 | case True | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 194 | then show ?thesis | 
| 70952 
f140135ff375
example applications of the 'metric' decision procedure, by Maximilian Schäffeler
 immler parents: 
70817diff
changeset | 195 | by metric | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 196 | next | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 197 | case False | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 198 | show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 199 | proof | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 200 | assume ?lhs | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 201 | then have "(cball a r \<subseteq> cball a' r')" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 202 | by (metis False closed_cball closure_ball closure_closed closure_mono not_less) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 203 | with False show ?rhs | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 204 | by (fastforce iff: cball_subset_cball_iff) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 205 | next | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 206 | assume ?rhs | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 207 | with False show ?lhs | 
| 70952 
f140135ff375
example applications of the 'metric' decision procedure, by Maximilian Schäffeler
 immler parents: 
70817diff
changeset | 208 | by metric | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 209 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 210 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 211 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 212 | lemma ball_subset_ball_iff: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 213 | fixes a :: "'a :: euclidean_space" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 214 | shows "ball a r \<subseteq> ball a' r' \<longleftrightarrow> dist a a' + r \<le> r' \<or> r \<le> 0" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 215 | (is "?lhs = ?rhs") | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 216 | proof (cases "r \<le> 0") | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 217 | case True then show ?thesis | 
| 70952 
f140135ff375
example applications of the 'metric' decision procedure, by Maximilian Schäffeler
 immler parents: 
70817diff
changeset | 218 | by metric | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 219 | next | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 220 | case False show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 221 | proof | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 222 | assume ?lhs | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 223 | then have "0 < r'" | 
| 70952 
f140135ff375
example applications of the 'metric' decision procedure, by Maximilian Schäffeler
 immler parents: 
70817diff
changeset | 224 | using False by metric | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 225 | then have "(cball a r \<subseteq> cball a' r')" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 226 | by (metis False\<open>?lhs\<close> closure_ball closure_mono not_less) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 227 | then show ?rhs | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 228 | using False cball_subset_cball_iff by fastforce | 
| 70952 
f140135ff375
example applications of the 'metric' decision procedure, by Maximilian Schäffeler
 immler parents: 
70817diff
changeset | 229 | qed metric | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 230 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 231 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 232 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 233 | lemma ball_eq_ball_iff: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 234 | fixes x :: "'a :: euclidean_space" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 235 | shows "ball x d = ball y e \<longleftrightarrow> d \<le> 0 \<and> e \<le> 0 \<or> x=y \<and> d=e" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 236 | (is "?lhs = ?rhs") | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 237 | proof | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 238 | assume ?lhs | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 239 | then show ?rhs | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 240 | proof (cases "d \<le> 0 \<or> e \<le> 0") | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 241 | case True | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 242 | with \<open>?lhs\<close> show ?rhs | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 243 | by safe (simp_all only: ball_eq_empty [of y e, symmetric] ball_eq_empty [of x d, symmetric]) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 244 | next | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 245 | case False | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 246 | with \<open>?lhs\<close> show ?rhs | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 247 | apply (auto simp: set_eq_subset ball_subset_ball_iff dist_norm norm_minus_commute algebra_simps) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 248 | apply (metis add_le_same_cancel1 le_add_same_cancel1 norm_ge_zero norm_pths(2) order_trans) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 249 | apply (metis add_increasing2 add_le_imp_le_right eq_iff norm_ge_zero) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 250 | done | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 251 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 252 | next | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 253 | assume ?rhs then show ?lhs | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 254 | by (auto simp: set_eq_subset ball_subset_ball_iff) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 255 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 256 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 257 | lemma cball_eq_cball_iff: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 258 | fixes x :: "'a :: euclidean_space" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 259 | shows "cball x d = cball y e \<longleftrightarrow> d < 0 \<and> e < 0 \<or> x=y \<and> d=e" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 260 | (is "?lhs = ?rhs") | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 261 | proof | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 262 | assume ?lhs | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 263 | then show ?rhs | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 264 | proof (cases "d < 0 \<or> e < 0") | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 265 | case True | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 266 | with \<open>?lhs\<close> show ?rhs | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 267 | by safe (simp_all only: cball_eq_empty [of y e, symmetric] cball_eq_empty [of x d, symmetric]) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 268 | next | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 269 | case False | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 270 | with \<open>?lhs\<close> show ?rhs | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 271 | apply (auto simp: set_eq_subset cball_subset_cball_iff dist_norm norm_minus_commute algebra_simps) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 272 | apply (metis add_le_same_cancel1 le_add_same_cancel1 norm_ge_zero norm_pths(2) order_trans) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 273 | apply (metis add_increasing2 add_le_imp_le_right eq_iff norm_ge_zero) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 274 | done | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 275 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 276 | next | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 277 | assume ?rhs then show ?lhs | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 278 | by (auto simp: set_eq_subset cball_subset_cball_iff) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 279 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 280 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 281 | lemma ball_eq_cball_iff: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 282 | fixes x :: "'a :: euclidean_space" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 283 | shows "ball x d = cball y e \<longleftrightarrow> d \<le> 0 \<and> e < 0" (is "?lhs = ?rhs") | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 284 | proof | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 285 | assume ?lhs | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 286 | then show ?rhs | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 287 | apply (auto simp: set_eq_subset ball_subset_cball_iff cball_subset_ball_iff algebra_simps) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 288 | apply (metis add_increasing2 add_le_cancel_right add_less_same_cancel1 dist_not_less_zero less_le_trans zero_le_dist) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 289 | apply (metis add_less_same_cancel1 dist_not_less_zero less_le_trans not_le) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 290 | using \<open>?lhs\<close> ball_eq_empty cball_eq_empty apply blast+ | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 291 | done | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 292 | next | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 293 | assume ?rhs then show ?lhs by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 294 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 295 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 296 | lemma cball_eq_ball_iff: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 297 | fixes x :: "'a :: euclidean_space" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 298 | shows "cball x d = ball y e \<longleftrightarrow> d < 0 \<and> e \<le> 0" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 299 | using ball_eq_cball_iff by blast | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 300 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 301 | lemma finite_ball_avoid: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 302 | fixes S :: "'a :: euclidean_space set" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 303 | assumes "open S" "finite X" "p \<in> S" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 304 | shows "\<exists>e>0. \<forall>w\<in>ball p e. w\<in>S \<and> (w\<noteq>p \<longrightarrow> w\<notin>X)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 305 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 306 | obtain e1 where "0 < e1" and e1_b:"ball p e1 \<subseteq> S" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 307 | using open_contains_ball_eq[OF \<open>open S\<close>] assms by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 308 | obtain e2 where "0 < e2" and "\<forall>x\<in>X. x \<noteq> p \<longrightarrow> e2 \<le> dist p x" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 309 | using finite_set_avoid[OF \<open>finite X\<close>,of p] by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 310 | hence "\<forall>w\<in>ball p (min e1 e2). w\<in>S \<and> (w\<noteq>p \<longrightarrow> w\<notin>X)" using e1_b by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 311 | thus "\<exists>e>0. \<forall>w\<in>ball p e. w \<in> S \<and> (w \<noteq> p \<longrightarrow> w \<notin> X)" using \<open>e2>0\<close> \<open>e1>0\<close> | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 312 | apply (rule_tac x="min e1 e2" in exI) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 313 | by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 314 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 315 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 316 | lemma finite_cball_avoid: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 317 | fixes S :: "'a :: euclidean_space set" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 318 | assumes "open S" "finite X" "p \<in> S" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 319 | shows "\<exists>e>0. \<forall>w\<in>cball p e. w\<in>S \<and> (w\<noteq>p \<longrightarrow> w\<notin>X)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 320 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 321 | obtain e1 where "e1>0" and e1: "\<forall>w\<in>ball p e1. w\<in>S \<and> (w\<noteq>p \<longrightarrow> w\<notin>X)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 322 | using finite_ball_avoid[OF assms] by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 323 | define e2 where "e2 \<equiv> e1/2" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 324 | have "e2>0" and "e2 < e1" unfolding e2_def using \<open>e1>0\<close> by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 325 | then have "cball p e2 \<subseteq> ball p e1" by (subst cball_subset_ball_iff,auto) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 326 | then show "\<exists>e>0. \<forall>w\<in>cball p e. w \<in> S \<and> (w \<noteq> p \<longrightarrow> w \<notin> X)" using \<open>e2>0\<close> e1 by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 327 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 328 | |
| 69619 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 329 | lemma dim_cball: | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 330 | assumes "e > 0" | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 331 |   shows "dim (cball (0 :: 'n::euclidean_space) e) = DIM('n)"
 | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 332 | proof - | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 333 |   {
 | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 334 | fix x :: "'n::euclidean_space" | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 335 | define y where "y = (e / norm x) *\<^sub>R x" | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 336 | then have "y \<in> cball 0 e" | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 337 | using assms by auto | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 338 | moreover have *: "x = (norm x / e) *\<^sub>R y" | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 339 | using y_def assms by simp | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 340 | moreover from * have "x = (norm x/e) *\<^sub>R y" | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 341 | by auto | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 342 | ultimately have "x \<in> span (cball 0 e)" | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 343 | using span_scale[of y "cball 0 e" "norm x/e"] | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 344 | span_superset[of "cball 0 e"] | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 345 | by (simp add: span_base) | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 346 | } | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 347 | then have "span (cball 0 e) = (UNIV :: 'n::euclidean_space set)" | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 348 | by auto | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 349 | then show ?thesis | 
| 71172 | 350 | using dim_span[of "cball (0 :: 'n::euclidean_space) e"] by (auto) | 
| 69619 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 351 | qed | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 352 | |
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 353 | |
| 60420 | 354 | subsection \<open>Boxes\<close> | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 355 | |
| 71033 | 356 | abbreviation\<^marker>\<open>tag important\<close> One :: "'a::euclidean_space" where | 
| 357 | "One \<equiv> \<Sum>Basis" | |
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 358 | |
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63105diff
changeset | 359 | lemma One_non_0: assumes "One = (0::'a::euclidean_space)" shows False | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63105diff
changeset | 360 | proof - | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63105diff
changeset | 361 | have "dependent (Basis :: 'a set)" | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63105diff
changeset | 362 | apply (simp add: dependent_finite) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63105diff
changeset | 363 | apply (rule_tac x="\<lambda>i. 1" in exI) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63105diff
changeset | 364 | using SOME_Basis apply (auto simp: assms) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63105diff
changeset | 365 | done | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63105diff
changeset | 366 | with independent_Basis show False by force | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63105diff
changeset | 367 | qed | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63105diff
changeset | 368 | |
| 71033 | 369 | corollary\<^marker>\<open>tag unimportant\<close> One_neq_0[iff]: "One \<noteq> 0" | 
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63105diff
changeset | 370 | by (metis One_non_0) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63105diff
changeset | 371 | |
| 71033 | 372 | corollary\<^marker>\<open>tag unimportant\<close> Zero_neq_One[iff]: "0 \<noteq> One" | 
| 63114 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63105diff
changeset | 373 | by (metis One_non_0) | 
| 
27afe7af7379
Lots of new material for multivariate analysis
 paulson <lp15@cam.ac.uk> parents: 
63105diff
changeset | 374 | |
| 71033 | 375 | definition\<^marker>\<open>tag important\<close> (in euclidean_space) eucl_less (infix "<e" 50) where | 
| 376 | "eucl_less a b \<longleftrightarrow> (\<forall>i\<in>Basis. a \<bullet> i < b \<bullet> i)" | |
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 377 | |
| 70136 | 378 | definition\<^marker>\<open>tag important\<close> box_eucl_less: "box a b = {x. a <e x \<and> x <e b}"
 | 
| 379 | definition\<^marker>\<open>tag important\<close> "cbox a b = {x. \<forall>i\<in>Basis. a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i}"
 | |
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 380 | |
| 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 381 | lemma box_def: "box a b = {x. \<forall>i\<in>Basis. a \<bullet> i < x \<bullet> i \<and> x \<bullet> i < b \<bullet> i}"
 | 
| 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 382 | and in_box_eucl_less: "x \<in> box a b \<longleftrightarrow> a <e x \<and> x <e b" | 
| 56188 | 383 | and mem_box: "x \<in> box a b \<longleftrightarrow> (\<forall>i\<in>Basis. a \<bullet> i < x \<bullet> i \<and> x \<bullet> i < b \<bullet> i)" | 
| 384 | "x \<in> cbox a b \<longleftrightarrow> (\<forall>i\<in>Basis. a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i)" | |
| 385 | by (auto simp: box_eucl_less eucl_less_def cbox_def) | |
| 386 | ||
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 387 | lemma cbox_Pair_eq: "cbox (a, c) (b, d) = cbox a b \<times> cbox c d" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 388 | by (force simp: cbox_def Basis_prod_def) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 389 | |
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 390 | lemma cbox_Pair_iff [iff]: "(x, y) \<in> cbox (a, c) (b, d) \<longleftrightarrow> x \<in> cbox a b \<and> y \<in> cbox c d" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 391 | by (force simp: cbox_Pair_eq) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 392 | |
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 393 | lemma cbox_Complex_eq: "cbox (Complex a c) (Complex b d) = (\<lambda>(x,y). Complex x y) ` (cbox a b \<times> cbox c d)" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 394 | apply (auto simp: cbox_def Basis_complex_def) | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 395 | apply (rule_tac x = "(Re x, Im x)" in image_eqI) | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 396 | using complex_eq by auto | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65585diff
changeset | 397 | |
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 398 | lemma cbox_Pair_eq_0: "cbox (a, c) (b, d) = {} \<longleftrightarrow> cbox a b = {} \<or> cbox c d = {}"
 | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 399 | by (force simp: cbox_Pair_eq) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 400 | |
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 401 | lemma swap_cbox_Pair [simp]: "prod.swap ` cbox (c, a) (d, b) = cbox (a,c) (b,d)" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 402 | by auto | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 403 | |
| 56188 | 404 | lemma mem_box_real[simp]: | 
| 405 | "(x::real) \<in> box a b \<longleftrightarrow> a < x \<and> x < b" | |
| 406 | "(x::real) \<in> cbox a b \<longleftrightarrow> a \<le> x \<and> x \<le> b" | |
| 407 | by (auto simp: mem_box) | |
| 408 | ||
| 409 | lemma box_real[simp]: | |
| 410 | fixes a b:: real | |
| 411 |   shows "box a b = {a <..< b}" "cbox a b = {a .. b}"
 | |
| 412 | by auto | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50324diff
changeset | 413 | |
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 414 | lemma box_Int_box: | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 415 | fixes a :: "'a::euclidean_space" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 416 | shows "box a b \<inter> box c d = | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 417 | box (\<Sum>i\<in>Basis. max (a\<bullet>i) (c\<bullet>i) *\<^sub>R i) (\<Sum>i\<in>Basis. min (b\<bullet>i) (d\<bullet>i) *\<^sub>R i)" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 418 | unfolding set_eq_iff and Int_iff and mem_box by auto | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 419 | |
| 50087 | 420 | lemma rational_boxes: | 
| 61076 | 421 | fixes x :: "'a::euclidean_space" | 
| 53291 | 422 | assumes "e > 0" | 
| 66643 
f7e38b8583a0
Correction of typos and a bit of streamlining
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 423 | shows "\<exists>a b. (\<forall>i\<in>Basis. a \<bullet> i \<in> \<rat> \<and> b \<bullet> i \<in> \<rat>) \<and> x \<in> box a b \<and> box a b \<subseteq> ball x e" | 
| 50087 | 424 | proof - | 
| 63040 | 425 |   define e' where "e' = e / (2 * sqrt (real (DIM ('a))))"
 | 
| 53291 | 426 | then have e: "e' > 0" | 
| 71172 | 427 | using assms by (auto) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50324diff
changeset | 428 | have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> y < x \<bullet> i \<and> x \<bullet> i - y < e'" (is "\<forall>i. ?th i") | 
| 50087 | 429 | proof | 
| 53255 | 430 | fix i | 
| 431 | from Rats_dense_in_real[of "x \<bullet> i - e'" "x \<bullet> i"] e | |
| 432 | show "?th i" by auto | |
| 50087 | 433 | qed | 
| 55522 | 434 | from choice[OF this] obtain a where | 
| 435 | a: "\<forall>xa. a xa \<in> \<rat> \<and> a xa < x \<bullet> xa \<and> x \<bullet> xa - a xa < e'" .. | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50324diff
changeset | 436 | have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> x \<bullet> i < y \<and> y - x \<bullet> i < e'" (is "\<forall>i. ?th i") | 
| 50087 | 437 | proof | 
| 53255 | 438 | fix i | 
| 439 | from Rats_dense_in_real[of "x \<bullet> i" "x \<bullet> i + e'"] e | |
| 440 | show "?th i" by auto | |
| 50087 | 441 | qed | 
| 55522 | 442 | from choice[OF this] obtain b where | 
| 443 | b: "\<forall>xa. b xa \<in> \<rat> \<and> x \<bullet> xa < b xa \<and> b xa - x \<bullet> xa < e'" .. | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50324diff
changeset | 444 | let ?a = "\<Sum>i\<in>Basis. a i *\<^sub>R i" and ?b = "\<Sum>i\<in>Basis. b i *\<^sub>R i" | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50324diff
changeset | 445 | show ?thesis | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50324diff
changeset | 446 | proof (rule exI[of _ ?a], rule exI[of _ ?b], safe) | 
| 53255 | 447 | fix y :: 'a | 
| 448 | assume *: "y \<in> box ?a ?b" | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52625diff
changeset | 449 | have "dist x y = sqrt (\<Sum>i\<in>Basis. (dist (x \<bullet> i) (y \<bullet> i))\<^sup>2)" | 
| 67155 | 450 | unfolding L2_set_def[symmetric] by (rule euclidean_dist_l2) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50324diff
changeset | 451 |     also have "\<dots> < sqrt (\<Sum>(i::'a)\<in>Basis. e^2 / real (DIM('a)))"
 | 
| 64267 | 452 | proof (rule real_sqrt_less_mono, rule sum_strict_mono) | 
| 53255 | 453 | fix i :: "'a" | 
| 454 | assume i: "i \<in> Basis" | |
| 455 | have "a i < y\<bullet>i \<and> y\<bullet>i < b i" | |
| 456 | using * i by (auto simp: box_def) | |
| 457 | moreover have "a i < x\<bullet>i" "x\<bullet>i - a i < e'" | |
| 458 | using a by auto | |
| 459 | moreover have "x\<bullet>i < b i" "b i - x\<bullet>i < e'" | |
| 460 | using b by auto | |
| 461 | ultimately have "\<bar>x\<bullet>i - y\<bullet>i\<bar> < 2 * e'" | |
| 462 | by auto | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50324diff
changeset | 463 |       then have "dist (x \<bullet> i) (y \<bullet> i) < e/sqrt (real (DIM('a)))"
 | 
| 50087 | 464 | unfolding e'_def by (auto simp: dist_real_def) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52625diff
changeset | 465 |       then have "(dist (x \<bullet> i) (y \<bullet> i))\<^sup>2 < (e/sqrt (real (DIM('a))))\<^sup>2"
 | 
| 50087 | 466 | by (rule power_strict_mono) auto | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52625diff
changeset | 467 |       then show "(dist (x \<bullet> i) (y \<bullet> i))\<^sup>2 < e\<^sup>2 / real DIM('a)"
 | 
| 50087 | 468 | by (simp add: power_divide) | 
| 469 | qed auto | |
| 53255 | 470 | also have "\<dots> = e" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61552diff
changeset | 471 | using \<open>0 < e\<close> by simp | 
| 53255 | 472 | finally show "y \<in> ball x e" | 
| 473 | by (auto simp: ball_def) | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50324diff
changeset | 474 | qed (insert a b, auto simp: box_def) | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50324diff
changeset | 475 | qed | 
| 51103 | 476 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50324diff
changeset | 477 | lemma open_UNION_box: | 
| 61076 | 478 | fixes M :: "'a::euclidean_space set" | 
| 53282 | 479 | assumes "open M" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50324diff
changeset | 480 | defines "a' \<equiv> \<lambda>f :: 'a \<Rightarrow> real \<times> real. (\<Sum>(i::'a)\<in>Basis. fst (f i) *\<^sub>R i)" | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50324diff
changeset | 481 | defines "b' \<equiv> \<lambda>f :: 'a \<Rightarrow> real \<times> real. (\<Sum>(i::'a)\<in>Basis. snd (f i) *\<^sub>R i)" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
52625diff
changeset | 482 |   defines "I \<equiv> {f\<in>Basis \<rightarrow>\<^sub>E \<rat> \<times> \<rat>. box (a' f) (b' f) \<subseteq> M}"
 | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50324diff
changeset | 483 | shows "M = (\<Union>f\<in>I. box (a' f) (b' f))" | 
| 52624 | 484 | proof - | 
| 60462 | 485 | have "x \<in> (\<Union>f\<in>I. box (a' f) (b' f))" if "x \<in> M" for x | 
| 486 | proof - | |
| 52624 | 487 | obtain e where e: "e > 0" "ball x e \<subseteq> M" | 
| 60420 | 488 | using openE[OF \<open>open M\<close> \<open>x \<in> M\<close>] by auto | 
| 53282 | 489 | moreover obtain a b where ab: | 
| 490 | "x \<in> box a b" | |
| 491 | "\<forall>i \<in> Basis. a \<bullet> i \<in> \<rat>" | |
| 492 | "\<forall>i\<in>Basis. b \<bullet> i \<in> \<rat>" | |
| 493 | "box a b \<subseteq> ball x e" | |
| 52624 | 494 | using rational_boxes[OF e(1)] by metis | 
| 60462 | 495 | ultimately show ?thesis | 
| 52624 | 496 | by (intro UN_I[of "\<lambda>i\<in>Basis. (a \<bullet> i, b \<bullet> i)"]) | 
| 497 | (auto simp: euclidean_representation I_def a'_def b'_def) | |
| 60462 | 498 | qed | 
| 52624 | 499 | then show ?thesis by (auto simp: I_def) | 
| 500 | qed | |
| 501 | ||
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 502 | corollary open_countable_Union_open_box: | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 503 | fixes S :: "'a :: euclidean_space set" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 504 | assumes "open S" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 505 | obtains \<D> where "countable \<D>" "\<D> \<subseteq> Pow S" "\<And>X. X \<in> \<D> \<Longrightarrow> \<exists>a b. X = box a b" "\<Union>\<D> = S" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 506 | proof - | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 507 | let ?a = "\<lambda>f. (\<Sum>(i::'a)\<in>Basis. fst (f i) *\<^sub>R i)" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 508 | let ?b = "\<lambda>f. (\<Sum>(i::'a)\<in>Basis. snd (f i) *\<^sub>R i)" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 509 |   let ?I = "{f\<in>Basis \<rightarrow>\<^sub>E \<rat> \<times> \<rat>. box (?a f) (?b f) \<subseteq> S}"
 | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 510 | let ?\<D> = "(\<lambda>f. box (?a f) (?b f)) ` ?I" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 511 | show ?thesis | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 512 | proof | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 513 | have "countable ?I" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 514 | by (simp add: countable_PiE countable_rat) | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 515 | then show "countable ?\<D>" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 516 | by blast | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 517 | show "\<Union>?\<D> = S" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 518 | using open_UNION_box [OF assms] by metis | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 519 | qed auto | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 520 | qed | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 521 | |
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 522 | lemma rational_cboxes: | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 523 | fixes x :: "'a::euclidean_space" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 524 | assumes "e > 0" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 525 | shows "\<exists>a b. (\<forall>i\<in>Basis. a \<bullet> i \<in> \<rat> \<and> b \<bullet> i \<in> \<rat>) \<and> x \<in> cbox a b \<and> cbox a b \<subseteq> ball x e" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 526 | proof - | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 527 |   define e' where "e' = e / (2 * sqrt (real (DIM ('a))))"
 | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 528 | then have e: "e' > 0" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 529 | using assms by auto | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 530 | have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> y < x \<bullet> i \<and> x \<bullet> i - y < e'" (is "\<forall>i. ?th i") | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 531 | proof | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 532 | fix i | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 533 | from Rats_dense_in_real[of "x \<bullet> i - e'" "x \<bullet> i"] e | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 534 | show "?th i" by auto | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 535 | qed | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 536 | from choice[OF this] obtain a where | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 537 | a: "\<forall>u. a u \<in> \<rat> \<and> a u < x \<bullet> u \<and> x \<bullet> u - a u < e'" .. | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 538 | have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> x \<bullet> i < y \<and> y - x \<bullet> i < e'" (is "\<forall>i. ?th i") | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 539 | proof | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 540 | fix i | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 541 | from Rats_dense_in_real[of "x \<bullet> i" "x \<bullet> i + e'"] e | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 542 | show "?th i" by auto | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 543 | qed | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 544 | from choice[OF this] obtain b where | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 545 | b: "\<forall>u. b u \<in> \<rat> \<and> x \<bullet> u < b u \<and> b u - x \<bullet> u < e'" .. | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 546 | let ?a = "\<Sum>i\<in>Basis. a i *\<^sub>R i" and ?b = "\<Sum>i\<in>Basis. b i *\<^sub>R i" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 547 | show ?thesis | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 548 | proof (rule exI[of _ ?a], rule exI[of _ ?b], safe) | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 549 | fix y :: 'a | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 550 | assume *: "y \<in> cbox ?a ?b" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 551 | have "dist x y = sqrt (\<Sum>i\<in>Basis. (dist (x \<bullet> i) (y \<bullet> i))\<^sup>2)" | 
| 67155 | 552 | unfolding L2_set_def[symmetric] by (rule euclidean_dist_l2) | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 553 |     also have "\<dots> < sqrt (\<Sum>(i::'a)\<in>Basis. e^2 / real (DIM('a)))"
 | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 554 | proof (rule real_sqrt_less_mono, rule sum_strict_mono) | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 555 | fix i :: "'a" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 556 | assume i: "i \<in> Basis" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 557 | have "a i \<le> y\<bullet>i \<and> y\<bullet>i \<le> b i" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 558 | using * i by (auto simp: cbox_def) | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 559 | moreover have "a i < x\<bullet>i" "x\<bullet>i - a i < e'" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 560 | using a by auto | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 561 | moreover have "x\<bullet>i < b i" "b i - x\<bullet>i < e'" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 562 | using b by auto | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 563 | ultimately have "\<bar>x\<bullet>i - y\<bullet>i\<bar> < 2 * e'" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 564 | by auto | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 565 |       then have "dist (x \<bullet> i) (y \<bullet> i) < e/sqrt (real (DIM('a)))"
 | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 566 | unfolding e'_def by (auto simp: dist_real_def) | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 567 |       then have "(dist (x \<bullet> i) (y \<bullet> i))\<^sup>2 < (e/sqrt (real (DIM('a))))\<^sup>2"
 | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 568 | by (rule power_strict_mono) auto | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 569 |       then show "(dist (x \<bullet> i) (y \<bullet> i))\<^sup>2 < e\<^sup>2 / real DIM('a)"
 | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 570 | by (simp add: power_divide) | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 571 | qed auto | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 572 | also have "\<dots> = e" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 573 | using \<open>0 < e\<close> by simp | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 574 | finally show "y \<in> ball x e" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 575 | by (auto simp: ball_def) | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 576 | next | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 577 | show "x \<in> cbox (\<Sum>i\<in>Basis. a i *\<^sub>R i) (\<Sum>i\<in>Basis. b i *\<^sub>R i)" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 578 | using a b less_imp_le by (auto simp: cbox_def) | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 579 | qed (use a b cbox_def in auto) | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 580 | qed | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 581 | |
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 582 | lemma open_UNION_cbox: | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 583 | fixes M :: "'a::euclidean_space set" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 584 | assumes "open M" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 585 | defines "a' \<equiv> \<lambda>f. (\<Sum>(i::'a)\<in>Basis. fst (f i) *\<^sub>R i)" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 586 | defines "b' \<equiv> \<lambda>f. (\<Sum>(i::'a)\<in>Basis. snd (f i) *\<^sub>R i)" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 587 |   defines "I \<equiv> {f\<in>Basis \<rightarrow>\<^sub>E \<rat> \<times> \<rat>. cbox (a' f) (b' f) \<subseteq> M}"
 | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 588 | shows "M = (\<Union>f\<in>I. cbox (a' f) (b' f))" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 589 | proof - | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 590 | have "x \<in> (\<Union>f\<in>I. cbox (a' f) (b' f))" if "x \<in> M" for x | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 591 | proof - | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 592 | obtain e where e: "e > 0" "ball x e \<subseteq> M" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 593 | using openE[OF \<open>open M\<close> \<open>x \<in> M\<close>] by auto | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 594 | moreover obtain a b where ab: "x \<in> cbox a b" "\<forall>i \<in> Basis. a \<bullet> i \<in> \<rat>" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 595 | "\<forall>i \<in> Basis. b \<bullet> i \<in> \<rat>" "cbox a b \<subseteq> ball x e" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 596 | using rational_cboxes[OF e(1)] by metis | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 597 | ultimately show ?thesis | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 598 | by (intro UN_I[of "\<lambda>i\<in>Basis. (a \<bullet> i, b \<bullet> i)"]) | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 599 | (auto simp: euclidean_representation I_def a'_def b'_def) | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 600 | qed | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 601 | then show ?thesis by (auto simp: I_def) | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 602 | qed | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 603 | |
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 604 | corollary open_countable_Union_open_cbox: | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 605 | fixes S :: "'a :: euclidean_space set" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 606 | assumes "open S" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 607 | obtains \<D> where "countable \<D>" "\<D> \<subseteq> Pow S" "\<And>X. X \<in> \<D> \<Longrightarrow> \<exists>a b. X = cbox a b" "\<Union>\<D> = S" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 608 | proof - | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 609 | let ?a = "\<lambda>f. (\<Sum>(i::'a)\<in>Basis. fst (f i) *\<^sub>R i)" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 610 | let ?b = "\<lambda>f. (\<Sum>(i::'a)\<in>Basis. snd (f i) *\<^sub>R i)" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 611 |   let ?I = "{f\<in>Basis \<rightarrow>\<^sub>E \<rat> \<times> \<rat>. cbox (?a f) (?b f) \<subseteq> S}"
 | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 612 | let ?\<D> = "(\<lambda>f. cbox (?a f) (?b f)) ` ?I" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 613 | show ?thesis | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 614 | proof | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 615 | have "countable ?I" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 616 | by (simp add: countable_PiE countable_rat) | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 617 | then show "countable ?\<D>" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 618 | by blast | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 619 | show "\<Union>?\<D> = S" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 620 | using open_UNION_cbox [OF assms] by metis | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 621 | qed auto | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 622 | qed | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 623 | |
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 624 | lemma box_eq_empty: | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 625 | fixes a :: "'a::euclidean_space" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 626 |   shows "(box a b = {} \<longleftrightarrow> (\<exists>i\<in>Basis. b\<bullet>i \<le> a\<bullet>i))" (is ?th1)
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 627 |     and "(cbox a b = {} \<longleftrightarrow> (\<exists>i\<in>Basis. b\<bullet>i < a\<bullet>i))" (is ?th2)
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 628 | proof - | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 629 |   {
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 630 | fix i x | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 631 | assume i: "i\<in>Basis" and as:"b\<bullet>i \<le> a\<bullet>i" and x:"x\<in>box a b" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 632 | then have "a \<bullet> i < x \<bullet> i \<and> x \<bullet> i < b \<bullet> i" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 633 | unfolding mem_box by (auto simp: box_def) | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 634 | then have "a\<bullet>i < b\<bullet>i" by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 635 | then have False using as by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 636 | } | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 637 | moreover | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 638 |   {
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 639 | assume as: "\<forall>i\<in>Basis. \<not> (b\<bullet>i \<le> a\<bullet>i)" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 640 | let ?x = "(1/2) *\<^sub>R (a + b)" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 641 |     {
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 642 | fix i :: 'a | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 643 | assume i: "i \<in> Basis" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 644 | have "a\<bullet>i < b\<bullet>i" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 645 | using as[THEN bspec[where x=i]] i by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 646 | then have "a\<bullet>i < ((1/2) *\<^sub>R (a+b)) \<bullet> i" "((1/2) *\<^sub>R (a+b)) \<bullet> i < b\<bullet>i" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 647 | by (auto simp: inner_add_left) | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 648 | } | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 649 |     then have "box a b \<noteq> {}"
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 650 | using mem_box(1)[of "?x" a b] by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 651 | } | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 652 | ultimately show ?th1 by blast | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 653 | |
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 654 |   {
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 655 | fix i x | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 656 | assume i: "i \<in> Basis" and as:"b\<bullet>i < a\<bullet>i" and x:"x\<in>cbox a b" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 657 | then have "a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 658 | unfolding mem_box by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 659 | then have "a\<bullet>i \<le> b\<bullet>i" by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 660 | then have False using as by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 661 | } | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 662 | moreover | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 663 |   {
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 664 | assume as:"\<forall>i\<in>Basis. \<not> (b\<bullet>i < a\<bullet>i)" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 665 | let ?x = "(1/2) *\<^sub>R (a + b)" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 666 |     {
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 667 | fix i :: 'a | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 668 | assume i:"i \<in> Basis" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 669 | have "a\<bullet>i \<le> b\<bullet>i" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 670 | using as[THEN bspec[where x=i]] i by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 671 | then have "a\<bullet>i \<le> ((1/2) *\<^sub>R (a+b)) \<bullet> i" "((1/2) *\<^sub>R (a+b)) \<bullet> i \<le> b\<bullet>i" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 672 | by (auto simp: inner_add_left) | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 673 | } | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 674 |     then have "cbox a b \<noteq> {}"
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 675 | using mem_box(2)[of "?x" a b] by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 676 | } | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 677 | ultimately show ?th2 by blast | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 678 | qed | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 679 | |
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 680 | lemma box_ne_empty: | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 681 | fixes a :: "'a::euclidean_space" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 682 |   shows "cbox a b \<noteq> {} \<longleftrightarrow> (\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i)"
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 683 |   and "box a b \<noteq> {} \<longleftrightarrow> (\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i)"
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 684 | unfolding box_eq_empty[of a b] by fastforce+ | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 685 | |
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 686 | lemma | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 687 | fixes a :: "'a::euclidean_space" | 
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 688 |   shows cbox_sing [simp]: "cbox a a = {a}"
 | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 689 |     and box_sing [simp]: "box a a = {}"
 | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 690 | unfolding set_eq_iff mem_box eq_iff [symmetric] | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 691 | by (auto intro!: euclidean_eqI[where 'a='a]) | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 692 | (metis all_not_in_conv nonempty_Basis) | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 693 | |
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 694 | lemma subset_box_imp: | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 695 | fixes a :: "'a::euclidean_space" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 696 | shows "(\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i) \<Longrightarrow> cbox c d \<subseteq> cbox a b" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 697 | and "(\<forall>i\<in>Basis. a\<bullet>i < c\<bullet>i \<and> d\<bullet>i < b\<bullet>i) \<Longrightarrow> cbox c d \<subseteq> box a b" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 698 | and "(\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i) \<Longrightarrow> box c d \<subseteq> cbox a b" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 699 | and "(\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i) \<Longrightarrow> box c d \<subseteq> box a b" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 700 | unfolding subset_eq[unfolded Ball_def] unfolding mem_box | 
| 58757 | 701 | by (best intro: order_trans less_le_trans le_less_trans less_imp_le)+ | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 702 | |
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 703 | lemma box_subset_cbox: | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 704 | fixes a :: "'a::euclidean_space" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 705 | shows "box a b \<subseteq> cbox a b" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 706 | unfolding subset_eq [unfolded Ball_def] mem_box | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 707 | by (fast intro: less_imp_le) | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 708 | |
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 709 | lemma subset_box: | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 710 | fixes a :: "'a::euclidean_space" | 
| 64539 | 711 | shows "cbox c d \<subseteq> cbox a b \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i \<le> d\<bullet>i) \<longrightarrow> (\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i)" (is ?th1) | 
| 712 | and "cbox c d \<subseteq> box a b \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i \<le> d\<bullet>i) \<longrightarrow> (\<forall>i\<in>Basis. a\<bullet>i < c\<bullet>i \<and> d\<bullet>i < b\<bullet>i)" (is ?th2) | |
| 713 | and "box c d \<subseteq> cbox a b \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i < d\<bullet>i) \<longrightarrow> (\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i)" (is ?th3) | |
| 714 | and "box c d \<subseteq> box a b \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i < d\<bullet>i) \<longrightarrow> (\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i)" (is ?th4) | |
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 715 | proof - | 
| 68120 | 716 | let ?lesscd = "\<forall>i\<in>Basis. c\<bullet>i < d\<bullet>i" | 
| 717 | let ?lerhs = "\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i" | |
| 718 | show ?th1 ?th2 | |
| 719 | by (fastforce simp: mem_box)+ | |
| 720 | have acdb: "a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i" | |
| 721 | if i: "i \<in> Basis" and box: "box c d \<subseteq> cbox a b" and cd: "\<And>i. i \<in> Basis \<Longrightarrow> c\<bullet>i < d\<bullet>i" for i | |
| 722 | proof - | |
| 723 |     have "box c d \<noteq> {}"
 | |
| 724 | using that | |
| 725 | unfolding box_eq_empty by force | |
| 726 |     { let ?x = "(\<Sum>j\<in>Basis. (if j=i then ((min (a\<bullet>j) (d\<bullet>j))+c\<bullet>j)/2 else (c\<bullet>j+d\<bullet>j)/2) *\<^sub>R j)::'a"
 | |
| 727 | assume *: "a\<bullet>i > c\<bullet>i" | |
| 728 | then have "c \<bullet> j < ?x \<bullet> j \<and> ?x \<bullet> j < d \<bullet> j" if "j \<in> Basis" for j | |
| 729 | using cd that by (fastforce simp add: i *) | |
| 730 | then have "?x \<in> box c d" | |
| 731 | unfolding mem_box by auto | |
| 732 | moreover have "?x \<notin> cbox a b" | |
| 733 | using i cd * by (force simp: mem_box) | |
| 734 | ultimately have False using box by auto | |
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 735 | } | 
| 68120 | 736 | then have "a\<bullet>i \<le> c\<bullet>i" by force | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 737 | moreover | 
| 68120 | 738 |     { let ?x = "(\<Sum>j\<in>Basis. (if j=i then ((max (b\<bullet>j) (c\<bullet>j))+d\<bullet>j)/2 else (c\<bullet>j+d\<bullet>j)/2) *\<^sub>R j)::'a"
 | 
| 739 | assume *: "b\<bullet>i < d\<bullet>i" | |
| 740 | then have "d \<bullet> j > ?x \<bullet> j \<and> ?x \<bullet> j > c \<bullet> j" if "j \<in> Basis" for j | |
| 741 | using cd that by (fastforce simp add: i *) | |
| 742 | then have "?x \<in> box c d" | |
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 743 | unfolding mem_box by auto | 
| 68120 | 744 | moreover have "?x \<notin> cbox a b" | 
| 745 | using i cd * by (force simp: mem_box) | |
| 746 | ultimately have False using box by auto | |
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 747 | } | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 748 | then have "b\<bullet>i \<ge> d\<bullet>i" by (rule ccontr) auto | 
| 68120 | 749 | ultimately show ?thesis by auto | 
| 750 | qed | |
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 751 | show ?th3 | 
| 68120 | 752 | using acdb by (fastforce simp add: mem_box) | 
| 753 | have acdb': "a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i" | |
| 754 | if "i \<in> Basis" "box c d \<subseteq> box a b" "\<And>i. i \<in> Basis \<Longrightarrow> c\<bullet>i < d\<bullet>i" for i | |
| 755 | using box_subset_cbox[of a b] that acdb by auto | |
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 756 | show ?th4 | 
| 68120 | 757 | using acdb' by (fastforce simp add: mem_box) | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 758 | qed | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 759 | |
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 760 | lemma eq_cbox: "cbox a b = cbox c d \<longleftrightarrow> cbox a b = {} \<and> cbox c d = {} \<or> a = c \<and> b = d"
 | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 761 | (is "?lhs = ?rhs") | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 762 | proof | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 763 | assume ?lhs | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 764 | then have "cbox a b \<subseteq> cbox c d" "cbox c d \<subseteq> cbox a b" | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 765 | by auto | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 766 | then show ?rhs | 
| 66643 
f7e38b8583a0
Correction of typos and a bit of streamlining
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 767 | by (force simp: subset_box box_eq_empty intro: antisym euclidean_eqI) | 
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 768 | next | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 769 | assume ?rhs | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 770 | then show ?lhs | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 771 | by force | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 772 | qed | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 773 | |
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 774 | lemma eq_cbox_box [simp]: "cbox a b = box c d \<longleftrightarrow> cbox a b = {} \<and> box c d = {}"
 | 
| 64539 | 775 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 776 | proof | 
| 68120 | 777 | assume L: ?lhs | 
| 778 | then have "cbox a b \<subseteq> box c d" "box c d \<subseteq> cbox a b" | |
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 779 | by auto | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 780 | then show ?rhs | 
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63955diff
changeset | 781 | apply (simp add: subset_box) | 
| 68120 | 782 | using L box_ne_empty box_sing apply (fastforce simp add:) | 
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 783 | done | 
| 68120 | 784 | qed force | 
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 785 | |
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 786 | lemma eq_box_cbox [simp]: "box a b = cbox c d \<longleftrightarrow> box a b = {} \<and> cbox c d = {}"
 | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 787 | by (metis eq_cbox_box) | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 788 | |
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 789 | lemma eq_box: "box a b = box c d \<longleftrightarrow> box a b = {} \<and> box c d = {} \<or> a = c \<and> b = d"
 | 
| 64539 | 790 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 791 | proof | 
| 68120 | 792 | assume L: ?lhs | 
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 793 | then have "box a b \<subseteq> box c d" "box c d \<subseteq> box a b" | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 794 | by auto | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 795 | then show ?rhs | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 796 | apply (simp add: subset_box) | 
| 68120 | 797 | using box_ne_empty(2) L | 
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 798 | apply auto | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 799 | apply (meson euclidean_eqI less_eq_real_def not_less)+ | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 800 | done | 
| 68120 | 801 | qed force | 
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 802 | |
| 66466 
aec5d9c88d69
More lemmas for HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 803 | lemma subset_box_complex: | 
| 66643 
f7e38b8583a0
Correction of typos and a bit of streamlining
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 804 | "cbox a b \<subseteq> cbox c d \<longleftrightarrow> | 
| 66466 
aec5d9c88d69
More lemmas for HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 805 | (Re a \<le> Re b \<and> Im a \<le> Im b) \<longrightarrow> Re a \<ge> Re c \<and> Im a \<ge> Im c \<and> Re b \<le> Re d \<and> Im b \<le> Im d" | 
| 66643 
f7e38b8583a0
Correction of typos and a bit of streamlining
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 806 | "cbox a b \<subseteq> box c d \<longleftrightarrow> | 
| 66466 
aec5d9c88d69
More lemmas for HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 807 | (Re a \<le> Re b \<and> Im a \<le> Im b) \<longrightarrow> Re a > Re c \<and> Im a > Im c \<and> Re b < Re d \<and> Im b < Im d" | 
| 
aec5d9c88d69
More lemmas for HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 808 | "box a b \<subseteq> cbox c d \<longleftrightarrow> | 
| 
aec5d9c88d69
More lemmas for HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 809 | (Re a < Re b \<and> Im a < Im b) \<longrightarrow> Re a \<ge> Re c \<and> Im a \<ge> Im c \<and> Re b \<le> Re d \<and> Im b \<le> Im d" | 
| 66643 
f7e38b8583a0
Correction of typos and a bit of streamlining
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 810 | "box a b \<subseteq> box c d \<longleftrightarrow> | 
| 66466 
aec5d9c88d69
More lemmas for HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 811 | (Re a < Re b \<and> Im a < Im b) \<longrightarrow> Re a \<ge> Re c \<and> Im a \<ge> Im c \<and> Re b \<le> Re d \<and> Im b \<le> Im d" | 
| 
aec5d9c88d69
More lemmas for HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 812 | by (subst subset_box; force simp: Basis_complex_def)+ | 
| 
aec5d9c88d69
More lemmas for HOL-Analysis
 Manuel Eberl <eberlm@in.tum.de> parents: 
66453diff
changeset | 813 | |
| 71029 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70952diff
changeset | 814 | lemma in_cbox_complex_iff: | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70952diff
changeset | 815 |   "x \<in> cbox a b \<longleftrightarrow> Re x \<in> {Re a..Re b} \<and> Im x \<in> {Im a..Im b}"
 | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70952diff
changeset | 816 | by (cases x; cases a; cases b) (auto simp: cbox_Complex_eq) | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70952diff
changeset | 817 | |
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70952diff
changeset | 818 | lemma box_Complex_eq: | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70952diff
changeset | 819 | "box (Complex a c) (Complex b d) = (\<lambda>(x,y). Complex x y) ` (box a b \<times> box c d)" | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70952diff
changeset | 820 | by (auto simp: box_def Basis_complex_def image_iff complex_eq_iff) | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70952diff
changeset | 821 | |
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70952diff
changeset | 822 | lemma in_box_complex_iff: | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70952diff
changeset | 823 |   "x \<in> box a b \<longleftrightarrow> Re x \<in> {Re a<..<Re b} \<and> Im x \<in> {Im a<..<Im b}"
 | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70952diff
changeset | 824 | by (cases x; cases a; cases b) (auto simp: box_Complex_eq) | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70952diff
changeset | 825 | |
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 826 | lemma Int_interval: | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 827 | fixes a :: "'a::euclidean_space" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 828 | shows "cbox a b \<inter> cbox c d = | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 829 | cbox (\<Sum>i\<in>Basis. max (a\<bullet>i) (c\<bullet>i) *\<^sub>R i) (\<Sum>i\<in>Basis. min (b\<bullet>i) (d\<bullet>i) *\<^sub>R i)" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 830 | unfolding set_eq_iff and Int_iff and mem_box | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 831 | by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 832 | |
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 833 | lemma disjoint_interval: | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 834 | fixes a::"'a::euclidean_space" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 835 |   shows "cbox a b \<inter> cbox c d = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i < a\<bullet>i \<or> d\<bullet>i < c\<bullet>i \<or> b\<bullet>i < c\<bullet>i \<or> d\<bullet>i < a\<bullet>i))" (is ?th1)
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 836 |     and "cbox a b \<inter> box c d = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i < a\<bullet>i \<or> d\<bullet>i \<le> c\<bullet>i \<or> b\<bullet>i \<le> c\<bullet>i \<or> d\<bullet>i \<le> a\<bullet>i))" (is ?th2)
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 837 |     and "box a b \<inter> cbox c d = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i \<le> a\<bullet>i \<or> d\<bullet>i < c\<bullet>i \<or> b\<bullet>i \<le> c\<bullet>i \<or> d\<bullet>i \<le> a\<bullet>i))" (is ?th3)
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 838 |     and "box a b \<inter> box c d = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i \<le> a\<bullet>i \<or> d\<bullet>i \<le> c\<bullet>i \<or> b\<bullet>i \<le> c\<bullet>i \<or> d\<bullet>i \<le> a\<bullet>i))" (is ?th4)
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 839 | proof - | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 840 | let ?z = "(\<Sum>i\<in>Basis. (((max (a\<bullet>i) (c\<bullet>i)) + (min (b\<bullet>i) (d\<bullet>i))) / 2) *\<^sub>R i)::'a" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 841 | have **: "\<And>P Q. (\<And>i :: 'a. i \<in> Basis \<Longrightarrow> Q ?z i \<Longrightarrow> P i) \<Longrightarrow> | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 842 | (\<And>i x :: 'a. i \<in> Basis \<Longrightarrow> P i \<Longrightarrow> Q x i) \<Longrightarrow> (\<forall>x. \<exists>i\<in>Basis. Q x i) \<longleftrightarrow> (\<exists>i\<in>Basis. P i)" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 843 | by blast | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 844 | note * = set_eq_iff Int_iff empty_iff mem_box ball_conj_distrib[symmetric] eq_False ball_simps(10) | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 845 | show ?th1 unfolding * by (intro **) auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 846 | show ?th2 unfolding * by (intro **) auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 847 | show ?th3 unfolding * by (intro **) auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 848 | show ?th4 unfolding * by (intro **) auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 849 | qed | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 850 | |
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 851 | lemma UN_box_eq_UNIV: "(\<Union>i::nat. box (- (real i *\<^sub>R One)) (real i *\<^sub>R One)) = UNIV" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 852 | proof - | 
| 61942 | 853 | have "\<bar>x \<bullet> b\<bar> < real_of_int (\<lceil>Max ((\<lambda>b. \<bar>x \<bullet> b\<bar>)`Basis)\<rceil> + 1)" | 
| 60462 | 854 | if [simp]: "b \<in> Basis" for x b :: 'a | 
| 855 | proof - | |
| 61942 | 856 | have "\<bar>x \<bullet> b\<bar> \<le> real_of_int \<lceil>\<bar>x \<bullet> b\<bar>\<rceil>" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61552diff
changeset | 857 | by (rule le_of_int_ceiling) | 
| 61942 | 858 | also have "\<dots> \<le> real_of_int \<lceil>Max ((\<lambda>b. \<bar>x \<bullet> b\<bar>)`Basis)\<rceil>" | 
| 59587 
8ea7b22525cb
Removed the obsolete functions "natfloor" and "natceiling"
 nipkow parents: 
58877diff
changeset | 859 | by (auto intro!: ceiling_mono) | 
| 61942 | 860 | also have "\<dots> < real_of_int (\<lceil>Max ((\<lambda>b. \<bar>x \<bullet> b\<bar>)`Basis)\<rceil> + 1)" | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 861 | by simp | 
| 60462 | 862 | finally show ?thesis . | 
| 863 | qed | |
| 864 | then have "\<exists>n::nat. \<forall>b\<in>Basis. \<bar>x \<bullet> b\<bar> < real n" for x :: 'a | |
| 59587 
8ea7b22525cb
Removed the obsolete functions "natfloor" and "natceiling"
 nipkow parents: 
58877diff
changeset | 865 | by (metis order.strict_trans reals_Archimedean2) | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 866 | moreover have "\<And>x b::'a. \<And>n::nat. \<bar>x \<bullet> b\<bar> < real n \<longleftrightarrow> - real n < x \<bullet> b \<and> x \<bullet> b < real n" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 867 | by auto | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 868 | ultimately show ?thesis | 
| 64267 | 869 | by (auto simp: box_def inner_sum_left inner_Basis sum.If_cases) | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 870 | qed | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 871 | |
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 872 | lemma image_affinity_cbox: fixes m::real | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 873 | fixes a b c :: "'a::euclidean_space" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 874 | shows "(\<lambda>x. m *\<^sub>R x + c) ` cbox a b = | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 875 |     (if cbox a b = {} then {}
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 876 | else (if 0 \<le> m then cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 877 | else cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c)))" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 878 | proof (cases "m = 0") | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 879 | case True | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 880 |   {
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 881 | fix x | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 882 | assume "\<forall>i\<in>Basis. x \<bullet> i \<le> c \<bullet> i" "\<forall>i\<in>Basis. c \<bullet> i \<le> x \<bullet> i" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 883 | then have "x = c" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 884 | by (simp add: dual_order.antisym euclidean_eqI) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 885 | } | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 886 | moreover have "c \<in> cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c)" | 
| 71172 | 887 | unfolding True by (auto) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 888 | ultimately show ?thesis using True by (auto simp: cbox_def) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 889 | next | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 890 | case False | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 891 |   {
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 892 | fix y | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 893 | assume "\<forall>i\<in>Basis. a \<bullet> i \<le> y \<bullet> i" "\<forall>i\<in>Basis. y \<bullet> i \<le> b \<bullet> i" "m > 0" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 894 | then have "\<forall>i\<in>Basis. (m *\<^sub>R a + c) \<bullet> i \<le> (m *\<^sub>R y + c) \<bullet> i" and "\<forall>i\<in>Basis. (m *\<^sub>R y + c) \<bullet> i \<le> (m *\<^sub>R b + c) \<bullet> i" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 895 | by (auto simp: inner_distrib) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 896 | } | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 897 | moreover | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 898 |   {
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 899 | fix y | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 900 | assume "\<forall>i\<in>Basis. a \<bullet> i \<le> y \<bullet> i" "\<forall>i\<in>Basis. y \<bullet> i \<le> b \<bullet> i" "m < 0" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 901 | then have "\<forall>i\<in>Basis. (m *\<^sub>R b + c) \<bullet> i \<le> (m *\<^sub>R y + c) \<bullet> i" and "\<forall>i\<in>Basis. (m *\<^sub>R y + c) \<bullet> i \<le> (m *\<^sub>R a + c) \<bullet> i" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 902 | by (auto simp: mult_left_mono_neg inner_distrib) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 903 | } | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 904 | moreover | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 905 |   {
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 906 | fix y | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 907 | assume "m > 0" and "\<forall>i\<in>Basis. (m *\<^sub>R a + c) \<bullet> i \<le> y \<bullet> i" and "\<forall>i\<in>Basis. y \<bullet> i \<le> (m *\<^sub>R b + c) \<bullet> i" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 908 | then have "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` cbox a b" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 909 | unfolding image_iff Bex_def mem_box | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 910 | apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 911 | apply (auto simp: pos_le_divide_eq pos_divide_le_eq mult.commute inner_distrib inner_diff_left) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 912 | done | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 913 | } | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 914 | moreover | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 915 |   {
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 916 | fix y | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 917 | assume "\<forall>i\<in>Basis. (m *\<^sub>R b + c) \<bullet> i \<le> y \<bullet> i" "\<forall>i\<in>Basis. y \<bullet> i \<le> (m *\<^sub>R a + c) \<bullet> i" "m < 0" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 918 | then have "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` cbox a b" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 919 | unfolding image_iff Bex_def mem_box | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 920 | apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 921 | apply (auto simp: neg_le_divide_eq neg_divide_le_eq mult.commute inner_distrib inner_diff_left) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 922 | done | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 923 | } | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 924 | ultimately show ?thesis using False by (auto simp: cbox_def) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 925 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 926 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 927 | lemma image_smult_cbox:"(\<lambda>x. m *\<^sub>R (x::_::euclidean_space)) ` cbox a b = | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 928 |   (if cbox a b = {} then {} else if 0 \<le> m then cbox (m *\<^sub>R a) (m *\<^sub>R b) else cbox (m *\<^sub>R b) (m *\<^sub>R a))"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 929 | using image_affinity_cbox[of m 0 a b] by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 930 | |
| 69619 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 931 | lemma swap_continuous: | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 932 | assumes "continuous_on (cbox (a,c) (b,d)) (\<lambda>(x,y). f x y)" | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 933 | shows "continuous_on (cbox (c,a) (d,b)) (\<lambda>(x, y). f y x)" | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 934 | proof - | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 935 | have "(\<lambda>(x, y). f y x) = (\<lambda>(x, y). f x y) \<circ> prod.swap" | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 936 | by auto | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 937 | then show ?thesis | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 938 | apply (rule ssubst) | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 939 | apply (rule continuous_on_compose) | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 940 | apply (simp add: split_def) | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 941 | apply (rule continuous_intros | simp add: assms)+ | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 942 | done | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 943 | qed | 
| 
3f7d8e05e0f2
split off Convex.thy: material that does not require Topology_Euclidean_Space
 immler parents: 
69618diff
changeset | 944 | |
| 69516 
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
 immler parents: 
69508diff
changeset | 945 | |
| 69508 | 946 | subsection \<open>General Intervals\<close> | 
| 67962 | 947 | |
| 70136 | 948 | definition\<^marker>\<open>tag important\<close> "is_interval (s::('a::euclidean_space) set) \<longleftrightarrow>
 | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 949 | (\<forall>a\<in>s. \<forall>b\<in>s. \<forall>x. (\<forall>i\<in>Basis. ((a\<bullet>i \<le> x\<bullet>i \<and> x\<bullet>i \<le> b\<bullet>i) \<or> (b\<bullet>i \<le> x\<bullet>i \<and> x\<bullet>i \<le> a\<bullet>i))) \<longrightarrow> x \<in> s)" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 950 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 951 | lemma is_interval_1: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 952 | "is_interval (s::real set) \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. \<forall> x. a \<le> x \<and> x \<le> b \<longrightarrow> x \<in> s)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 953 | unfolding is_interval_def by auto | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 954 | |
| 70019 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69922diff
changeset | 955 | lemma is_interval_Int: "is_interval X \<Longrightarrow> is_interval Y \<Longrightarrow> is_interval (X \<inter> Y)" | 
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 956 | unfolding is_interval_def | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 957 | by blast | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 958 | |
| 66089 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 959 | lemma is_interval_cbox [simp]: "is_interval (cbox a (b::'a::euclidean_space))" (is ?th1) | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 960 | and is_interval_box [simp]: "is_interval (box a b)" (is ?th2) | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 961 | unfolding is_interval_def mem_box Ball_def atLeastAtMost_iff | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 962 | by (meson order_trans le_less_trans less_le_trans less_trans)+ | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 963 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61552diff
changeset | 964 | lemma is_interval_empty [iff]: "is_interval {}"
 | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61552diff
changeset | 965 | unfolding is_interval_def by simp | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61552diff
changeset | 966 | |
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61552diff
changeset | 967 | lemma is_interval_univ [iff]: "is_interval UNIV" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61552diff
changeset | 968 | unfolding is_interval_def by simp | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 969 | |
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 970 | lemma mem_is_intervalI: | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 971 | assumes "is_interval s" | 
| 64539 | 972 | and "a \<in> s" "b \<in> s" | 
| 973 | and "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i \<or> b \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> a \<bullet> i" | |
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 974 | shows "x \<in> s" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 975 | by (rule assms(1)[simplified is_interval_def, rule_format, OF assms(2,3,4)]) | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 976 | |
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 977 | lemma interval_subst: | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 978 | fixes S::"'a::euclidean_space set" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 979 | assumes "is_interval S" | 
| 64539 | 980 | and "x \<in> S" "y j \<in> S" | 
| 981 | and "j \<in> Basis" | |
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 982 | shows "(\<Sum>i\<in>Basis. (if i = j then y i \<bullet> i else x \<bullet> i) *\<^sub>R i) \<in> S" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 983 | by (rule mem_is_intervalI[OF assms(1,2)]) (auto simp: assms) | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 984 | |
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 985 | lemma mem_box_componentwiseI: | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 986 | fixes S::"'a::euclidean_space set" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 987 | assumes "is_interval S" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 988 | assumes "\<And>i. i \<in> Basis \<Longrightarrow> x \<bullet> i \<in> ((\<lambda>x. x \<bullet> i) ` S)" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 989 | shows "x \<in> S" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 990 | proof - | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 991 | from assms have "\<forall>i \<in> Basis. \<exists>s \<in> S. x \<bullet> i = s \<bullet> i" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 992 | by auto | 
| 64539 | 993 | with finite_Basis obtain s and bs::"'a list" | 
| 994 | where s: "\<And>i. i \<in> Basis \<Longrightarrow> x \<bullet> i = s i \<bullet> i" "\<And>i. i \<in> Basis \<Longrightarrow> s i \<in> S" | |
| 995 | and bs: "set bs = Basis" "distinct bs" | |
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 996 | by (metis finite_distinct_list) | 
| 64539 | 997 | from nonempty_Basis s obtain j where j: "j \<in> Basis" "s j \<in> S" | 
| 998 | by blast | |
| 63040 | 999 | define y where | 
| 1000 | "y = rec_list (s j) (\<lambda>j _ Y. (\<Sum>i\<in>Basis. (if i = j then s i \<bullet> i else Y \<bullet> i) *\<^sub>R i))" | |
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1001 | have "x = (\<Sum>i\<in>Basis. (if i \<in> set bs then s i \<bullet> i else s j \<bullet> i) *\<^sub>R i)" | 
| 66643 
f7e38b8583a0
Correction of typos and a bit of streamlining
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 1002 | using bs by (auto simp: s(1)[symmetric] euclidean_representation) | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1003 | also have [symmetric]: "y bs = \<dots>" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1004 | using bs(2) bs(1)[THEN equalityD1] | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1005 | by (induct bs) (auto simp: y_def euclidean_representation intro!: euclidean_eqI[where 'a='a]) | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1006 | also have "y bs \<in> S" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1007 | using bs(1)[THEN equalityD1] | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1008 | apply (induct bs) | 
| 64539 | 1009 | apply (auto simp: y_def j) | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1010 | apply (rule interval_subst[OF assms(1)]) | 
| 64539 | 1011 | apply (auto simp: s) | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1012 | done | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1013 | finally show ?thesis . | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1014 | qed | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1015 | |
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62948diff
changeset | 1016 | lemma cbox01_nonempty [simp]: "cbox 0 One \<noteq> {}"
 | 
| 64267 | 1017 | by (simp add: box_ne_empty inner_Basis inner_sum_left sum_nonneg) | 
| 63007 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62948diff
changeset | 1018 | |
| 
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
 paulson <lp15@cam.ac.uk> parents: 
62948diff
changeset | 1019 | lemma box01_nonempty [simp]: "box 0 One \<noteq> {}"
 | 
| 66089 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1020 | by (simp add: box_ne_empty inner_Basis inner_sum_left) | 
| 63075 
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
 paulson <lp15@cam.ac.uk> parents: 
63040diff
changeset | 1021 | |
| 64773 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64758diff
changeset | 1022 | lemma empty_as_interval: "{} = cbox One (0::'a::euclidean_space)"
 | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64758diff
changeset | 1023 | using nonempty_Basis box01_nonempty box_eq_empty(1) box_ne_empty(1) by blast | 
| 
223b2ebdda79
Many new theorems, and more tidying
 paulson <lp15@cam.ac.uk> parents: 
64758diff
changeset | 1024 | |
| 66089 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1025 | lemma interval_subset_is_interval: | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1026 | assumes "is_interval S" | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1027 |   shows "cbox a b \<subseteq> S \<longleftrightarrow> cbox a b = {} \<or> a \<in> S \<and> b \<in> S" (is "?lhs = ?rhs")
 | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1028 | proof | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1029 | assume ?lhs | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1030 | then show ?rhs using box_ne_empty(1) mem_box(2) by fastforce | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1031 | next | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1032 | assume ?rhs | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1033 | have "cbox a b \<subseteq> S" if "a \<in> S" "b \<in> S" | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1034 | using assms unfolding is_interval_def | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1035 | apply (clarsimp simp add: mem_box) | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1036 | using that by blast | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1037 | with \<open>?rhs\<close> show ?lhs | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1038 | by blast | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1039 | qed | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 1040 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1041 | lemma is_real_interval_union: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1042 | "is_interval (X \<union> Y)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1043 |   if X: "is_interval X" and Y: "is_interval Y" and I: "(X \<noteq> {} \<Longrightarrow> Y \<noteq> {} \<Longrightarrow> X \<inter> Y \<noteq> {})"
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1044 | for X Y::"real set" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1045 | proof - | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1046 |   consider "X \<noteq> {}" "Y \<noteq> {}" | "X = {}" | "Y = {}" by blast
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1047 | then show ?thesis | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1048 | proof cases | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1049 | case 1 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1050 |     then obtain r where "r \<in> X \<or> X \<inter> Y = {}" "r \<in> Y \<or> X \<inter> Y = {}"
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1051 | by blast | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1052 | then show ?thesis | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1053 | using I 1 X Y unfolding is_interval_1 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1054 | by (metis (full_types) Un_iff le_cases) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1055 | qed (use that in auto) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1056 | qed | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1057 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1058 | lemma is_interval_translationI: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1059 | assumes "is_interval X" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1060 | shows "is_interval ((+) x ` X)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1061 | unfolding is_interval_def | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1062 | proof safe | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1063 | fix b d e | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1064 | assume "b \<in> X" "d \<in> X" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1065 | "\<forall>i\<in>Basis. (x + b) \<bullet> i \<le> e \<bullet> i \<and> e \<bullet> i \<le> (x + d) \<bullet> i \<or> | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1066 | (x + d) \<bullet> i \<le> e \<bullet> i \<and> e \<bullet> i \<le> (x + b) \<bullet> i" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1067 | hence "e - x \<in> X" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1068 | by (intro mem_is_intervalI[OF assms \<open>b \<in> X\<close> \<open>d \<in> X\<close>, of "e - x"]) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1069 | (auto simp: algebra_simps) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1070 | thus "e \<in> (+) x ` X" by force | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1071 | qed | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1072 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1073 | lemma is_interval_uminusI: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1074 | assumes "is_interval X" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1075 | shows "is_interval (uminus ` X)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1076 | unfolding is_interval_def | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1077 | proof safe | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1078 | fix b d e | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1079 | assume "b \<in> X" "d \<in> X" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1080 | "\<forall>i\<in>Basis. (- b) \<bullet> i \<le> e \<bullet> i \<and> e \<bullet> i \<le> (- d) \<bullet> i \<or> | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1081 | (- d) \<bullet> i \<le> e \<bullet> i \<and> e \<bullet> i \<le> (- b) \<bullet> i" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1082 | hence "- e \<in> X" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1083 | by (intro mem_is_intervalI[OF assms \<open>b \<in> X\<close> \<open>d \<in> X\<close>, of "- e"]) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1084 | (auto simp: algebra_simps) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1085 | thus "e \<in> uminus ` X" by force | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1086 | qed | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1087 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1088 | lemma is_interval_uminus[simp]: "is_interval (uminus ` x) = is_interval x" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1089 | using is_interval_uminusI[of x] is_interval_uminusI[of "uminus ` x"] | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1090 | by (auto simp: image_image) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1091 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1092 | lemma is_interval_neg_translationI: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1093 | assumes "is_interval X" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1094 | shows "is_interval ((-) x ` X)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1095 | proof - | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1096 | have "(-) x ` X = (+) x ` uminus ` X" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1097 | by (force simp: algebra_simps) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1098 | also have "is_interval \<dots>" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1099 | by (metis is_interval_uminusI is_interval_translationI assms) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1100 | finally show ?thesis . | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1101 | qed | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1102 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1103 | lemma is_interval_translation[simp]: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1104 | "is_interval ((+) x ` X) = is_interval X" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1105 | using is_interval_neg_translationI[of "(+) x ` X" x] | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1106 | by (auto intro!: is_interval_translationI simp: image_image) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1107 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1108 | lemma is_interval_minus_translation[simp]: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1109 | shows "is_interval ((-) x ` X) = is_interval X" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1110 | proof - | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1111 | have "(-) x ` X = (+) x ` uminus ` X" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1112 | by (force simp: algebra_simps) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1113 | also have "is_interval \<dots> = is_interval X" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1114 | by simp | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1115 | finally show ?thesis . | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1116 | qed | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1117 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1118 | lemma is_interval_minus_translation'[simp]: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1119 | shows "is_interval ((\<lambda>x. x - c) ` X) = is_interval X" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1120 | using is_interval_translation[of "-c" X] | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1121 | by (metis image_cong uminus_add_conv_diff) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67613diff
changeset | 1122 | |
| 71028 
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
 immler parents: 
71025diff
changeset | 1123 | lemma is_interval_cball_1[intro, simp]: "is_interval (cball a b)" for a b::real | 
| 
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
 immler parents: 
71025diff
changeset | 1124 | by (simp add: cball_eq_atLeastAtMost is_interval_def) | 
| 
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
 immler parents: 
71025diff
changeset | 1125 | |
| 
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
 immler parents: 
71025diff
changeset | 1126 | lemma is_interval_ball_real: "is_interval (ball a b)" for a b::real | 
| 
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
 immler parents: 
71025diff
changeset | 1127 | by (simp add: ball_eq_greaterThanLessThan is_interval_def) | 
| 
c2465b429e6e
Line_Segment is independent of Convex_Euclidean_Space
 immler parents: 
71025diff
changeset | 1128 | |
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1129 | |
| 70136 | 1130 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Bounded Projections\<close> | 
| 62127 | 1131 | |
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1132 | lemma bounded_inner_imp_bdd_above: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1133 | assumes "bounded s" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1134 | shows "bdd_above ((\<lambda>x. x \<bullet> a) ` s)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1135 | by (simp add: assms bounded_imp_bdd_above bounded_linear_image bounded_linear_inner_left) | 
| 33175 | 1136 | |
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1137 | lemma bounded_inner_imp_bdd_below: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1138 | assumes "bounded s" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1139 | shows "bdd_below ((\<lambda>x. x \<bullet> a) ` s)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1140 | by (simp add: assms bounded_imp_bdd_below bounded_linear_image bounded_linear_inner_left) | 
| 44632 | 1141 | |
| 53282 | 1142 | |
| 70136 | 1143 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Structural rules for pointwise continuity\<close> | 
| 33175 | 1144 | |
| 51361 
21e5b6efb317
changed continuous_intros into a named theorems collection
 hoelzl parents: 
51351diff
changeset | 1145 | lemma continuous_infnorm[continuous_intros]: | 
| 53282 | 1146 | "continuous F f \<Longrightarrow> continuous F (\<lambda>x. infnorm (f x))" | 
| 44647 
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
 huffman parents: 
44632diff
changeset | 1147 | unfolding continuous_def by (rule tendsto_infnorm) | 
| 33175 | 1148 | |
| 51361 
21e5b6efb317
changed continuous_intros into a named theorems collection
 hoelzl parents: 
51351diff
changeset | 1149 | lemma continuous_inner[continuous_intros]: | 
| 53282 | 1150 | assumes "continuous F f" | 
| 1151 | and "continuous F g" | |
| 44647 
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
 huffman parents: 
44632diff
changeset | 1152 | shows "continuous F (\<lambda>x. inner (f x) (g x))" | 
| 
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
 huffman parents: 
44632diff
changeset | 1153 | using assms unfolding continuous_def by (rule tendsto_inner) | 
| 
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
 huffman parents: 
44632diff
changeset | 1154 | |
| 69516 
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
 immler parents: 
69508diff
changeset | 1155 | |
| 70136 | 1156 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Structural rules for setwise continuity\<close> | 
| 33175 | 1157 | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56290diff
changeset | 1158 | lemma continuous_on_infnorm[continuous_intros]: | 
| 53282 | 1159 | "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. infnorm (f x))" | 
| 44647 
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
 huffman parents: 
44632diff
changeset | 1160 | unfolding continuous_on by (fast intro: tendsto_infnorm) | 
| 
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
 huffman parents: 
44632diff
changeset | 1161 | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56290diff
changeset | 1162 | lemma continuous_on_inner[continuous_intros]: | 
| 44531 
1d477a2b1572
replace some continuous_on lemmas with more general versions
 huffman parents: 
44530diff
changeset | 1163 | fixes g :: "'a::topological_space \<Rightarrow> 'b::real_inner" | 
| 53282 | 1164 | assumes "continuous_on s f" | 
| 1165 | and "continuous_on s g" | |
| 44531 
1d477a2b1572
replace some continuous_on lemmas with more general versions
 huffman parents: 
44530diff
changeset | 1166 | shows "continuous_on s (\<lambda>x. inner (f x) (g x))" | 
| 
1d477a2b1572
replace some continuous_on lemmas with more general versions
 huffman parents: 
44530diff
changeset | 1167 | using bounded_bilinear_inner assms | 
| 
1d477a2b1572
replace some continuous_on lemmas with more general versions
 huffman parents: 
44530diff
changeset | 1168 | by (rule bounded_bilinear.continuous_on) | 
| 
1d477a2b1572
replace some continuous_on lemmas with more general versions
 huffman parents: 
44530diff
changeset | 1169 | |
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1170 | |
| 70136 | 1171 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Openness of halfspaces.\<close> | 
| 33175 | 1172 | |
| 1173 | lemma open_halfspace_lt: "open {x. inner a x < b}"
 | |
| 71172 | 1174 | by (simp add: open_Collect_less continuous_on_inner) | 
| 33175 | 1175 | |
| 1176 | lemma open_halfspace_gt: "open {x. inner a x > b}"
 | |
| 71172 | 1177 | by (simp add: open_Collect_less continuous_on_inner) | 
| 33175 | 1178 | |
| 53282 | 1179 | lemma open_halfspace_component_lt: "open {x::'a::euclidean_space. x\<bullet>i < a}"
 | 
| 71172 | 1180 | by (simp add: open_Collect_less continuous_on_inner) | 
| 33175 | 1181 | |
| 53282 | 1182 | lemma open_halfspace_component_gt: "open {x::'a::euclidean_space. x\<bullet>i > a}"
 | 
| 71172 | 1183 | by (simp add: open_Collect_less continuous_on_inner) | 
| 33175 | 1184 | |
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1185 | lemma eucl_less_eq_halfspaces: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1186 | fixes a :: "'a::euclidean_space" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1187 |   shows "{x. x <e a} = (\<Inter>i\<in>Basis. {x. x \<bullet> i < a \<bullet> i})"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1188 |         "{x. a <e x} = (\<Inter>i\<in>Basis. {x. a \<bullet> i < x \<bullet> i})"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1189 | by (auto simp: eucl_less_def) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1190 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1191 | lemma open_Collect_eucl_less[simp, intro]: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1192 | fixes a :: "'a::euclidean_space" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1193 |   shows "open {x. x <e a}" "open {x. a <e x}"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1194 | by (auto simp: eucl_less_eq_halfspaces open_halfspace_component_lt open_halfspace_component_gt) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1195 | |
| 71025 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1196 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Closure and Interior of halfspaces and hyperplanes\<close> | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1197 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1198 | lemma continuous_at_inner: "continuous (at x) (inner a)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1199 | unfolding continuous_at by (intro tendsto_intros) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1200 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1201 | lemma closed_halfspace_le: "closed {x. inner a x \<le> b}"
 | 
| 71172 | 1202 | by (simp add: closed_Collect_le continuous_on_inner) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1203 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1204 | lemma closed_halfspace_ge: "closed {x. inner a x \<ge> b}"
 | 
| 71172 | 1205 | by (simp add: closed_Collect_le continuous_on_inner) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1206 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1207 | lemma closed_hyperplane: "closed {x. inner a x = b}"
 | 
| 71172 | 1208 | by (simp add: closed_Collect_eq continuous_on_inner) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1209 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1210 | lemma closed_halfspace_component_le: "closed {x::'a::euclidean_space. x\<bullet>i \<le> a}"
 | 
| 71172 | 1211 | by (simp add: closed_Collect_le continuous_on_inner) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1212 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1213 | lemma closed_halfspace_component_ge: "closed {x::'a::euclidean_space. x\<bullet>i \<ge> a}"
 | 
| 71172 | 1214 | by (simp add: closed_Collect_le continuous_on_inner) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1215 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1216 | lemma closed_interval_left: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1217 | fixes b :: "'a::euclidean_space" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1218 |   shows "closed {x::'a. \<forall>i\<in>Basis. x\<bullet>i \<le> b\<bullet>i}"
 | 
| 71172 | 1219 | by (simp add: Collect_ball_eq closed_INT closed_Collect_le continuous_on_inner) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1220 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1221 | lemma closed_interval_right: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1222 | fixes a :: "'a::euclidean_space" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1223 |   shows "closed {x::'a. \<forall>i\<in>Basis. a\<bullet>i \<le> x\<bullet>i}"
 | 
| 71172 | 1224 | by (simp add: Collect_ball_eq closed_INT closed_Collect_le continuous_on_inner) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1225 | |
| 71025 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1226 | lemma interior_halfspace_le [simp]: | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1227 | assumes "a \<noteq> 0" | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1228 |     shows "interior {x. a \<bullet> x \<le> b} = {x. a \<bullet> x < b}"
 | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1229 | proof - | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1230 |   have *: "a \<bullet> x < b" if x: "x \<in> S" and S: "S \<subseteq> {x. a \<bullet> x \<le> b}" and "open S" for S x
 | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1231 | proof - | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1232 | obtain e where "e>0" and e: "cball x e \<subseteq> S" | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1233 | using \<open>open S\<close> open_contains_cball x by blast | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1234 | then have "x + (e / norm a) *\<^sub>R a \<in> cball x e" | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1235 | by (simp add: dist_norm) | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1236 | then have "x + (e / norm a) *\<^sub>R a \<in> S" | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1237 | using e by blast | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1238 |     then have "x + (e / norm a) *\<^sub>R a \<in> {x. a \<bullet> x \<le> b}"
 | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1239 | using S by blast | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1240 | moreover have "e * (a \<bullet> a) / norm a > 0" | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1241 | by (simp add: \<open>0 < e\<close> assms) | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1242 | ultimately show ?thesis | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1243 | by (simp add: algebra_simps) | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1244 | qed | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1245 | show ?thesis | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1246 | by (rule interior_unique) (auto simp: open_halfspace_lt *) | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1247 | qed | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1248 | |
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1249 | lemma interior_halfspace_ge [simp]: | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1250 |    "a \<noteq> 0 \<Longrightarrow> interior {x. a \<bullet> x \<ge> b} = {x. a \<bullet> x > b}"
 | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1251 | using interior_halfspace_le [of "-a" "-b"] by simp | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1252 | |
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1253 | lemma closure_halfspace_lt [simp]: | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1254 | assumes "a \<noteq> 0" | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1255 |     shows "closure {x. a \<bullet> x < b} = {x. a \<bullet> x \<le> b}"
 | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1256 | proof - | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1257 |   have [simp]: "-{x. a \<bullet> x < b} = {x. a \<bullet> x \<ge> b}"
 | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1258 | by (force simp:) | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1259 | then show ?thesis | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1260 | using interior_halfspace_ge [of a b] assms | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1261 | by (force simp: closure_interior) | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1262 | qed | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1263 | |
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1264 | lemma closure_halfspace_gt [simp]: | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1265 |    "a \<noteq> 0 \<Longrightarrow> closure {x. a \<bullet> x > b} = {x. a \<bullet> x \<ge> b}"
 | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1266 | using closure_halfspace_lt [of "-a" "-b"] by simp | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1267 | |
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1268 | lemma interior_hyperplane [simp]: | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1269 | assumes "a \<noteq> 0" | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1270 |     shows "interior {x. a \<bullet> x = b} = {}"
 | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1271 | proof - | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1272 |   have [simp]: "{x. a \<bullet> x = b} = {x. a \<bullet> x \<le> b} \<inter> {x. a \<bullet> x \<ge> b}"
 | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1273 | by (force simp:) | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1274 | then show ?thesis | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1275 | by (auto simp: assms) | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1276 | qed | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1277 | |
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1278 | lemma frontier_halfspace_le: | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1279 | assumes "a \<noteq> 0 \<or> b \<noteq> 0" | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1280 |     shows "frontier {x. a \<bullet> x \<le> b} = {x. a \<bullet> x = b}"
 | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1281 | proof (cases "a = 0") | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1282 | case True with assms show ?thesis by simp | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1283 | next | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1284 | case False then show ?thesis | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1285 | by (force simp: frontier_def closed_halfspace_le) | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1286 | qed | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1287 | |
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1288 | lemma frontier_halfspace_ge: | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1289 | assumes "a \<noteq> 0 \<or> b \<noteq> 0" | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1290 |     shows "frontier {x. a \<bullet> x \<ge> b} = {x. a \<bullet> x = b}"
 | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1291 | proof (cases "a = 0") | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1292 | case True with assms show ?thesis by simp | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1293 | next | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1294 | case False then show ?thesis | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1295 | by (force simp: frontier_def closed_halfspace_ge) | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1296 | qed | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1297 | |
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1298 | lemma frontier_halfspace_lt: | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1299 | assumes "a \<noteq> 0 \<or> b \<noteq> 0" | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1300 |     shows "frontier {x. a \<bullet> x < b} = {x. a \<bullet> x = b}"
 | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1301 | proof (cases "a = 0") | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1302 | case True with assms show ?thesis by simp | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1303 | next | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1304 | case False then show ?thesis | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1305 | by (force simp: frontier_def interior_open open_halfspace_lt) | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1306 | qed | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1307 | |
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1308 | lemma frontier_halfspace_gt: | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1309 | assumes "a \<noteq> 0 \<or> b \<noteq> 0" | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1310 |     shows "frontier {x. a \<bullet> x > b} = {x. a \<bullet> x = b}"
 | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1311 | proof (cases "a = 0") | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1312 | case True with assms show ?thesis by simp | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1313 | next | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1314 | case False then show ?thesis | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1315 | by (force simp: frontier_def interior_open open_halfspace_gt) | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70952diff
changeset | 1316 | qed | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1317 | |
| 70136 | 1318 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Some more convenient intermediate-value theorem formulations\<close> | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1319 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1320 | lemma connected_ivt_hyperplane: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1321 | assumes "connected S" and xy: "x \<in> S" "y \<in> S" and b: "inner a x \<le> b" "b \<le> inner a y" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1322 | shows "\<exists>z \<in> S. inner a z = b" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1323 | proof (rule ccontr) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1324 | assume as:"\<not> (\<exists>z\<in>S. inner a z = b)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1325 |   let ?A = "{x. inner a x < b}"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1326 |   let ?B = "{x. inner a x > b}"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1327 | have "open ?A" "open ?B" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1328 | using open_halfspace_lt and open_halfspace_gt by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1329 |   moreover have "?A \<inter> ?B = {}" by auto
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1330 | moreover have "S \<subseteq> ?A \<union> ?B" using as by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1331 | ultimately show False | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1332 | using \<open>connected S\<close>[unfolded connected_def not_ex, | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1333 | THEN spec[where x="?A"], THEN spec[where x="?B"]] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1334 | using xy b by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1335 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1336 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1337 | lemma connected_ivt_component: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1338 | fixes x::"'a::euclidean_space" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1339 | shows "connected S \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> x\<bullet>k \<le> a \<Longrightarrow> a \<le> y\<bullet>k \<Longrightarrow> (\<exists>z\<in>S. z\<bullet>k = a)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1340 | using connected_ivt_hyperplane[of S x y "k::'a" a] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1341 | by (auto simp: inner_commute) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1342 | |
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1343 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1344 | subsection \<open>Limit Component Bounds\<close> | 
| 33175 | 1345 | |
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1346 | lemma Lim_component_le: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1347 | fixes f :: "'a \<Rightarrow> 'b::euclidean_space" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1348 | assumes "(f \<longlongrightarrow> l) net" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1349 | and "\<not> (trivial_limit net)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1350 | and "eventually (\<lambda>x. f(x)\<bullet>i \<le> b) net" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1351 | shows "l\<bullet>i \<le> b" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1352 | by (rule tendsto_le[OF assms(2) tendsto_const tendsto_inner[OF assms(1) tendsto_const] assms(3)]) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1353 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1354 | lemma Lim_component_ge: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1355 | fixes f :: "'a \<Rightarrow> 'b::euclidean_space" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1356 | assumes "(f \<longlongrightarrow> l) net" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1357 | and "\<not> (trivial_limit net)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1358 | and "eventually (\<lambda>x. b \<le> (f x)\<bullet>i) net" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1359 | shows "b \<le> l\<bullet>i" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1360 | by (rule tendsto_le[OF assms(2) tendsto_inner[OF assms(1) tendsto_const] tendsto_const assms(3)]) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1361 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1362 | lemma Lim_component_eq: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1363 | fixes f :: "'a \<Rightarrow> 'b::euclidean_space" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1364 | assumes net: "(f \<longlongrightarrow> l) net" "\<not> trivial_limit net" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1365 | and ev:"eventually (\<lambda>x. f(x)\<bullet>i = b) net" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1366 | shows "l\<bullet>i = b" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1367 | using ev[unfolded order_eq_iff eventually_conj_iff] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1368 | using Lim_component_ge[OF net, of b i] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1369 | using Lim_component_le[OF net, of i b] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1370 | by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 1371 | |
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1372 | lemma open_box[intro]: "open (box a b)" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1373 | proof - | 
| 67399 | 1374 |   have "open (\<Inter>i\<in>Basis. ((\<bullet>) i) -` {a \<bullet> i <..< b \<bullet> i})"
 | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62466diff
changeset | 1375 | by (auto intro!: continuous_open_vimage continuous_inner continuous_ident continuous_const) | 
| 67399 | 1376 |   also have "(\<Inter>i\<in>Basis. ((\<bullet>) i) -` {a \<bullet> i <..< b \<bullet> i}) = box a b"
 | 
| 66643 
f7e38b8583a0
Correction of typos and a bit of streamlining
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 1377 | by (auto simp: box_def inner_commute) | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1378 | finally show ?thesis . | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1379 | qed | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1380 | |
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1381 | lemma closed_cbox[intro]: | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1382 | fixes a b :: "'a::euclidean_space" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1383 | shows "closed (cbox a b)" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1384 | proof - | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1385 |   have "closed (\<Inter>i\<in>Basis. (\<lambda>x. x\<bullet>i) -` {a\<bullet>i .. b\<bullet>i})"
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1386 | by (intro closed_INT ballI continuous_closed_vimage allI | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1387 | linear_continuous_at closed_real_atLeastAtMost finite_Basis bounded_linear_inner_left) | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1388 |   also have "(\<Inter>i\<in>Basis. (\<lambda>x. x\<bullet>i) -` {a\<bullet>i .. b\<bullet>i}) = cbox a b"
 | 
| 66643 
f7e38b8583a0
Correction of typos and a bit of streamlining
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 1389 | by (auto simp: cbox_def) | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1390 | finally show "closed (cbox a b)" . | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1391 | qed | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1392 | |
| 62618 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1393 | lemma interior_cbox [simp]: | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1394 | fixes a b :: "'a::euclidean_space" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1395 | shows "interior (cbox a b) = box a b" (is "?L = ?R") | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1396 | proof(rule subset_antisym) | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1397 | show "?R \<subseteq> ?L" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1398 | using box_subset_cbox open_box | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1399 | by (rule interior_maximal) | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1400 |   {
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1401 | fix x | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1402 | assume "x \<in> interior (cbox a b)" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1403 | then obtain s where s: "open s" "x \<in> s" "s \<subseteq> cbox a b" .. | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1404 | then obtain e where "e>0" and e:"\<forall>x'. dist x' x < e \<longrightarrow> x' \<in> cbox a b" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1405 | unfolding open_dist and subset_eq by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1406 |     {
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1407 | fix i :: 'a | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1408 | assume i: "i \<in> Basis" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1409 | have "dist (x - (e / 2) *\<^sub>R i) x < e" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1410 | and "dist (x + (e / 2) *\<^sub>R i) x < e" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1411 | unfolding dist_norm | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1412 | apply auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1413 | unfolding norm_minus_cancel | 
| 60420 | 1414 | using norm_Basis[OF i] \<open>e>0\<close> | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1415 | apply auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1416 | done | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1417 | then have "a \<bullet> i \<le> (x - (e / 2) *\<^sub>R i) \<bullet> i" and "(x + (e / 2) *\<^sub>R i) \<bullet> i \<le> b \<bullet> i" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1418 | using e[THEN spec[where x="x - (e/2) *\<^sub>R i"]] | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1419 | and e[THEN spec[where x="x + (e/2) *\<^sub>R i"]] | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1420 | unfolding mem_box | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1421 | using i | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1422 | by blast+ | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1423 | then have "a \<bullet> i < x \<bullet> i" and "x \<bullet> i < b \<bullet> i" | 
| 60420 | 1424 | using \<open>e>0\<close> i | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1425 | by (auto simp: inner_diff_left inner_Basis inner_add_left) | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1426 | } | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1427 | then have "x \<in> box a b" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1428 | unfolding mem_box by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1429 | } | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1430 | then show "?L \<subseteq> ?R" .. | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1431 | qed | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1432 | |
| 63928 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63881diff
changeset | 1433 | lemma bounded_cbox [simp]: | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1434 | fixes a :: "'a::euclidean_space" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1435 | shows "bounded (cbox a b)" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1436 | proof - | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1437 | let ?b = "\<Sum>i\<in>Basis. \<bar>a\<bullet>i\<bar> + \<bar>b\<bullet>i\<bar>" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1438 |   {
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1439 | fix x :: "'a" | 
| 68120 | 1440 | assume "\<And>i. i\<in>Basis \<Longrightarrow> a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i" | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1441 | then have "(\<Sum>i\<in>Basis. \<bar>x \<bullet> i\<bar>) \<le> ?b" | 
| 68120 | 1442 | by (force simp: intro!: sum_mono) | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1443 | then have "norm x \<le> ?b" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1444 | using norm_le_l1[of x] by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1445 | } | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1446 | then show ?thesis | 
| 68120 | 1447 | unfolding cbox_def bounded_iff by force | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1448 | qed | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1449 | |
| 62618 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1450 | lemma bounded_box [simp]: | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1451 | fixes a :: "'a::euclidean_space" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1452 | shows "bounded (box a b)" | 
| 68120 | 1453 | using bounded_cbox[of a b] box_subset_cbox[of a b] bounded_subset[of "cbox a b" "box a b"] | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1454 | by simp | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1455 | |
| 62618 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1456 | lemma not_interval_UNIV [simp]: | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1457 | fixes a :: "'a::euclidean_space" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1458 | shows "cbox a b \<noteq> UNIV" "box a b \<noteq> UNIV" | 
| 62618 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1459 | using bounded_box[of a b] bounded_cbox[of a b] by force+ | 
| 
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1460 | |
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 1461 | lemma not_interval_UNIV2 [simp]: | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 1462 | fixes a :: "'a::euclidean_space" | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 1463 | shows "UNIV \<noteq> cbox a b" "UNIV \<noteq> box a b" | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 1464 | using bounded_box[of a b] bounded_cbox[of a b] by force+ | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 1465 | |
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1466 | lemma box_midpoint: | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1467 | fixes a :: "'a::euclidean_space" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1468 |   assumes "box a b \<noteq> {}"
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1469 | shows "((1/2) *\<^sub>R (a + b)) \<in> box a b" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1470 | proof - | 
| 68120 | 1471 | have "a \<bullet> i < ((1 / 2) *\<^sub>R (a + b)) \<bullet> i \<and> ((1 / 2) *\<^sub>R (a + b)) \<bullet> i < b \<bullet> i" if "i \<in> Basis" for i | 
| 1472 | using assms that by (auto simp: inner_add_left box_ne_empty) | |
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1473 | then show ?thesis unfolding mem_box by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1474 | qed | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1475 | |
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1476 | lemma open_cbox_convex: | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1477 | fixes x :: "'a::euclidean_space" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1478 | assumes x: "x \<in> box a b" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1479 | and y: "y \<in> cbox a b" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1480 | and e: "0 < e" "e \<le> 1" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1481 | shows "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \<in> box a b" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1482 | proof - | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1483 |   {
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1484 | fix i :: 'a | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1485 | assume i: "i \<in> Basis" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1486 | have "a \<bullet> i = e * (a \<bullet> i) + (1 - e) * (a \<bullet> i)" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1487 | unfolding left_diff_distrib by simp | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1488 | also have "\<dots> < e * (x \<bullet> i) + (1 - e) * (y \<bullet> i)" | 
| 68120 | 1489 | proof (rule add_less_le_mono) | 
| 1490 | show "e * (a \<bullet> i) < e * (x \<bullet> i)" | |
| 1491 | using \<open>0 < e\<close> i mem_box(1) x by auto | |
| 1492 | show "(1 - e) * (a \<bullet> i) \<le> (1 - e) * (y \<bullet> i)" | |
| 1493 | by (meson diff_ge_0_iff_ge \<open>e \<le> 1\<close> i mem_box(2) mult_left_mono y) | |
| 1494 | qed | |
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1495 | finally have "a \<bullet> i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1496 | unfolding inner_simps by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1497 | moreover | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1498 |     {
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1499 | have "b \<bullet> i = e * (b\<bullet>i) + (1 - e) * (b\<bullet>i)" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1500 | unfolding left_diff_distrib by simp | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1501 | also have "\<dots> > e * (x \<bullet> i) + (1 - e) * (y \<bullet> i)" | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: 
66794diff
changeset | 1502 | proof (rule add_less_le_mono) | 
| 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: 
66794diff
changeset | 1503 | show "e * (x \<bullet> i) < e * (b \<bullet> i)" | 
| 68120 | 1504 | using \<open>0 < e\<close> i mem_box(1) x by auto | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: 
66794diff
changeset | 1505 | show "(1 - e) * (y \<bullet> i) \<le> (1 - e) * (b \<bullet> i)" | 
| 68120 | 1506 | by (meson diff_ge_0_iff_ge \<open>e \<le> 1\<close> i mem_box(2) mult_left_mono y) | 
| 66827 
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
 paulson <lp15@cam.ac.uk> parents: 
66794diff
changeset | 1507 | qed | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1508 | finally have "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i < b \<bullet> i" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1509 | unfolding inner_simps by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1510 | } | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1511 | ultimately have "a \<bullet> i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i \<and> (e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i < b \<bullet> i" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1512 | by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1513 | } | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1514 | then show ?thesis | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1515 | unfolding mem_box by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1516 | qed | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1517 | |
| 63881 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1518 | lemma closure_cbox [simp]: "closure (cbox a b) = cbox a b" | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1519 | by (simp add: closed_cbox) | 
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1520 | |
| 
b746b19197bd
lots of new results about topology, affine dimension etc
 paulson <lp15@cam.ac.uk> parents: 
63627diff
changeset | 1521 | lemma closure_box [simp]: | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1522 | fixes a :: "'a::euclidean_space" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1523 |    assumes "box a b \<noteq> {}"
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1524 | shows "closure (box a b) = cbox a b" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1525 | proof - | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1526 | have ab: "a <e b" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1527 | using assms by (simp add: eucl_less_def box_ne_empty) | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1528 | let ?c = "(1 / 2) *\<^sub>R (a + b)" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1529 |   {
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1530 | fix x | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1531 | assume as:"x \<in> cbox a b" | 
| 63040 | 1532 | define f where [abs_def]: "f n = x + (inverse (real n + 1)) *\<^sub>R (?c - x)" for n | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1533 |     {
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1534 | fix n | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1535 | assume fn: "f n <e b \<longrightarrow> a <e f n \<longrightarrow> f n = x" and xc: "x \<noteq> ?c" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1536 | have *: "0 < inverse (real n + 1)" "inverse (real n + 1) \<le> 1" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1537 | unfolding inverse_le_1_iff by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1538 | have "(inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b)) + (1 - inverse (real n + 1)) *\<^sub>R x = | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1539 | x + (inverse (real n + 1)) *\<^sub>R (((1 / 2) *\<^sub>R (a + b)) - x)" | 
| 66643 
f7e38b8583a0
Correction of typos and a bit of streamlining
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 1540 | by (auto simp: algebra_simps) | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1541 | then have "f n <e b" and "a <e f n" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1542 | using open_cbox_convex[OF box_midpoint[OF assms] as *] | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1543 | unfolding f_def by (auto simp: box_def eucl_less_def) | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1544 | then have False | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1545 | using fn unfolding f_def using xc by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1546 | } | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1547 | moreover | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1548 |     {
 | 
| 61973 | 1549 | assume "\<not> (f \<longlongrightarrow> x) sequentially" | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1550 |       {
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1551 | fix e :: real | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1552 | assume "e > 0" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61552diff
changeset | 1553 | then obtain N :: nat where N: "inverse (real (N + 1)) < e" | 
| 68120 | 1554 | using reals_Archimedean by auto | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61552diff
changeset | 1555 | have "inverse (real n + 1) < e" if "N \<le> n" for n | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61552diff
changeset | 1556 | by (auto intro!: that le_less_trans [OF _ N]) | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1557 | then have "\<exists>N::nat. \<forall>n\<ge>N. inverse (real n + 1) < e" by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1558 | } | 
| 61973 | 1559 | then have "((\<lambda>n. inverse (real n + 1)) \<longlongrightarrow> 0) sequentially" | 
| 66643 
f7e38b8583a0
Correction of typos and a bit of streamlining
 paulson <lp15@cam.ac.uk> parents: 
66641diff
changeset | 1560 | unfolding lim_sequentially by(auto simp: dist_norm) | 
| 61973 | 1561 | then have "(f \<longlongrightarrow> x) sequentially" | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1562 | unfolding f_def | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1563 | using tendsto_add[OF tendsto_const, of "\<lambda>n::nat. (inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x)" 0 sequentially x] | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1564 | using tendsto_scaleR [OF _ tendsto_const, of "\<lambda>n::nat. inverse (real n + 1)" 0 sequentially "((1 / 2) *\<^sub>R (a + b) - x)"] | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1565 | by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1566 | } | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1567 | ultimately have "x \<in> closure (box a b)" | 
| 68120 | 1568 | using as box_midpoint[OF assms] | 
| 1569 | unfolding closure_def islimpt_sequential | |
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1570 | by (cases "x=?c") (auto simp: in_box_eucl_less) | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1571 | } | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1572 | then show ?thesis | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1573 | using closure_minimal[OF box_subset_cbox, of a b] by blast | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1574 | qed | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1575 | |
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1576 | lemma bounded_subset_box_symmetric: | 
| 68120 | 1577 |   fixes S :: "('a::euclidean_space) set"
 | 
| 1578 | assumes "bounded S" | |
| 1579 | obtains a where "S \<subseteq> box (-a) a" | |
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1580 | proof - | 
| 68120 | 1581 | obtain b where "b>0" and b: "\<forall>x\<in>S. norm x \<le> b" | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1582 | using assms[unfolded bounded_pos] by auto | 
| 63040 | 1583 | define a :: 'a where "a = (\<Sum>i\<in>Basis. (b + 1) *\<^sub>R i)" | 
| 68120 | 1584 | have "(-a)\<bullet>i < x\<bullet>i" and "x\<bullet>i < a\<bullet>i" if "x \<in> S" and i: "i \<in> Basis" for x i | 
| 1585 | using b Basis_le_norm[OF i, of x] that by (auto simp: a_def) | |
| 1586 | then have "S \<subseteq> box (-a) a" | |
| 1587 | by (auto simp: simp add: box_def) | |
| 1588 | then show ?thesis .. | |
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1589 | qed | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1590 | |
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1591 | lemma bounded_subset_cbox_symmetric: | 
| 68120 | 1592 |   fixes S :: "('a::euclidean_space) set"
 | 
| 1593 | assumes "bounded S" | |
| 1594 | obtains a where "S \<subseteq> cbox (-a) a" | |
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1595 | proof - | 
| 68120 | 1596 | obtain a where "S \<subseteq> box (-a) a" | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1597 | using bounded_subset_box_symmetric[OF assms] by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1598 | then show ?thesis | 
| 68120 | 1599 | by (meson box_subset_cbox dual_order.trans that) | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1600 | qed | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1601 | |
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1602 | lemma frontier_cbox: | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1603 | fixes a b :: "'a::euclidean_space" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1604 | shows "frontier (cbox a b) = cbox a b - box a b" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1605 | unfolding frontier_def unfolding interior_cbox and closure_closed[OF closed_cbox] .. | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1606 | |
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1607 | lemma frontier_box: | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1608 | fixes a b :: "'a::euclidean_space" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1609 |   shows "frontier (box a b) = (if box a b = {} then {} else cbox a b - box a b)"
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1610 | proof (cases "box a b = {}")
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1611 | case True | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1612 | then show ?thesis | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1613 | using frontier_empty by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1614 | next | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1615 | case False | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1616 | then show ?thesis | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1617 | unfolding frontier_def and closure_box[OF False] and interior_open[OF open_box] | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1618 | by auto | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1619 | qed | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1620 | |
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66827diff
changeset | 1621 | lemma Int_interval_mixed_eq_empty: | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1622 | fixes a :: "'a::euclidean_space" | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1623 |    assumes "box c d \<noteq> {}"
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1624 |   shows "box a b \<inter> cbox c d = {} \<longleftrightarrow> box a b \<inter> box c d = {}"
 | 
| 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1625 | unfolding closure_box[OF assms, symmetric] | 
| 62843 
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 1626 | unfolding open_Int_closure_eq_empty[OF open_box] .. | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 1627 | |
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1628 | subsection \<open>Class Instances\<close> | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1629 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1630 | lemma compact_lemma: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1631 | fixes f :: "nat \<Rightarrow> 'a::euclidean_space" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1632 | assumes "bounded (range f)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1633 | shows "\<forall>d\<subseteq>Basis. \<exists>l::'a. \<exists> r. | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1634 | strict_mono r \<and> (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1635 | by (rule compact_lemma_general[where unproj="\<lambda>e. \<Sum>i\<in>Basis. e i *\<^sub>R i"]) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1636 | (auto intro!: assms bounded_linear_inner_left bounded_linear_image | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1637 | simp: euclidean_representation) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1638 | |
| 70136 | 1639 | instance\<^marker>\<open>tag important\<close> euclidean_space \<subseteq> heine_borel | 
| 1640 | proof | |
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1641 | fix f :: "nat \<Rightarrow> 'a" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1642 | assume f: "bounded (range f)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1643 | then obtain l::'a and r where r: "strict_mono r" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1644 | and l: "\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1645 | using compact_lemma [OF f] by blast | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1646 |   {
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1647 | fix e::real | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1648 | assume "e > 0" | 
| 71172 | 1649 |     hence "e / real_of_nat DIM('a) > 0" by (simp)
 | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1650 |     with l have "eventually (\<lambda>n. \<forall>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i) < e / (real_of_nat DIM('a))) sequentially"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1651 | by simp | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1652 | moreover | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1653 |     {
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1654 | fix n | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1655 |       assume n: "\<forall>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i) < e / (real_of_nat DIM('a))"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1656 | have "dist (f (r n)) l \<le> (\<Sum>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i))" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1657 | apply (subst euclidean_dist_l2) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1658 | using zero_le_dist | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1659 | apply (rule L2_set_le_sum) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1660 | done | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1661 |       also have "\<dots> < (\<Sum>i\<in>(Basis::'a set). e / (real_of_nat DIM('a)))"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1662 | apply (rule sum_strict_mono) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1663 | using n | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1664 | apply auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1665 | done | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1666 | finally have "dist (f (r n)) l < e" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1667 | by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1668 | } | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1669 | ultimately have "eventually (\<lambda>n. dist (f (r n)) l < e) sequentially" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1670 | by (rule eventually_mono) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1671 | } | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1672 | then have *: "((f \<circ> r) \<longlongrightarrow> l) sequentially" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1673 | unfolding o_def tendsto_iff by simp | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1674 | with r show "\<exists>l r. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1675 | by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1676 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1677 | |
| 70136 | 1678 | instance\<^marker>\<open>tag important\<close> euclidean_space \<subseteq> banach .. | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1679 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1680 | instance euclidean_space \<subseteq> second_countable_topology | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1681 | proof | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1682 | define a where "a f = (\<Sum>i\<in>Basis. fst (f i) *\<^sub>R i)" for f :: "'a \<Rightarrow> real \<times> real" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1683 | then have a: "\<And>f. (\<Sum>i\<in>Basis. fst (f i) *\<^sub>R i) = a f" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1684 | by simp | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1685 | define b where "b f = (\<Sum>i\<in>Basis. snd (f i) *\<^sub>R i)" for f :: "'a \<Rightarrow> real \<times> real" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1686 | then have b: "\<And>f. (\<Sum>i\<in>Basis. snd (f i) *\<^sub>R i) = b f" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1687 | by simp | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1688 | define B where "B = (\<lambda>f. box (a f) (b f)) ` (Basis \<rightarrow>\<^sub>E (\<rat> \<times> \<rat>))" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1689 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1690 | have "Ball B open" by (simp add: B_def open_box) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1691 | moreover have "(\<forall>A. open A \<longrightarrow> (\<exists>B'\<subseteq>B. \<Union>B' = A))" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1692 | proof safe | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1693 | fix A::"'a set" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1694 | assume "open A" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1695 | show "\<exists>B'\<subseteq>B. \<Union>B' = A" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1696 |       apply (rule exI[of _ "{b\<in>B. b \<subseteq> A}"])
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1697 | apply (subst (3) open_UNION_box[OF \<open>open A\<close>]) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1698 | apply (auto simp: a b B_def) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1699 | done | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1700 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1701 | ultimately | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1702 | have "topological_basis B" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1703 | unfolding topological_basis_def by blast | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1704 | moreover | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1705 | have "countable B" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1706 | unfolding B_def | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1707 | by (intro countable_image countable_PiE finite_Basis countable_SIGMA countable_rat) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1708 | ultimately show "\<exists>B::'a set set. countable B \<and> open = generate_topology B" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1709 | by (blast intro: topological_basis_imp_subbasis) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1710 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1711 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1712 | instance euclidean_space \<subseteq> polish_space .. | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1713 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1714 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1715 | subsection \<open>Compact Boxes\<close> | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1716 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1717 | lemma compact_cbox [simp]: | 
| 61076 | 1718 | fixes a :: "'a::euclidean_space" | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1719 | shows "compact (cbox a b)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1720 | using bounded_closed_imp_seq_compact[of "cbox a b"] using bounded_cbox[of a b] | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1721 | by (auto simp: compact_eq_seq_compact_metric) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1722 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1723 | proposition is_interval_compact: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1724 | "is_interval S \<and> compact S \<longleftrightarrow> (\<exists>a b. S = cbox a b)" (is "?lhs = ?rhs") | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1725 | proof (cases "S = {}")
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1726 | case True | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1727 | with empty_as_interval show ?thesis by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1728 | next | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1729 | case False | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1730 | show ?thesis | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1731 | proof | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1732 | assume L: ?lhs | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1733 | then have "is_interval S" "compact S" by auto | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1734 | define a where "a \<equiv> \<Sum>i\<in>Basis. (INF x\<in>S. x \<bullet> i) *\<^sub>R i" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1735 | define b where "b \<equiv> \<Sum>i\<in>Basis. (SUP x\<in>S. x \<bullet> i) *\<^sub>R i" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1736 | have 1: "\<And>x i. \<lbrakk>x \<in> S; i \<in> Basis\<rbrakk> \<Longrightarrow> (INF x\<in>S. x \<bullet> i) \<le> x \<bullet> i" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1737 | by (simp add: cInf_lower bounded_inner_imp_bdd_below compact_imp_bounded L) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1738 | have 2: "\<And>x i. \<lbrakk>x \<in> S; i \<in> Basis\<rbrakk> \<Longrightarrow> x \<bullet> i \<le> (SUP x\<in>S. x \<bullet> i)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1739 | by (simp add: cSup_upper bounded_inner_imp_bdd_above compact_imp_bounded L) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1740 | have 3: "x \<in> S" if inf: "\<And>i. i \<in> Basis \<Longrightarrow> (INF x\<in>S. x \<bullet> i) \<le> x \<bullet> i" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1741 | and sup: "\<And>i. i \<in> Basis \<Longrightarrow> x \<bullet> i \<le> (SUP x\<in>S. x \<bullet> i)" for x | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1742 | proof (rule mem_box_componentwiseI [OF \<open>is_interval S\<close>]) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1743 | fix i::'a | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1744 | assume i: "i \<in> Basis" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1745 | have cont: "continuous_on S (\<lambda>x. x \<bullet> i)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1746 | by (intro continuous_intros) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1747 | obtain a where "a \<in> S" and a: "\<And>y. y\<in>S \<Longrightarrow> a \<bullet> i \<le> y \<bullet> i" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1748 | using continuous_attains_inf [OF \<open>compact S\<close> False cont] by blast | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1749 | obtain b where "b \<in> S" and b: "\<And>y. y\<in>S \<Longrightarrow> y \<bullet> i \<le> b \<bullet> i" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1750 | using continuous_attains_sup [OF \<open>compact S\<close> False cont] by blast | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1751 | have "a \<bullet> i \<le> (INF x\<in>S. x \<bullet> i)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1752 | by (simp add: False a cINF_greatest) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1753 | also have "\<dots> \<le> x \<bullet> i" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1754 | by (simp add: i inf) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1755 | finally have ai: "a \<bullet> i \<le> x \<bullet> i" . | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1756 | have "x \<bullet> i \<le> (SUP x\<in>S. x \<bullet> i)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1757 | by (simp add: i sup) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1758 | also have "(SUP x\<in>S. x \<bullet> i) \<le> b \<bullet> i" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1759 | by (simp add: False b cSUP_least) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1760 | finally have bi: "x \<bullet> i \<le> b \<bullet> i" . | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1761 | show "x \<bullet> i \<in> (\<lambda>x. x \<bullet> i) ` S" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1762 | apply (rule_tac x="\<Sum>j\<in>Basis. (if j = i then x \<bullet> i else a \<bullet> j) *\<^sub>R j" in image_eqI) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1763 | apply (simp add: i) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1764 | apply (rule mem_is_intervalI [OF \<open>is_interval S\<close> \<open>a \<in> S\<close> \<open>b \<in> S\<close>]) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1765 | using i ai bi apply force | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1766 | done | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1767 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1768 | have "S = cbox a b" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1769 | by (auto simp: a_def b_def mem_box intro: 1 2 3) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1770 | then show ?rhs | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1771 | by blast | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1772 | next | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1773 | assume R: ?rhs | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1774 | then show ?lhs | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1775 | using compact_cbox is_interval_cbox by blast | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1776 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1777 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1778 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 1779 | |
| 70136 | 1780 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Componentwise limits and continuity\<close> | 
| 69615 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1781 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1782 | text\<open>But is the premise really necessary? Need to generalise @{thm euclidean_dist_l2}\<close>
 | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1783 | lemma Euclidean_dist_upper: "i \<in> Basis \<Longrightarrow> dist (x \<bullet> i) (y \<bullet> i) \<le> dist x y" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1784 | by (metis (no_types) member_le_L2_set euclidean_dist_l2 finite_Basis) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1785 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1786 | text\<open>But is the premise \<^term>\<open>i \<in> Basis\<close> really necessary?\<close> | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1787 | lemma open_preimage_inner: | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1788 | assumes "open S" "i \<in> Basis" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1789 |     shows "open {x. x \<bullet> i \<in> S}"
 | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1790 | proof (rule openI, simp) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1791 | fix x | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1792 | assume x: "x \<bullet> i \<in> S" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1793 | with assms obtain e where "0 < e" and e: "ball (x \<bullet> i) e \<subseteq> S" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1794 | by (auto simp: open_contains_ball_eq) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1795 | have "\<exists>e>0. ball (y \<bullet> i) e \<subseteq> S" if dxy: "dist x y < e / 2" for y | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1796 | proof (intro exI conjI) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1797 | have "dist (x \<bullet> i) (y \<bullet> i) < e / 2" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1798 | by (meson \<open>i \<in> Basis\<close> dual_order.trans Euclidean_dist_upper not_le that) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1799 | then have "dist (x \<bullet> i) z < e" if "dist (y \<bullet> i) z < e / 2" for z | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1800 | by (metis dist_commute dist_triangle_half_l that) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1801 | then have "ball (y \<bullet> i) (e / 2) \<subseteq> ball (x \<bullet> i) e" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1802 | using mem_ball by blast | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1803 | with e show "ball (y \<bullet> i) (e / 2) \<subseteq> S" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1804 | by (metis order_trans) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1805 | qed (simp add: \<open>0 < e\<close>) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1806 |   then show "\<exists>e>0. ball x e \<subseteq> {s. s \<bullet> i \<in> S}"
 | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1807 | by (metis (no_types, lifting) \<open>0 < e\<close> \<open>open S\<close> half_gt_zero_iff mem_Collect_eq mem_ball open_contains_ball_eq subsetI) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1808 | qed | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1809 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1810 | proposition tendsto_componentwise_iff: | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1811 | fixes f :: "_ \<Rightarrow> 'b::euclidean_space" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1812 | shows "(f \<longlongrightarrow> l) F \<longleftrightarrow> (\<forall>i \<in> Basis. ((\<lambda>x. (f x \<bullet> i)) \<longlongrightarrow> (l \<bullet> i)) F)" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1813 | (is "?lhs = ?rhs") | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1814 | proof | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1815 | assume ?lhs | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1816 | then show ?rhs | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1817 | unfolding tendsto_def | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1818 | apply clarify | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1819 |     apply (drule_tac x="{s. s \<bullet> i \<in> S}" in spec)
 | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1820 | apply (auto simp: open_preimage_inner) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1821 | done | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1822 | next | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1823 | assume R: ?rhs | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1824 | then have "\<And>e. e > 0 \<Longrightarrow> \<forall>i\<in>Basis. \<forall>\<^sub>F x in F. dist (f x \<bullet> i) (l \<bullet> i) < e" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1825 | unfolding tendsto_iff by blast | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1826 | then have R': "\<And>e. e > 0 \<Longrightarrow> \<forall>\<^sub>F x in F. \<forall>i\<in>Basis. dist (f x \<bullet> i) (l \<bullet> i) < e" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1827 | by (simp add: eventually_ball_finite_distrib [symmetric]) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1828 | show ?lhs | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1829 | unfolding tendsto_iff | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1830 | proof clarify | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1831 | fix e::real | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1832 | assume "0 < e" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1833 | have *: "L2_set (\<lambda>i. dist (f x \<bullet> i) (l \<bullet> i)) Basis < e" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1834 |              if "\<forall>i\<in>Basis. dist (f x \<bullet> i) (l \<bullet> i) < e / real DIM('b)" for x
 | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1835 | proof - | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1836 | have "L2_set (\<lambda>i. dist (f x \<bullet> i) (l \<bullet> i)) Basis \<le> sum (\<lambda>i. dist (f x \<bullet> i) (l \<bullet> i)) Basis" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1837 | by (simp add: L2_set_le_sum) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1838 |       also have "... < DIM('b) * (e / real DIM('b))"
 | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1839 | apply (rule sum_bounded_above_strict) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1840 | using that by auto | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1841 | also have "... = e" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1842 | by (simp add: field_simps) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1843 | finally show "L2_set (\<lambda>i. dist (f x \<bullet> i) (l \<bullet> i)) Basis < e" . | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1844 | qed | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1845 |     have "\<forall>\<^sub>F x in F. \<forall>i\<in>Basis. dist (f x \<bullet> i) (l \<bullet> i) < e / DIM('b)"
 | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1846 | apply (rule R') | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1847 | using \<open>0 < e\<close> by simp | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1848 | then show "\<forall>\<^sub>F x in F. dist (f x) l < e" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1849 | apply (rule eventually_mono) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1850 | apply (subst euclidean_dist_l2) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1851 | using * by blast | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1852 | qed | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1853 | qed | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1854 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1855 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1856 | corollary continuous_componentwise: | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1857 | "continuous F f \<longleftrightarrow> (\<forall>i \<in> Basis. continuous F (\<lambda>x. (f x \<bullet> i)))" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1858 | by (simp add: continuous_def tendsto_componentwise_iff [symmetric]) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1859 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1860 | corollary continuous_on_componentwise: | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1861 | fixes S :: "'a :: t2_space set" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1862 | shows "continuous_on S f \<longleftrightarrow> (\<forall>i \<in> Basis. continuous_on S (\<lambda>x. (f x \<bullet> i)))" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1863 | apply (simp add: continuous_on_eq_continuous_within) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1864 | using continuous_componentwise by blast | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1865 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1866 | lemma linear_componentwise_iff: | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1867 | "(linear f') \<longleftrightarrow> (\<forall>i\<in>Basis. linear (\<lambda>x. f' x \<bullet> i))" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1868 | apply (auto simp: linear_iff inner_left_distrib) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1869 | apply (metis inner_left_distrib euclidean_eq_iff) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1870 | by (metis euclidean_eqI inner_scaleR_left) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1871 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1872 | lemma bounded_linear_componentwise_iff: | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1873 | "(bounded_linear f') \<longleftrightarrow> (\<forall>i\<in>Basis. bounded_linear (\<lambda>x. f' x \<bullet> i))" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1874 | (is "?lhs = ?rhs") | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1875 | proof | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1876 | assume ?lhs then show ?rhs | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1877 | by (simp add: bounded_linear_inner_left_comp) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1878 | next | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1879 | assume ?rhs | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1880 | then have "(\<forall>i\<in>Basis. \<exists>K. \<forall>x. \<bar>f' x \<bullet> i\<bar> \<le> norm x * K)" "linear f'" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1881 | by (auto simp: bounded_linear_def bounded_linear_axioms_def linear_componentwise_iff [symmetric] ball_conj_distrib) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1882 | then obtain F where F: "\<And>i x. i \<in> Basis \<Longrightarrow> \<bar>f' x \<bullet> i\<bar> \<le> norm x * F i" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1883 | by metis | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1884 | have "norm (f' x) \<le> norm x * sum F Basis" for x | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1885 | proof - | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1886 | have "norm (f' x) \<le> (\<Sum>i\<in>Basis. \<bar>f' x \<bullet> i\<bar>)" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1887 | by (rule norm_le_l1) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1888 | also have "... \<le> (\<Sum>i\<in>Basis. norm x * F i)" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1889 | by (metis F sum_mono) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1890 | also have "... = norm x * sum F Basis" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1891 | by (simp add: sum_distrib_left) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1892 | finally show ?thesis . | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1893 | qed | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1894 | then show ?lhs | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1895 | by (force simp: bounded_linear_def bounded_linear_axioms_def \<open>linear f'\<close>) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1896 | qed | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1897 | |
| 70136 | 1898 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Continuous Extension\<close> | 
| 69615 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1899 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1900 | definition clamp :: "'a::euclidean_space \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" where | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1901 | "clamp a b x = (if (\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1902 | then (\<Sum>i\<in>Basis. (if x\<bullet>i < a\<bullet>i then a\<bullet>i else if x\<bullet>i \<le> b\<bullet>i then x\<bullet>i else b\<bullet>i) *\<^sub>R i) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1903 | else a)" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1904 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1905 | lemma clamp_in_interval[simp]: | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1906 | assumes "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1907 | shows "clamp a b x \<in> cbox a b" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1908 | unfolding clamp_def | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1909 | using box_ne_empty(1)[of a b] assms by (auto simp: cbox_def) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1910 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1911 | lemma clamp_cancel_cbox[simp]: | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1912 | fixes x a b :: "'a::euclidean_space" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1913 | assumes x: "x \<in> cbox a b" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1914 | shows "clamp a b x = x" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1915 | using assms | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1916 | by (auto simp: clamp_def mem_box intro!: euclidean_eqI[where 'a='a]) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1917 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1918 | lemma clamp_empty_interval: | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1919 | assumes "i \<in> Basis" "a \<bullet> i > b \<bullet> i" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1920 | shows "clamp a b = (\<lambda>_. a)" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1921 | using assms | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1922 | by (force simp: clamp_def[abs_def] split: if_splits intro!: ext) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1923 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1924 | lemma dist_clamps_le_dist_args: | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1925 | fixes x :: "'a::euclidean_space" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1926 | shows "dist (clamp a b y) (clamp a b x) \<le> dist y x" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1927 | proof cases | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1928 | assume le: "(\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i)" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1929 | then have "(\<Sum>i\<in>Basis. (dist (clamp a b y \<bullet> i) (clamp a b x \<bullet> i))\<^sup>2) \<le> | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1930 | (\<Sum>i\<in>Basis. (dist (y \<bullet> i) (x \<bullet> i))\<^sup>2)" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1931 | by (auto intro!: sum_mono simp: clamp_def dist_real_def abs_le_square_iff[symmetric]) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1932 | then show ?thesis | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1933 | by (auto intro: real_sqrt_le_mono | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1934 | simp: euclidean_dist_l2[where y=x] euclidean_dist_l2[where y="clamp a b x"] L2_set_def) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1935 | qed (auto simp: clamp_def) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1936 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1937 | lemma clamp_continuous_at: | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1938 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::metric_space" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1939 | and x :: 'a | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1940 | assumes f_cont: "continuous_on (cbox a b) f" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1941 | shows "continuous (at x) (\<lambda>x. f (clamp a b x))" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1942 | proof cases | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1943 | assume le: "(\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i)" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1944 | show ?thesis | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1945 | unfolding continuous_at_eps_delta | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1946 | proof safe | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1947 | fix x :: 'a | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1948 | fix e :: real | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1949 | assume "e > 0" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1950 | moreover have "clamp a b x \<in> cbox a b" | 
| 71172 | 1951 | by (simp add: le) | 
| 69615 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1952 | moreover note f_cont[simplified continuous_on_iff] | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1953 | ultimately | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1954 | obtain d where d: "0 < d" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1955 | "\<And>x'. x' \<in> cbox a b \<Longrightarrow> dist x' (clamp a b x) < d \<Longrightarrow> dist (f x') (f (clamp a b x)) < e" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1956 | by force | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1957 | show "\<exists>d>0. \<forall>x'. dist x' x < d \<longrightarrow> | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1958 | dist (f (clamp a b x')) (f (clamp a b x)) < e" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1959 | using le | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1960 | by (auto intro!: d clamp_in_interval dist_clamps_le_dist_args[THEN le_less_trans]) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1961 | qed | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1962 | qed (auto simp: clamp_empty_interval) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1963 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1964 | lemma clamp_continuous_on: | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1965 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::metric_space" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1966 | assumes f_cont: "continuous_on (cbox a b) f" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1967 | shows "continuous_on S (\<lambda>x. f (clamp a b x))" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1968 | using assms | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1969 | by (auto intro: continuous_at_imp_continuous_on clamp_continuous_at) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1970 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1971 | lemma clamp_bounded: | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1972 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::metric_space" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1973 | assumes bounded: "bounded (f ` (cbox a b))" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1974 | shows "bounded (range (\<lambda>x. f (clamp a b x)))" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1975 | proof cases | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1976 | assume le: "(\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i)" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1977 | from bounded obtain c where f_bound: "\<forall>x\<in>f ` cbox a b. dist undefined x \<le> c" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1978 | by (auto simp: bounded_any_center[where a=undefined]) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1979 | then show ?thesis | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1980 | by (auto intro!: exI[where x=c] clamp_in_interval[OF le[rule_format]] | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1981 | simp: bounded_any_center[where a=undefined]) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1982 | qed (auto simp: clamp_empty_interval image_def) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1983 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1984 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1985 | definition ext_cont :: "('a::euclidean_space \<Rightarrow> 'b::metric_space) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b"
 | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1986 | where "ext_cont f a b = (\<lambda>x. f (clamp a b x))" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1987 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1988 | lemma ext_cont_cancel_cbox[simp]: | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1989 | fixes x a b :: "'a::euclidean_space" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1990 | assumes x: "x \<in> cbox a b" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1991 | shows "ext_cont f a b x = f x" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1992 | using assms | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1993 | unfolding ext_cont_def | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1994 | by (auto simp: clamp_def mem_box intro!: euclidean_eqI[where 'a='a] arg_cong[where f=f]) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1995 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1996 | lemma continuous_on_ext_cont[continuous_intros]: | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1997 | "continuous_on (cbox a b) f \<Longrightarrow> continuous_on S (ext_cont f a b)" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1998 | by (auto intro!: clamp_continuous_on simp: ext_cont_def) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 1999 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2000 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2001 | subsection \<open>Separability\<close> | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2002 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2003 | lemma univ_second_countable_sequence: | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2004 | obtains B :: "nat \<Rightarrow> 'a::euclidean_space set" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2005 |     where "inj B" "\<And>n. open(B n)" "\<And>S. open S \<Longrightarrow> \<exists>k. S = \<Union>{B n |n. n \<in> k}"
 | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2006 | proof - | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2007 | obtain \<B> :: "'a set set" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2008 | where "countable \<B>" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2009 | and opn: "\<And>C. C \<in> \<B> \<Longrightarrow> open C" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2010 | and Un: "\<And>S. open S \<Longrightarrow> \<exists>U. U \<subseteq> \<B> \<and> S = \<Union>U" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2011 | using univ_second_countable by blast | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2012 | have *: "infinite (range (\<lambda>n. ball (0::'a) (inverse(Suc n))))" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2013 | apply (rule Infinite_Set.range_inj_infinite) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2014 | apply (simp add: inj_on_def ball_eq_ball_iff) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2015 | done | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2016 | have "infinite \<B>" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2017 | proof | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2018 | assume "finite \<B>" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2019 | then have "finite (Union ` (Pow \<B>))" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2020 | by simp | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2021 | then have "finite (range (\<lambda>n. ball (0::'a) (inverse(Suc n))))" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2022 | apply (rule rev_finite_subset) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2023 | by (metis (no_types, lifting) PowI image_eqI image_subset_iff Un [OF open_ball]) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2024 | with * show False by simp | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2025 | qed | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2026 | obtain f :: "nat \<Rightarrow> 'a set" where "\<B> = range f" "inj f" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2027 | by (blast intro: countable_as_injective_image [OF \<open>countable \<B>\<close> \<open>infinite \<B>\<close>]) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2028 |   have *: "\<exists>k. S = \<Union>{f n |n. n \<in> k}" if "open S" for S
 | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2029 | using Un [OF that] | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2030 | apply clarify | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2031 | apply (rule_tac x="f-`U" in exI) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2032 | using \<open>inj f\<close> \<open>\<B> = range f\<close> apply force | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2033 | done | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2034 | show ?thesis | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2035 | apply (rule that [OF \<open>inj f\<close> _ *]) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2036 | apply (auto simp: \<open>\<B> = range f\<close> opn) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2037 | done | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2038 | qed | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2039 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2040 | proposition separable: | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2041 |   fixes S :: "'a::{metric_space, second_countable_topology} set"
 | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2042 | obtains T where "countable T" "T \<subseteq> S" "S \<subseteq> closure T" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2043 | proof - | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2044 | obtain \<B> :: "'a set set" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2045 | where "countable \<B>" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2046 |       and "{} \<notin> \<B>"
 | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2047 | and ope: "\<And>C. C \<in> \<B> \<Longrightarrow> openin(top_of_set S) C" | 
| 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2048 | and if_ope: "\<And>T. openin(top_of_set S) T \<Longrightarrow> \<exists>\<U>. \<U> \<subseteq> \<B> \<and> T = \<Union>\<U>" | 
| 69615 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2049 | by (meson subset_second_countable) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2050 | then obtain f where f: "\<And>C. C \<in> \<B> \<Longrightarrow> f C \<in> C" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2051 | by (metis equals0I) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2052 | show ?thesis | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2053 | proof | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2054 | show "countable (f ` \<B>)" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2055 | by (simp add: \<open>countable \<B>\<close>) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2056 | show "f ` \<B> \<subseteq> S" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2057 | using ope f openin_imp_subset by blast | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2058 | show "S \<subseteq> closure (f ` \<B>)" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2059 | proof (clarsimp simp: closure_approachable) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2060 | fix x and e::real | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2061 | assume "x \<in> S" "0 < e" | 
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2062 | have "openin (top_of_set S) (S \<inter> ball x e)" | 
| 69615 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2063 | by (simp add: openin_Int_open) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2064 | with if_ope obtain \<U> where \<U>: "\<U> \<subseteq> \<B>" "S \<inter> ball x e = \<Union>\<U>" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2065 | by meson | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2066 | show "\<exists>C \<in> \<B>. dist (f C) x < e" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2067 |       proof (cases "\<U> = {}")
 | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2068 | case True | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2069 | then show ?thesis | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2070 | using \<open>0 < e\<close> \<U> \<open>x \<in> S\<close> by auto | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2071 | next | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2072 | case False | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2073 | then obtain C where "C \<in> \<U>" by blast | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2074 | show ?thesis | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2075 | proof | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2076 | show "dist (f C) x < e" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2077 | by (metis Int_iff Union_iff \<U> \<open>C \<in> \<U>\<close> dist_commute f mem_ball subsetCE) | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2078 | show "C \<in> \<B>" | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2079 | using \<open>\<U> \<subseteq> \<B>\<close> \<open>C \<in> \<U>\<close> by blast | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2080 | qed | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2081 | qed | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2082 | qed | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2083 | qed | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2084 | qed | 
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2085 | |
| 
e502cd4d7062
moved material from Connected.thy to more appropriate places
 immler parents: 
69613diff
changeset | 2086 | |
| 70136 | 2087 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Diameter\<close> | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2088 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2089 | lemma diameter_cball [simp]: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2090 | fixes a :: "'a::euclidean_space" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2091 | shows "diameter(cball a r) = (if r < 0 then 0 else 2*r)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2092 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2093 | have "diameter(cball a r) = 2*r" if "r \<ge> 0" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2094 | proof (rule order_antisym) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2095 | show "diameter (cball a r) \<le> 2*r" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2096 | proof (rule diameter_le) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2097 | fix x y assume "x \<in> cball a r" "y \<in> cball a r" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2098 | then have "norm (x - a) \<le> r" "norm (a - y) \<le> r" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2099 | by (auto simp: dist_norm norm_minus_commute) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2100 | then have "norm (x - y) \<le> r+r" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2101 | using norm_diff_triangle_le by blast | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2102 | then show "norm (x - y) \<le> 2*r" by simp | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2103 | qed (simp add: that) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2104 | have "2*r = dist (a + r *\<^sub>R (SOME i. i \<in> Basis)) (a - r *\<^sub>R (SOME i. i \<in> Basis))" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2105 | apply (simp add: dist_norm) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2106 | by (metis abs_of_nonneg mult.right_neutral norm_numeral norm_scaleR norm_some_Basis real_norm_def scaleR_2 that) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2107 | also have "... \<le> diameter (cball a r)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2108 | apply (rule diameter_bounded_bound) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2109 | using that by (auto simp: dist_norm) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2110 | finally show "2*r \<le> diameter (cball a r)" . | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2111 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2112 | then show ?thesis by simp | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2113 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2114 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2115 | lemma diameter_ball [simp]: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2116 | fixes a :: "'a::euclidean_space" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2117 | shows "diameter(ball a r) = (if r < 0 then 0 else 2*r)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2118 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2119 | have "diameter(ball a r) = 2*r" if "r > 0" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2120 | by (metis bounded_ball diameter_closure closure_ball diameter_cball less_eq_real_def linorder_not_less that) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2121 | then show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2122 | by (simp add: diameter_def) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2123 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2124 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2125 | lemma diameter_closed_interval [simp]: "diameter {a..b} = (if b < a then 0 else b-a)"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2126 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2127 |   have "{a .. b} = cball ((a+b)/2) ((b-a)/2)"
 | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 2128 | by (auto simp: dist_norm abs_if field_split_simps split: if_split_asm) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2129 | then show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2130 | by simp | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2131 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2132 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2133 | lemma diameter_open_interval [simp]: "diameter {a<..<b} = (if b < a then 0 else b-a)"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2134 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2135 |   have "{a <..< b} = ball ((a+b)/2) ((b-a)/2)"
 | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 2136 | by (auto simp: dist_norm abs_if field_split_simps split: if_split_asm) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2137 | then show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2138 | by simp | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2139 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2140 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2141 | lemma diameter_cbox: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2142 | fixes a b::"'a::euclidean_space" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2143 | shows "(\<forall>i \<in> Basis. a \<bullet> i \<le> b \<bullet> i) \<Longrightarrow> diameter (cbox a b) = dist a b" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2144 | by (force simp: diameter_def intro!: cSup_eq_maximum L2_set_mono | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2145 | simp: euclidean_dist_l2[where 'a='a] cbox_def dist_norm) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2146 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2147 | |
| 70136 | 2148 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Relating linear images to open/closed/interior/closure/connected\<close> | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 2149 | |
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2150 | proposition open_surjective_linear_image: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2151 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2152 | assumes "open A" "linear f" "surj f" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2153 | shows "open(f ` A)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2154 | unfolding open_dist | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2155 | proof clarify | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2156 | fix x | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2157 | assume "x \<in> A" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2158 | have "bounded (inv f ` Basis)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2159 | by (simp add: finite_imp_bounded) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2160 | with bounded_pos obtain B where "B > 0" and B: "\<And>x. x \<in> inv f ` Basis \<Longrightarrow> norm x \<le> B" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2161 | by metis | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2162 | obtain e where "e > 0" and e: "\<And>z. dist z x < e \<Longrightarrow> z \<in> A" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2163 | by (metis open_dist \<open>x \<in> A\<close> \<open>open A\<close>) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2164 |   define \<delta> where "\<delta> \<equiv> e / B / DIM('b)"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2165 | show "\<exists>e>0. \<forall>y. dist y (f x) < e \<longrightarrow> y \<in> f ` A" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2166 | proof (intro exI conjI) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2167 | show "\<delta> > 0" | 
| 70817 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 haftmann parents: 
70802diff
changeset | 2168 | using \<open>e > 0\<close> \<open>B > 0\<close> by (simp add: \<delta>_def field_split_simps) | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2169 |     have "y \<in> f ` A" if "dist y (f x) * (B * real DIM('b)) < e" for y
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2170 | proof - | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2171 | define u where "u \<equiv> y - f x" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2172 | show ?thesis | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2173 | proof (rule image_eqI) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2174 | show "y = f (x + (\<Sum>i\<in>Basis. (u \<bullet> i) *\<^sub>R inv f i))" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2175 | apply (simp add: linear_add linear_sum linear.scaleR \<open>linear f\<close> surj_f_inv_f \<open>surj f\<close>) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2176 | apply (simp add: euclidean_representation u_def) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2177 | done | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2178 | have "dist (x + (\<Sum>i\<in>Basis. (u \<bullet> i) *\<^sub>R inv f i)) x \<le> (\<Sum>i\<in>Basis. norm ((u \<bullet> i) *\<^sub>R inv f i))" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2179 | by (simp add: dist_norm sum_norm_le) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2180 | also have "... = (\<Sum>i\<in>Basis. \<bar>u \<bullet> i\<bar> * norm (inv f i))" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2181 | by simp | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2182 | also have "... \<le> (\<Sum>i\<in>Basis. \<bar>u \<bullet> i\<bar>) * B" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2183 | by (simp add: B sum_distrib_right sum_mono mult_left_mono) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2184 |         also have "... \<le> DIM('b) * dist y (f x) * B"
 | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2185 | apply (rule mult_right_mono [OF sum_bounded_above]) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2186 | using \<open>0 < B\<close> by (auto simp: Basis_le_norm dist_norm u_def) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2187 | also have "... < e" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2188 | by (metis mult.commute mult.left_commute that) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2189 | finally show "x + (\<Sum>i\<in>Basis. (u \<bullet> i) *\<^sub>R inv f i) \<in> A" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2190 | by (rule e) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2191 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2192 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2193 | then show "\<forall>y. dist y (f x) < \<delta> \<longrightarrow> y \<in> f ` A" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2194 | using \<open>e > 0\<close> \<open>B > 0\<close> | 
| 71172 | 2195 | by (auto simp: \<delta>_def field_split_simps) | 
| 69611 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2196 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2197 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2198 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2199 | corollary open_bijective_linear_image_eq: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2200 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2201 | assumes "linear f" "bij f" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2202 | shows "open(f ` A) \<longleftrightarrow> open A" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2203 | proof | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2204 | assume "open(f ` A)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2205 | then have "open(f -` (f ` A))" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2206 | using assms by (force simp: linear_continuous_at linear_conv_bounded_linear continuous_open_vimage) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2207 | then show "open A" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2208 | by (simp add: assms bij_is_inj inj_vimage_image_eq) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2209 | next | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2210 | assume "open A" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2211 | then show "open(f ` A)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2212 | by (simp add: assms bij_is_surj open_surjective_linear_image) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2213 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2214 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2215 | corollary interior_bijective_linear_image: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2216 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2217 | assumes "linear f" "bij f" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2218 | shows "interior (f ` S) = f ` interior S" (is "?lhs = ?rhs") | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2219 | proof safe | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2220 | fix x | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2221 | assume x: "x \<in> ?lhs" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2222 | then obtain T where "open T" and "x \<in> T" and "T \<subseteq> f ` S" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2223 | by (metis interiorE) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2224 | then show "x \<in> ?rhs" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2225 | by (metis (no_types, hide_lams) assms subsetD interior_maximal open_bijective_linear_image_eq subset_image_iff) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2226 | next | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2227 | fix x | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2228 | assume x: "x \<in> interior S" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2229 | then show "f x \<in> interior (f ` S)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2230 | by (meson assms imageI image_mono interiorI interior_subset open_bijective_linear_image_eq open_interior) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2231 | qed | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2232 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2233 | lemma interior_injective_linear_image: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2234 | fixes f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2235 | assumes "linear f" "inj f" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2236 | shows "interior(f ` S) = f ` (interior S)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2237 | by (simp add: linear_injective_imp_surjective assms bijI interior_bijective_linear_image) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2238 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2239 | lemma interior_surjective_linear_image: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2240 | fixes f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2241 | assumes "linear f" "surj f" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2242 | shows "interior(f ` S) = f ` (interior S)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2243 | by (simp add: assms interior_injective_linear_image linear_surjective_imp_injective) | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2244 | |
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2245 | lemma interior_negations: | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2246 | fixes S :: "'a::euclidean_space set" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2247 | shows "interior(uminus ` S) = image uminus (interior S)" | 
| 
42cc3609fedf
moved some material from Connected.thy to more appropriate places
 immler parents: 
69544diff
changeset | 2248 | by (simp add: bij_uminus interior_bijective_linear_image linear_uminus) | 
| 56189 
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
 immler parents: 
56188diff
changeset | 2249 | |
| 69617 | 2250 | lemma connected_linear_image: | 
| 2251 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | |
| 2252 | assumes "linear f" and "connected s" | |
| 2253 | shows "connected (f ` s)" | |
| 2254 | using connected_continuous_image assms linear_continuous_on linear_conv_bounded_linear by blast | |
| 2255 | ||
| 2256 | ||
| 70136 | 2257 | subsection\<^marker>\<open>tag unimportant\<close> \<open>"Isometry" (up to constant bounds) of Injective Linear Map\<close> | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2258 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2259 | proposition injective_imp_isometric: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2260 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2261 | assumes s: "closed s" "subspace s" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2262 | and f: "bounded_linear f" "\<forall>x\<in>s. f x = 0 \<longrightarrow> x = 0" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2263 | shows "\<exists>e>0. \<forall>x\<in>s. norm (f x) \<ge> e * norm x" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2264 | proof (cases "s \<subseteq> {0::'a}")
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2265 | case True | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2266 | have "norm x \<le> norm (f x)" if "x \<in> s" for x | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2267 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2268 | from True that have "x = 0" by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2269 | then show ?thesis by simp | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2270 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2271 | then show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2272 | by (auto intro!: exI[where x=1]) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2273 | next | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2274 | case False | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2275 | interpret f: bounded_linear f by fact | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2276 | from False obtain a where a: "a \<noteq> 0" "a \<in> s" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2277 | by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2278 |   from False have "s \<noteq> {}"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2279 | by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2280 |   let ?S = "{f x| x. x \<in> s \<and> norm x = norm a}"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2281 |   let ?S' = "{x::'a. x\<in>s \<and> norm x = norm a}"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2282 |   let ?S'' = "{x::'a. norm x = norm a}"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2283 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2284 | have "?S'' = frontier (cball 0 (norm a))" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2285 | by (simp add: sphere_def dist_norm) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2286 | then have "compact ?S''" by (metis compact_cball compact_frontier) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2287 | moreover have "?S' = s \<inter> ?S''" by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2288 | ultimately have "compact ?S'" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2289 | using closed_Int_compact[of s ?S''] using s(1) by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2290 | moreover have *:"f ` ?S' = ?S" by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2291 | ultimately have "compact ?S" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2292 | using compact_continuous_image[OF linear_continuous_on[OF f(1)], of ?S'] by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2293 | then have "closed ?S" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2294 | using compact_imp_closed by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2295 |   moreover from a have "?S \<noteq> {}" by auto
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2296 | ultimately obtain b' where "b'\<in>?S" "\<forall>y\<in>?S. norm b' \<le> norm y" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2297 | using distance_attains_inf[of ?S 0] unfolding dist_0_norm by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2298 | then obtain b where "b\<in>s" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2299 | and ba: "norm b = norm a" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2300 |     and b: "\<forall>x\<in>{x \<in> s. norm x = norm a}. norm (f b) \<le> norm (f x)"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2301 | unfolding *[symmetric] unfolding image_iff by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2302 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2303 | let ?e = "norm (f b) / norm b" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2304 | have "norm b > 0" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2305 | using ba and a and norm_ge_zero by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2306 | moreover have "norm (f b) > 0" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2307 | using f(2)[THEN bspec[where x=b], OF \<open>b\<in>s\<close>] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2308 | using \<open>norm b >0\<close> by simp | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2309 | ultimately have "0 < norm (f b) / norm b" by simp | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2310 | moreover | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2311 | have "norm (f b) / norm b * norm x \<le> norm (f x)" if "x\<in>s" for x | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2312 | proof (cases "x = 0") | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2313 | case True | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2314 | then show "norm (f b) / norm b * norm x \<le> norm (f x)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2315 | by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2316 | next | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2317 | case False | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2318 | with \<open>a \<noteq> 0\<close> have *: "0 < norm a / norm x" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2319 | unfolding zero_less_norm_iff[symmetric] by simp | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2320 | have "\<forall>x\<in>s. c *\<^sub>R x \<in> s" for c | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2321 | using s[unfolded subspace_def] by simp | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2322 |     with \<open>x \<in> s\<close> \<open>x \<noteq> 0\<close> have "(norm a / norm x) *\<^sub>R x \<in> {x \<in> s. norm x = norm a}"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2323 | by simp | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2324 | with \<open>x \<noteq> 0\<close> \<open>a \<noteq> 0\<close> show "norm (f b) / norm b * norm x \<le> norm (f x)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2325 | using b[THEN bspec[where x="(norm a / norm x) *\<^sub>R x"]] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2326 | unfolding f.scaleR and ba | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2327 | by (auto simp: mult.commute pos_le_divide_eq pos_divide_le_eq) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2328 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2329 | ultimately show ?thesis by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2330 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2331 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2332 | proposition closed_injective_image_subspace: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2333 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2334 | assumes "subspace s" "bounded_linear f" "\<forall>x\<in>s. f x = 0 \<longrightarrow> x = 0" "closed s" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2335 | shows "closed(f ` s)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2336 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2337 | obtain e where "e > 0" and e: "\<forall>x\<in>s. e * norm x \<le> norm (f x)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2338 | using injective_imp_isometric[OF assms(4,1,2,3)] by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2339 | show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2340 | using complete_isometric_image[OF \<open>e>0\<close> assms(1,2) e] and assms(4) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2341 | unfolding complete_eq_closed[symmetric] by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2342 | qed | 
| 71255 | 2343 | |
| 2344 | ||
| 2345 | lemma closure_bounded_linear_image_subset: | |
| 2346 | assumes f: "bounded_linear f" | |
| 2347 | shows "f ` closure S \<subseteq> closure (f ` S)" | |
| 2348 | using linear_continuous_on [OF f] closed_closure closure_subset | |
| 2349 | by (rule image_closure_subset) | |
| 2350 | ||
| 2351 | lemma closure_linear_image_subset: | |
| 2352 | fixes f :: "'m::euclidean_space \<Rightarrow> 'n::real_normed_vector" | |
| 2353 | assumes "linear f" | |
| 2354 | shows "f ` (closure S) \<subseteq> closure (f ` S)" | |
| 2355 | using assms unfolding linear_conv_bounded_linear | |
| 2356 | by (rule closure_bounded_linear_image_subset) | |
| 2357 | ||
| 2358 | lemma closed_injective_linear_image: | |
| 2359 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | |
| 2360 | assumes S: "closed S" and f: "linear f" "inj f" | |
| 2361 | shows "closed (f ` S)" | |
| 2362 | proof - | |
| 2363 | obtain g where g: "linear g" "g \<circ> f = id" | |
| 2364 | using linear_injective_left_inverse [OF f] by blast | |
| 2365 | then have confg: "continuous_on (range f) g" | |
| 2366 | using linear_continuous_on linear_conv_bounded_linear by blast | |
| 2367 | have [simp]: "g ` f ` S = S" | |
| 2368 | using g by (simp add: image_comp) | |
| 2369 | have cgf: "closed (g ` f ` S)" | |
| 2370 | by (simp add: \<open>g \<circ> f = id\<close> S image_comp) | |
| 2371 | have [simp]: "(range f \<inter> g -` S) = f ` S" | |
| 2372 | using g unfolding o_def id_def image_def by auto metis+ | |
| 2373 | show ?thesis | |
| 2374 | proof (rule closedin_closed_trans [of "range f"]) | |
| 2375 | show "closedin (top_of_set (range f)) (f ` S)" | |
| 2376 | using continuous_closedin_preimage [OF confg cgf] by simp | |
| 2377 | show "closed (range f)" | |
| 2378 | apply (rule closed_injective_image_subspace) | |
| 2379 | using f apply (auto simp: linear_linear linear_injective_0) | |
| 2380 | done | |
| 2381 | qed | |
| 2382 | qed | |
| 2383 | ||
| 2384 | lemma closed_injective_linear_image_eq: | |
| 2385 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | |
| 2386 | assumes f: "linear f" "inj f" | |
| 2387 | shows "(closed(image f s) \<longleftrightarrow> closed s)" | |
| 2388 | by (metis closed_injective_linear_image closure_eq closure_linear_image_subset closure_subset_eq f(1) f(2) inj_image_subset_iff) | |
| 2389 | ||
| 2390 | lemma closure_injective_linear_image: | |
| 2391 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | |
| 2392 | shows "\<lbrakk>linear f; inj f\<rbrakk> \<Longrightarrow> f ` (closure S) = closure (f ` S)" | |
| 2393 | apply (rule subset_antisym) | |
| 2394 | apply (simp add: closure_linear_image_subset) | |
| 2395 | by (simp add: closure_minimal closed_injective_linear_image closure_subset image_mono) | |
| 2396 | ||
| 2397 | lemma closure_bounded_linear_image: | |
| 2398 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | |
| 2399 | shows "\<lbrakk>linear f; bounded S\<rbrakk> \<Longrightarrow> f ` (closure S) = closure (f ` S)" | |
| 2400 | apply (rule subset_antisym, simp add: closure_linear_image_subset) | |
| 2401 | apply (rule closure_minimal, simp add: closure_subset image_mono) | |
| 2402 | by (meson bounded_closure closed_closure compact_continuous_image compact_eq_bounded_closed linear_continuous_on linear_conv_bounded_linear) | |
| 2403 | ||
| 2404 | lemma closure_scaleR: | |
| 2405 | fixes S :: "'a::real_normed_vector set" | |
| 2406 | shows "((*\<^sub>R) c) ` (closure S) = closure (((*\<^sub>R) c) ` S)" | |
| 2407 | proof | |
| 2408 | show "((*\<^sub>R) c) ` (closure S) \<subseteq> closure (((*\<^sub>R) c) ` S)" | |
| 2409 | using bounded_linear_scaleR_right | |
| 2410 | by (rule closure_bounded_linear_image_subset) | |
| 2411 | show "closure (((*\<^sub>R) c) ` S) \<subseteq> ((*\<^sub>R) c) ` (closure S)" | |
| 2412 | by (intro closure_minimal image_mono closure_subset closed_scaling closed_closure) | |
| 2413 | qed | |
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2414 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2415 | |
| 70136 | 2416 | subsection\<^marker>\<open>tag unimportant\<close> \<open>Some properties of a canonical subspace\<close> | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2417 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2418 | lemma closed_substandard: "closed {x::'a::euclidean_space. \<forall>i\<in>Basis. P i \<longrightarrow> x\<bullet>i = 0}"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2419 | (is "closed ?A") | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2420 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2421 |   let ?D = "{i\<in>Basis. P i}"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2422 |   have "closed (\<Inter>i\<in>?D. {x::'a. x\<bullet>i = 0})"
 | 
| 71172 | 2423 | by (simp add: closed_INT closed_Collect_eq continuous_on_inner) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2424 |   also have "(\<Inter>i\<in>?D. {x::'a. x\<bullet>i = 0}) = ?A"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2425 | by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2426 | finally show "closed ?A" . | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2427 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2428 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2429 | lemma closed_subspace: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2430 | fixes s :: "'a::euclidean_space set" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2431 | assumes "subspace s" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2432 | shows "closed s" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2433 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2434 | have "dim s \<le> card (Basis :: 'a set)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2435 | using dim_subset_UNIV by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2436 | with ex_card[OF this] obtain d :: "'a set" where t: "card d = dim s" and d: "d \<subseteq> Basis" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2437 | by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2438 |   let ?t = "{x::'a. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2439 |   have "\<exists>f. linear f \<and> f ` {x::'a. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0} = s \<and>
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2440 |       inj_on f {x::'a. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0}"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2441 | using dim_substandard[of d] t d assms | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2442 | by (intro subspace_isomorphism[OF subspace_substandard[of "\<lambda>i. i \<notin> d"]]) (auto simp: inner_Basis) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2443 | then obtain f where f: | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2444 | "linear f" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2445 |       "f ` {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0} = s"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2446 |       "inj_on f {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0}"
 | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2447 | by blast | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2448 | interpret f: bounded_linear f | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2449 | using f by (simp add: linear_conv_bounded_linear) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2450 | have "x \<in> ?t \<Longrightarrow> f x = 0 \<Longrightarrow> x = 0" for x | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2451 | using f.zero d f(3)[THEN inj_onD, of x 0] by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2452 | moreover have "closed ?t" by (rule closed_substandard) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2453 | moreover have "subspace ?t" by (rule subspace_substandard) | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2454 | ultimately show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2455 | using closed_injective_image_subspace[of ?t f] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2456 | unfolding f(2) using f(1) unfolding linear_conv_bounded_linear by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2457 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2458 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2459 | lemma complete_subspace: "subspace s \<Longrightarrow> complete s" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2460 | for s :: "'a::euclidean_space set" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2461 | using complete_eq_closed closed_subspace by auto | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2462 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2463 | lemma closed_span [iff]: "closed (span s)" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2464 | for s :: "'a::euclidean_space set" | 
| 71172 | 2465 | by (simp add: closed_subspace) | 
| 69613 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2466 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2467 | lemma dim_closure [simp]: "dim (closure s) = dim s" (is "?dc = ?d") | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2468 | for s :: "'a::euclidean_space set" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2469 | proof - | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2470 | have "?dc \<le> ?d" | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2471 | using closure_minimal[OF span_superset, of s] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2472 | using closed_subspace[OF subspace_span, of s] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2473 | using dim_subset[of "closure s" "span s"] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2474 | by simp | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2475 | then show ?thesis | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2476 | using dim_subset[OF closure_subset, of s] | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2477 | by simp | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2478 | qed | 
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2479 | |
| 
1331e57b54c6
moved material from Connected.thy to more appropriate places
 immler parents: 
69611diff
changeset | 2480 | |
| 69618 | 2481 | subsection \<open>Set Distance\<close> | 
| 2482 | ||
| 2483 | lemma setdist_compact_closed: | |
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2484 | fixes A :: "'a::heine_borel set" | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2485 | assumes A: "compact A" and B: "closed B" | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2486 |     and "A \<noteq> {}" "B \<noteq> {}"
 | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2487 | shows "\<exists>x \<in> A. \<exists>y \<in> B. dist x y = setdist A B" | 
| 69618 | 2488 | proof - | 
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2489 | obtain x where "x \<in> A" "setdist A B = infdist x B" | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2490 | by (metis A assms(3) setdist_attains_inf setdist_sym) | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2491 | moreover | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2492 | obtain y where"y \<in> B" "infdist x B = dist x y" | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2493 |     using B \<open>B \<noteq> {}\<close> infdist_attains_inf by blast
 | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2494 | ultimately show ?thesis | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2495 | using \<open>x \<in> A\<close> \<open>y \<in> B\<close> by auto | 
| 69618 | 2496 | qed | 
| 2497 | ||
| 2498 | lemma setdist_closed_compact: | |
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2499 | fixes S :: "'a::heine_borel set" | 
| 69618 | 2500 | assumes S: "closed S" and T: "compact T" | 
| 2501 |       and "S \<noteq> {}" "T \<noteq> {}"
 | |
| 2502 | shows "\<exists>x \<in> S. \<exists>y \<in> T. dist x y = setdist S T" | |
| 2503 |   using setdist_compact_closed [OF T S \<open>T \<noteq> {}\<close> \<open>S \<noteq> {}\<close>]
 | |
| 2504 | by (metis dist_commute setdist_sym) | |
| 2505 | ||
| 2506 | lemma setdist_eq_0_compact_closed: | |
| 2507 | assumes S: "compact S" and T: "closed T" | |
| 2508 |     shows "setdist S T = 0 \<longleftrightarrow> S = {} \<or> T = {} \<or> S \<inter> T \<noteq> {}"
 | |
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2509 | proof (cases "S = {} \<or> T = {}")
 | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2510 | case True | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2511 | then show ?thesis | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2512 | by force | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2513 | next | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2514 | case False | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2515 | then show ?thesis | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2516 | by (metis S T disjoint_iff_not_equal in_closed_iff_infdist_zero setdist_attains_inf setdist_eq_0I setdist_sym) | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2517 | qed | 
| 69618 | 2518 | |
| 2519 | corollary setdist_gt_0_compact_closed: | |
| 2520 | assumes S: "compact S" and T: "closed T" | |
| 2521 |     shows "setdist S T > 0 \<longleftrightarrow> (S \<noteq> {} \<and> T \<noteq> {} \<and> S \<inter> T = {})"
 | |
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2522 | using setdist_pos_le [of S T] setdist_eq_0_compact_closed [OF assms] by linarith | 
| 69618 | 2523 | |
| 2524 | lemma setdist_eq_0_closed_compact: | |
| 2525 | assumes S: "closed S" and T: "compact T" | |
| 2526 |     shows "setdist S T = 0 \<longleftrightarrow> S = {} \<or> T = {} \<or> S \<inter> T \<noteq> {}"
 | |
| 2527 | using setdist_eq_0_compact_closed [OF T S] | |
| 2528 | by (metis Int_commute setdist_sym) | |
| 2529 | ||
| 2530 | lemma setdist_eq_0_bounded: | |
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2531 | fixes S :: "'a::heine_borel set" | 
| 69618 | 2532 | assumes "bounded S \<or> bounded T" | 
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2533 |   shows "setdist S T = 0 \<longleftrightarrow> S = {} \<or> T = {} \<or> closure S \<inter> closure T \<noteq> {}"
 | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2534 | proof (cases "S = {} \<or> T = {}")
 | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2535 | case False | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2536 | then show ?thesis | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2537 | using setdist_eq_0_compact_closed [of "closure S" "closure T"] | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2538 | setdist_eq_0_closed_compact [of "closure S" "closure T"] assms | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2539 | by (force simp: bounded_closure compact_eq_bounded_closed) | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2540 | qed force | 
| 69618 | 2541 | |
| 2542 | lemma setdist_eq_0_sing_1: | |
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2543 |   "setdist {x} S = 0 \<longleftrightarrow> S = {} \<or> x \<in> closure S"
 | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2544 | by (metis in_closure_iff_infdist_zero infdist_def infdist_eq_setdist) | 
| 69618 | 2545 | |
| 2546 | lemma setdist_eq_0_sing_2: | |
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2547 |   "setdist S {x} = 0 \<longleftrightarrow> S = {} \<or> x \<in> closure S"
 | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2548 | by (metis setdist_eq_0_sing_1 setdist_sym) | 
| 69618 | 2549 | |
| 2550 | lemma setdist_neq_0_sing_1: | |
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2551 |   "\<lbrakk>setdist {x} S = a; a \<noteq> 0\<rbrakk> \<Longrightarrow> S \<noteq> {} \<and> x \<notin> closure S"
 | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2552 | by (metis setdist_closure_2 setdist_empty2 setdist_eq_0I singletonI) | 
| 69618 | 2553 | |
| 2554 | lemma setdist_neq_0_sing_2: | |
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2555 |   "\<lbrakk>setdist S {x} = a; a \<noteq> 0\<rbrakk> \<Longrightarrow> S \<noteq> {} \<and> x \<notin> closure S"
 | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2556 | by (simp add: setdist_neq_0_sing_1 setdist_sym) | 
| 69618 | 2557 | |
| 2558 | lemma setdist_sing_in_set: | |
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2559 |    "x \<in> S \<Longrightarrow> setdist {x} S = 0"
 | 
| 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2560 | by (simp add: setdist_eq_0I) | 
| 69618 | 2561 | |
| 2562 | lemma setdist_eq_0_closed: | |
| 69918 
eddcc7c726f3
new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
 paulson <lp15@cam.ac.uk> parents: 
69750diff
changeset | 2563 |    "closed S \<Longrightarrow> (setdist {x} S = 0 \<longleftrightarrow> S = {} \<or> x \<in> S)"
 | 
| 69618 | 2564 | by (simp add: setdist_eq_0_sing_1) | 
| 2565 | ||
| 2566 | lemma setdist_eq_0_closedin: | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2567 | shows "\<lbrakk>closedin (top_of_set U) S; x \<in> U\<rbrakk> | 
| 69618 | 2568 |          \<Longrightarrow> (setdist {x} S = 0 \<longleftrightarrow> S = {} \<or> x \<in> S)"
 | 
| 2569 | by (auto simp: closedin_limpt setdist_eq_0_sing_1 closure_def) | |
| 2570 | ||
| 2571 | lemma setdist_gt_0_closedin: | |
| 69922 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 paulson <lp15@cam.ac.uk> parents: 
69918diff
changeset | 2572 |   shows "\<lbrakk>closedin (top_of_set U) S; x \<in> U; S \<noteq> {}; x \<notin> S\<rbrakk>
 | 
| 69618 | 2573 |          \<Longrightarrow> setdist {x} S > 0"
 | 
| 2574 | using less_eq_real_def setdist_eq_0_closedin by fastforce | |
| 2575 | ||
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 2576 | no_notation | 
| 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 2577 | eucl_less (infix "<e" 50) | 
| 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 2578 | |
| 33175 | 2579 | end |