| author | wenzelm | 
| Sun, 17 Dec 2023 11:15:08 +0100 | |
| changeset 79268 | 154166613b40 | 
| parent 77943 | ffdad62bc235 | 
| child 82538 | 4b132ea7d575 | 
| permissions | -rw-r--r-- | 
| 68038 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 1 | (* Title: HOL/Analysis/Cartesian_Euclidean_Space.thy | 
| 68043 
d345e9c35ae1
some of Jose Divasón's material from Rank_Nullity_Theorem/Miscellaneous
 paulson <lp15@cam.ac.uk> parents: 
68041diff
changeset | 2 | Some material by Jose Divasón, Tim Makarios and L C Paulson | 
| 68038 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 3 | *) | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 4 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 5 | section \<open>Finite Cartesian Products of Euclidean Spaces\<close> | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 6 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 7 | theory Cartesian_Euclidean_Space | 
| 71175 
a1e94be66bd6
reduced imports; deleted unusewd minor lemmas for that purpose
 nipkow parents: 
71172diff
changeset | 8 | imports Derivative | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 9 | begin | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 10 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 11 | lemma subspace_special_hyperplane: "subspace {x. x $ k = 0}"
 | 
| 63016 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 12 | by (simp add: subspace_def) | 
| 
3590590699b1
numerous theorems about affine hulls, hyperplanes, etc.
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 13 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 14 | lemma sum_mult_product: | 
| 64267 | 15 |   "sum h {..<A * B :: nat} = (\<Sum>i\<in>{..<A}. \<Sum>j\<in>{..<B}. h (j + i * B))"
 | 
| 70113 
c8deb8ba6d05
Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
 paulson <lp15@cam.ac.uk> parents: 
69723diff
changeset | 16 | unfolding sum.nat_group[of h B A, unfolded atLeast0LessThan, symmetric] | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 17 | proof (rule sum.cong, simp, rule sum.reindex_cong) | 
| 49644 | 18 | fix i | 
| 19 |   show "inj_on (\<lambda>j. j + i * B) {..<B}" by (auto intro!: inj_onI)
 | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 20 |   show "{i * B..<i * B + B} = (\<lambda>j. j + i * B) ` {..<B}"
 | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 21 | proof safe | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 22 |     fix j assume "j \<in> {i * B..<i * B + B}"
 | 
| 49644 | 23 |     then show "j \<in> (\<lambda>j. j + i * B) ` {..<B}"
 | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 24 | by (auto intro!: image_eqI[of _ _ "j - i * B"]) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 25 | qed simp | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 26 | qed simp | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 27 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 28 | lemma interval_cbox_cart: "{a::real^'n..b} = cbox a b"
 | 
| 56188 | 29 | by (auto simp add: less_eq_vec_def mem_box Basis_vec_def inner_axis) | 
| 30 | ||
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 31 | lemma differentiable_vec: | 
| 67979 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 32 | fixes S :: "'a::euclidean_space set" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 33 | shows "vec differentiable_on S" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 34 | by (simp add: linear_linear bounded_linear_imp_differentiable_on) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 35 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 36 | lemma continuous_vec [continuous_intros]: | 
| 67979 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 37 | fixes x :: "'a::euclidean_space" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 38 | shows "isCont vec x" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 39 | apply (clarsimp simp add: continuous_def LIM_def dist_vec_def L2_set_def) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 40 |   apply (rule_tac x="r / sqrt (real CARD('b))" in exI)
 | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 41 | by (simp add: mult.commute pos_less_divide_eq real_sqrt_mult) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 42 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 43 | lemma box_vec_eq_empty [simp]: | 
| 67979 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 44 |   shows "cbox (vec a) (vec b) = {} \<longleftrightarrow> cbox a b = {}"
 | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 45 |         "box (vec a) (vec b) = {} \<longleftrightarrow> box a b = {}"
 | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 46 | by (auto simp: Basis_vec_def mem_box box_eq_empty inner_axis) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 47 | |
| 69683 | 48 | subsection\<open>Closures and interiors of halfspaces\<close> | 
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 49 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 50 | lemma interior_halfspace_component_le [simp]: | 
| 67731 | 51 |      "interior {x. x$k \<le> a} = {x :: (real^'n). x$k < a}" (is "?LE")
 | 
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 52 | and interior_halfspace_component_ge [simp]: | 
| 67731 | 53 |      "interior {x. x$k \<ge> a} = {x :: (real^'n). x$k > a}" (is "?GE")
 | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 54 | proof - | 
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 55 | have "axis k (1::real) \<noteq> 0" | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 56 | by (simp add: axis_def vec_eq_iff) | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 57 | moreover have "axis k (1::real) \<bullet> x = x$k" for x | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 58 | by (simp add: cart_eq_inner_axis inner_commute) | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 59 | ultimately show ?LE ?GE | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 60 | using interior_halfspace_le [of "axis k (1::real)" a] | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 61 | interior_halfspace_ge [of "axis k (1::real)" a] by auto | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 62 | qed | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 63 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 64 | lemma closure_halfspace_component_lt [simp]: | 
| 67731 | 65 |      "closure {x. x$k < a} = {x :: (real^'n). x$k \<le> a}" (is "?LE")
 | 
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 66 | and closure_halfspace_component_gt [simp]: | 
| 67731 | 67 |      "closure {x. x$k > a} = {x :: (real^'n). x$k \<ge> a}" (is "?GE")
 | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 68 | proof - | 
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 69 | have "axis k (1::real) \<noteq> 0" | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 70 | by (simp add: axis_def vec_eq_iff) | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 71 | moreover have "axis k (1::real) \<bullet> x = x$k" for x | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 72 | by (simp add: cart_eq_inner_axis inner_commute) | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 73 | ultimately show ?LE ?GE | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 74 | using closure_halfspace_lt [of "axis k (1::real)" a] | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 75 | closure_halfspace_gt [of "axis k (1::real)" a] by auto | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 76 | qed | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 77 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 78 | lemma interior_standard_hyperplane: | 
| 67731 | 79 |    "interior {x :: (real^'n). x$k = a} = {}"
 | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 80 | proof - | 
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 81 | have "axis k (1::real) \<noteq> 0" | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 82 | by (simp add: axis_def vec_eq_iff) | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 83 | moreover have "axis k (1::real) \<bullet> x = x$k" for x | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 84 | by (simp add: cart_eq_inner_axis inner_commute) | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 85 | ultimately show ?thesis | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 86 | using interior_hyperplane [of "axis k (1::real)" a] | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 87 | by force | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 88 | qed | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 89 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 90 | lemma matrix_vector_mul_bounded_linear[intro, simp]: "bounded_linear ((*v) A)" for A :: "'a::{euclidean_space,real_algebra_1}^'n^'m"
 | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 91 | using matrix_vector_mul_linear[of A] | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 92 | by (simp add: linear_conv_bounded_linear linear_matrix_vector_mul_eq) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 93 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 94 | lemma | 
| 68073 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immlerdiff
changeset | 95 |   fixes A :: "'a::{euclidean_space,real_algebra_1}^'n^'m"
 | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 96 | shows matrix_vector_mult_linear_continuous_at [continuous_intros]: "isCont ((*v) A) z" | 
| 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 97 | and matrix_vector_mult_linear_continuous_on [continuous_intros]: "continuous_on S ((*v) A)" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 98 | by (simp_all add: linear_continuous_at linear_continuous_on) | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 99 | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 100 | |
| 69683 | 101 | subsection\<open>Bounds on components etc.\ relative to operator norm\<close> | 
| 67719 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 102 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 103 | lemma norm_column_le_onorm: | 
| 67719 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 104 | fixes A :: "real^'n^'m" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 105 | shows "norm(column i A) \<le> onorm((*v) A)" | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 106 | proof - | 
| 67719 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 107 | have "norm (\<chi> j. A $ j $ i) \<le> norm (A *v axis i 1)" | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 108 | by (simp add: matrix_mult_dot cart_eq_inner_axis) | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 109 | also have "\<dots> \<le> onorm ((*v) A)" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 110 | using onorm [OF matrix_vector_mul_bounded_linear, of A "axis i 1"] by auto | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 111 | finally have "norm (\<chi> j. A $ j $ i) \<le> onorm ((*v) A)" . | 
| 67719 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 112 | then show ?thesis | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 113 | unfolding column_def . | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 114 | qed | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 115 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 116 | lemma matrix_component_le_onorm: | 
| 67719 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 117 | fixes A :: "real^'n^'m" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 118 | shows "\<bar>A $ i $ j\<bar> \<le> onorm((*v) A)" | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 119 | proof - | 
| 67719 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 120 | have "\<bar>A $ i $ j\<bar> \<le> norm (\<chi> n. (A $ n $ j))" | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 121 | by (metis (full_types, lifting) component_le_norm_cart vec_lambda_beta) | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 122 | also have "\<dots> \<le> onorm ((*v) A)" | 
| 67719 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 123 | by (metis (no_types) column_def norm_column_le_onorm) | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 124 | finally show ?thesis . | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 125 | qed | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 126 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 127 | lemma component_le_onorm: | 
| 67719 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 128 | fixes f :: "real^'m \<Rightarrow> real^'n" | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 129 | shows "linear f \<Longrightarrow> \<bar>matrix f $ i $ j\<bar> \<le> onorm f" | 
| 71172 | 130 | by (metis matrix_component_le_onorm matrix_vector_mul(2)) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 131 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 132 | lemma onorm_le_matrix_component_sum: | 
| 67719 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 133 | fixes A :: "real^'n^'m" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 134 | shows "onorm((*v) A) \<le> (\<Sum>i\<in>UNIV. \<Sum>j\<in>UNIV. \<bar>A $ i $ j\<bar>)" | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 135 | proof (rule onorm_le) | 
| 67719 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 136 | fix x | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 137 | have "norm (A *v x) \<le> (\<Sum>i\<in>UNIV. \<bar>(A *v x) $ i\<bar>)" | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 138 | by (rule norm_le_l1_cart) | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 139 | also have "\<dots> \<le> (\<Sum>i\<in>UNIV. \<Sum>j\<in>UNIV. \<bar>A $ i $ j\<bar> * norm x)" | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 140 | proof (rule sum_mono) | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 141 | fix i | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 142 | have "\<bar>(A *v x) $ i\<bar> \<le> \<bar>\<Sum>j\<in>UNIV. A $ i $ j * x $ j\<bar>" | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 143 | by (simp add: matrix_vector_mult_def) | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 144 | also have "\<dots> \<le> (\<Sum>j\<in>UNIV. \<bar>A $ i $ j * x $ j\<bar>)" | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 145 | by (rule sum_abs) | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 146 | also have "\<dots> \<le> (\<Sum>j\<in>UNIV. \<bar>A $ i $ j\<bar> * norm x)" | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 147 | by (rule sum_mono) (simp add: abs_mult component_le_norm_cart mult_left_mono) | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 148 | finally show "\<bar>(A *v x) $ i\<bar> \<le> (\<Sum>j\<in>UNIV. \<bar>A $ i $ j\<bar> * norm x)" . | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 149 | qed | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 150 | finally show "norm (A *v x) \<le> (\<Sum>i\<in>UNIV. \<Sum>j\<in>UNIV. \<bar>A $ i $ j\<bar>) * norm x" | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 151 | by (simp add: sum_distrib_right) | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 152 | qed | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 153 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 154 | lemma onorm_le_matrix_component: | 
| 67719 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 155 | fixes A :: "real^'n^'m" | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 156 | assumes "\<And>i j. abs(A$i$j) \<le> B" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 157 |   shows "onorm((*v) A) \<le> real (CARD('m)) * real (CARD('n)) * B"
 | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 158 | proof (rule onorm_le) | 
| 67731 | 159 | fix x :: "real^'n::_" | 
| 67719 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 160 | have "norm (A *v x) \<le> (\<Sum>i\<in>UNIV. \<bar>(A *v x) $ i\<bar>)" | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 161 | by (rule norm_le_l1_cart) | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 162 |   also have "\<dots> \<le> (\<Sum>i::'m \<in>UNIV. real (CARD('n)) * B * norm x)"
 | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 163 | proof (rule sum_mono) | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 164 | fix i | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 165 | have "\<bar>(A *v x) $ i\<bar> \<le> norm(A $ i) * norm x" | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 166 | by (simp add: matrix_mult_dot Cauchy_Schwarz_ineq2) | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 167 | also have "\<dots> \<le> (\<Sum>j\<in>UNIV. \<bar>A $ i $ j\<bar>) * norm x" | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 168 | by (simp add: mult_right_mono norm_le_l1_cart) | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 169 |     also have "\<dots> \<le> real (CARD('n)) * B * norm x"
 | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 170 | by (simp add: assms sum_bounded_above mult_right_mono) | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 171 |     finally show "\<bar>(A *v x) $ i\<bar> \<le> real (CARD('n)) * B * norm x" .
 | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 172 | qed | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 173 |   also have "\<dots> \<le> CARD('m) * real (CARD('n)) * B * norm x"
 | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 174 | by simp | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 175 |   finally show "norm (A *v x) \<le> CARD('m) * real (CARD('n)) * B * norm x" .
 | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 176 | qed | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 177 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 178 | lemma vector_sub_project_orthogonal_cart: "(b::real^'n) \<bullet> (x - ((b \<bullet> x) / (b \<bullet> b)) *s b) = 0" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49962diff
changeset | 179 | unfolding inner_simps scalar_mult_eq_scaleR by auto | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 180 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 181 | lemma infnorm_cart:"infnorm (x::real^'n) = Sup {\<bar>x$i\<bar> |i. i\<in>UNIV}"
 | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49962diff
changeset | 182 | by (simp add: infnorm_def inner_axis Basis_vec_def) (metis (lifting) inner_axis real_inner_1_right) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 183 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 184 | lemma component_le_infnorm_cart: "\<bar>x$i\<bar> \<le> infnorm (x::real^'n)" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49962diff
changeset | 185 | using Basis_le_infnorm[of "axis i 1" x] | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49962diff
changeset | 186 | by (simp add: Basis_vec_def axis_eq_axis inner_axis) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 187 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 188 | lemma continuous_component[continuous_intros]: "continuous F f \<Longrightarrow> continuous F (\<lambda>x. f x $ i)" | 
| 44647 
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
 huffman parents: 
44571diff
changeset | 189 | unfolding continuous_def by (rule tendsto_vec_nth) | 
| 44213 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 huffman parents: 
44211diff
changeset | 190 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 191 | lemma continuous_on_component[continuous_intros]: "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. f x $ i)" | 
| 44647 
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
 huffman parents: 
44571diff
changeset | 192 | unfolding continuous_on_def by (fast intro: tendsto_vec_nth) | 
| 44213 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 huffman parents: 
44211diff
changeset | 193 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 194 | lemma continuous_on_vec_lambda[continuous_intros]: | 
| 63334 
bd37a72a940a
Multivariate_Analysis: add continuous_on_vec_lambda
 hoelzl parents: 
63332diff
changeset | 195 | "(\<And>i. continuous_on S (f i)) \<Longrightarrow> continuous_on S (\<lambda>x. \<chi> i. f i x)" | 
| 
bd37a72a940a
Multivariate_Analysis: add continuous_on_vec_lambda
 hoelzl parents: 
63332diff
changeset | 196 | unfolding continuous_on_def by (auto intro: tendsto_vec_lambda) | 
| 
bd37a72a940a
Multivariate_Analysis: add continuous_on_vec_lambda
 hoelzl parents: 
63332diff
changeset | 197 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 198 | lemma closed_positive_orthant: "closed {x::real^'n. \<forall>i. 0 \<le>x$i}"
 | 
| 71172 | 199 | by (simp add: Collect_all_eq closed_INT closed_Collect_le continuous_on_component) | 
| 44213 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 huffman parents: 
44211diff
changeset | 200 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 201 | lemma bounded_component_cart: "bounded s \<Longrightarrow> bounded ((\<lambda>x. x $ i) ` s)" | 
| 49644 | 202 | unfolding bounded_def | 
| 203 | apply clarify | |
| 204 | apply (rule_tac x="x $ i" in exI) | |
| 205 | apply (rule_tac x="e" in exI) | |
| 206 | apply clarify | |
| 207 | apply (rule order_trans [OF dist_vec_nth_le], simp) | |
| 208 | done | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 209 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 210 | lemma compact_lemma_cart: | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 211 | fixes f :: "nat \<Rightarrow> 'a::heine_borel ^ 'n" | 
| 50998 | 212 | assumes f: "bounded (range f)" | 
| 66447 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 213 | shows "\<exists>l r. strict_mono r \<and> | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 214 | (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) $ i) (l $ i) < e) sequentially)" | 
| 62127 | 215 | (is "?th d") | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 216 | proof - | 
| 62127 | 217 | have "\<forall>d' \<subseteq> d. ?th d'" | 
| 218 | by (rule compact_lemma_general[where unproj=vec_lambda]) | |
| 71172 | 219 | (auto intro!: f bounded_component_cart) | 
| 62127 | 220 | then show "?th d" by simp | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 221 | qed | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 222 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 223 | instance vec :: (heine_borel, finite) heine_borel | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 224 | proof | 
| 50998 | 225 | fix f :: "nat \<Rightarrow> 'a ^ 'b" | 
| 226 | assume f: "bounded (range f)" | |
| 66447 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 227 | then obtain l r where r: "strict_mono r" | 
| 49644 | 228 | and l: "\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>UNIV. dist (f (r n) $ i) (l $ i) < e) sequentially" | 
| 50998 | 229 | using compact_lemma_cart [OF f] by blast | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 230 | let ?d = "UNIV::'b set" | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 231 |   { fix e::real assume "e>0"
 | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 232 | hence "0 < e / (real_of_nat (card ?d))" | 
| 49644 | 233 | using zero_less_card_finite divide_pos_pos[of e, of "real_of_nat (card ?d)"] by auto | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 234 | with l have "eventually (\<lambda>n. \<forall>i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))) sequentially" | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 235 | by simp | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 236 | moreover | 
| 49644 | 237 |     { fix n
 | 
| 238 | assume n: "\<forall>i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))" | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 239 | have "dist (f (r n)) l \<le> (\<Sum>i\<in>?d. dist (f (r n) $ i) (l $ i))" | 
| 67155 | 240 | unfolding dist_vec_def using zero_le_dist by (rule L2_set_le_sum) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 241 | also have "\<dots> < (\<Sum>i\<in>?d. e / (real_of_nat (card ?d)))" | 
| 64267 | 242 | by (rule sum_strict_mono) (simp_all add: n) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 243 | finally have "dist (f (r n)) l < e" by simp | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 244 | } | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 245 | ultimately have "eventually (\<lambda>n. dist (f (r n)) l < e) sequentially" | 
| 61810 | 246 | by (rule eventually_mono) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 247 | } | 
| 61973 | 248 | hence "((f \<circ> r) \<longlongrightarrow> l) sequentially" unfolding o_def tendsto_iff by simp | 
| 66447 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 249 | with r show "\<exists>l r. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially" by auto | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 250 | qed | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 251 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 252 | lemma interval_cart: | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 253 | fixes a :: "real^'n" | 
| 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 254 |   shows "box a b = {x::real^'n. \<forall>i. a$i < x$i \<and> x$i < b$i}"
 | 
| 56188 | 255 |     and "cbox a b = {x::real^'n. \<forall>i. a$i \<le> x$i \<and> x$i \<le> b$i}"
 | 
| 256 | by (auto simp add: set_eq_iff less_vec_def less_eq_vec_def mem_box Basis_vec_def inner_axis) | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 257 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 258 | lemma mem_box_cart: | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 259 | fixes a :: "real^'n" | 
| 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 260 | shows "x \<in> box a b \<longleftrightarrow> (\<forall>i. a$i < x$i \<and> x$i < b$i)" | 
| 56188 | 261 | and "x \<in> cbox a b \<longleftrightarrow> (\<forall>i. a$i \<le> x$i \<and> x$i \<le> b$i)" | 
| 49644 | 262 | using interval_cart[of a b] by (auto simp add: set_eq_iff less_vec_def less_eq_vec_def) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 263 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 264 | lemma interval_eq_empty_cart: | 
| 49644 | 265 | fixes a :: "real^'n" | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 266 |   shows "(box a b = {} \<longleftrightarrow> (\<exists>i. b$i \<le> a$i))" (is ?th1)
 | 
| 56188 | 267 |     and "(cbox a b = {} \<longleftrightarrow> (\<exists>i. b$i < a$i))" (is ?th2)
 | 
| 49644 | 268 | proof - | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 269 |   { fix i x assume as:"b$i \<le> a$i" and x:"x\<in>box a b"
 | 
| 67673 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 270 | hence "a $ i < x $ i \<and> x $ i < b $ i" unfolding mem_box_cart by auto | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 271 | hence "a$i < b$i" by auto | 
| 49644 | 272 | hence False using as by auto } | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 273 | moreover | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 274 |   { assume as:"\<forall>i. \<not> (b$i \<le> a$i)"
 | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 275 | let ?x = "(1/2) *\<^sub>R (a + b)" | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 276 |     { fix i
 | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 277 | have "a$i < b$i" using as[THEN spec[where x=i]] by auto | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 278 | hence "a$i < ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i < b$i" | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 279 | unfolding vector_smult_component and vector_add_component | 
| 49644 | 280 | by auto } | 
| 67673 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 281 |     hence "box a b \<noteq> {}" using mem_box_cart(1)[of "?x" a b] by auto }
 | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 282 | ultimately show ?th1 by blast | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 283 | |
| 56188 | 284 |   { fix i x assume as:"b$i < a$i" and x:"x\<in>cbox a b"
 | 
| 67673 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 285 | hence "a $ i \<le> x $ i \<and> x $ i \<le> b $ i" unfolding mem_box_cart by auto | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 286 | hence "a$i \<le> b$i" by auto | 
| 49644 | 287 | hence False using as by auto } | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 288 | moreover | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 289 |   { assume as:"\<forall>i. \<not> (b$i < a$i)"
 | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 290 | let ?x = "(1/2) *\<^sub>R (a + b)" | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 291 |     { fix i
 | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 292 | have "a$i \<le> b$i" using as[THEN spec[where x=i]] by auto | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 293 | hence "a$i \<le> ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i \<le> b$i" | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 294 | unfolding vector_smult_component and vector_add_component | 
| 49644 | 295 | by auto } | 
| 67673 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 296 |     hence "cbox a b \<noteq> {}" using mem_box_cart(2)[of "?x" a b] by auto  }
 | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 297 | ultimately show ?th2 by blast | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 298 | qed | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 299 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 300 | lemma interval_ne_empty_cart: | 
| 49644 | 301 | fixes a :: "real^'n" | 
| 56188 | 302 |   shows "cbox a b \<noteq> {} \<longleftrightarrow> (\<forall>i. a$i \<le> b$i)"
 | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 303 |     and "box a b \<noteq> {} \<longleftrightarrow> (\<forall>i. a$i < b$i)"
 | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 304 | unfolding interval_eq_empty_cart[of a b] by (auto simp add: not_less not_le) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 305 | (* BH: Why doesn't just "auto" work here? *) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 306 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 307 | lemma subset_interval_imp_cart: | 
| 49644 | 308 | fixes a :: "real^'n" | 
| 56188 | 309 | shows "(\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> cbox c d \<subseteq> cbox a b" | 
| 310 | and "(\<forall>i. a$i < c$i \<and> d$i < b$i) \<Longrightarrow> cbox c d \<subseteq> box a b" | |
| 311 | and "(\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> box c d \<subseteq> cbox a b" | |
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 312 | and "(\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> box c d \<subseteq> box a b" | 
| 67673 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 313 | unfolding subset_eq[unfolded Ball_def] unfolding mem_box_cart | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 314 | by (auto intro: order_trans less_le_trans le_less_trans less_imp_le) (* BH: Why doesn't just "auto" work here? *) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 315 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 316 | lemma interval_sing: | 
| 49644 | 317 | fixes a :: "'a::linorder^'n" | 
| 318 |   shows "{a .. a} = {a} \<and> {a<..<a} = {}"
 | |
| 319 | apply (auto simp add: set_eq_iff less_vec_def less_eq_vec_def vec_eq_iff) | |
| 320 | done | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 321 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 322 | lemma subset_interval_cart: | 
| 49644 | 323 | fixes a :: "real^'n" | 
| 56188 | 324 | shows "cbox c d \<subseteq> cbox a b \<longleftrightarrow> (\<forall>i. c$i \<le> d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th1) | 
| 325 | and "cbox c d \<subseteq> box a b \<longleftrightarrow> (\<forall>i. c$i \<le> d$i) --> (\<forall>i. a$i < c$i \<and> d$i < b$i)" (is ?th2) | |
| 326 | and "box c d \<subseteq> cbox a b \<longleftrightarrow> (\<forall>i. c$i < d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th3) | |
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 327 | and "box c d \<subseteq> box a b \<longleftrightarrow> (\<forall>i. c$i < d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th4) | 
| 56188 | 328 | using subset_box[of c d a b] by (simp_all add: Basis_vec_def inner_axis) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 329 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 330 | lemma disjoint_interval_cart: | 
| 49644 | 331 | fixes a::"real^'n" | 
| 56188 | 332 |   shows "cbox a b \<inter> cbox c d = {} \<longleftrightarrow> (\<exists>i. (b$i < a$i \<or> d$i < c$i \<or> b$i < c$i \<or> d$i < a$i))" (is ?th1)
 | 
| 333 |     and "cbox a b \<inter> box c d = {} \<longleftrightarrow> (\<exists>i. (b$i < a$i \<or> d$i \<le> c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th2)
 | |
| 334 |     and "box a b \<inter> cbox c d = {} \<longleftrightarrow> (\<exists>i. (b$i \<le> a$i \<or> d$i < c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th3)
 | |
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 335 |     and "box a b \<inter> box c d = {} \<longleftrightarrow> (\<exists>i. (b$i \<le> a$i \<or> d$i \<le> c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th4)
 | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49962diff
changeset | 336 | using disjoint_interval[of a b c d] by (simp_all add: Basis_vec_def inner_axis) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 337 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 338 | lemma Int_interval_cart: | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 339 | fixes a :: "real^'n" | 
| 56188 | 340 |   shows "cbox a b \<inter> cbox c d =  {(\<chi> i. max (a$i) (c$i)) .. (\<chi> i. min (b$i) (d$i))}"
 | 
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63938diff
changeset | 341 | unfolding Int_interval | 
| 56188 | 342 | by (auto simp: mem_box less_eq_vec_def) | 
| 343 | (auto simp: Basis_vec_def inner_axis) | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 344 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 345 | lemma closed_interval_left_cart: | 
| 49644 | 346 | fixes b :: "real^'n" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 347 |   shows "closed {x::real^'n. \<forall>i. x$i \<le> b$i}"
 | 
| 71172 | 348 | by (simp add: Collect_all_eq closed_INT closed_Collect_le continuous_on_component) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 349 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 350 | lemma closed_interval_right_cart: | 
| 49644 | 351 | fixes a::"real^'n" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 352 |   shows "closed {x::real^'n. \<forall>i. a$i \<le> x$i}"
 | 
| 71172 | 353 | by (simp add: Collect_all_eq closed_INT closed_Collect_le continuous_on_component) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 354 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 355 | lemma is_interval_cart: | 
| 49644 | 356 | "is_interval (s::(real^'n) set) \<longleftrightarrow> | 
| 357 | (\<forall>a\<in>s. \<forall>b\<in>s. \<forall>x. (\<forall>i. ((a$i \<le> x$i \<and> x$i \<le> b$i) \<or> (b$i \<le> x$i \<and> x$i \<le> a$i))) \<longrightarrow> x \<in> s)" | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49962diff
changeset | 358 | by (simp add: is_interval_def Ball_def Basis_vec_def inner_axis imp_ex) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 359 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 360 | lemma closed_halfspace_component_le_cart: "closed {x::real^'n. x$i \<le> a}"
 | 
| 69669 | 361 | by (simp add: closed_Collect_le continuous_on_component) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 362 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 363 | lemma closed_halfspace_component_ge_cart: "closed {x::real^'n. x$i \<ge> a}"
 | 
| 69669 | 364 | by (simp add: closed_Collect_le continuous_on_component) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 365 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 366 | lemma open_halfspace_component_lt_cart: "open {x::real^'n. x$i < a}"
 | 
| 69669 | 367 | by (simp add: open_Collect_less continuous_on_component) | 
| 49644 | 368 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 369 | lemma open_halfspace_component_gt_cart: "open {x::real^'n. x$i  > a}"
 | 
| 69669 | 370 | by (simp add: open_Collect_less continuous_on_component) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 371 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 372 | lemma Lim_component_le_cart: | 
| 49644 | 373 | fixes f :: "'a \<Rightarrow> real^'n" | 
| 61973 | 374 | assumes "(f \<longlongrightarrow> l) net" "\<not> (trivial_limit net)" "eventually (\<lambda>x. f x $i \<le> b) net" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 375 | shows "l$i \<le> b" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49962diff
changeset | 376 | by (rule tendsto_le[OF assms(2) tendsto_const tendsto_vec_nth, OF assms(1, 3)]) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 377 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 378 | lemma Lim_component_ge_cart: | 
| 49644 | 379 | fixes f :: "'a \<Rightarrow> real^'n" | 
| 61973 | 380 | assumes "(f \<longlongrightarrow> l) net" "\<not> (trivial_limit net)" "eventually (\<lambda>x. b \<le> (f x)$i) net" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 381 | shows "b \<le> l$i" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49962diff
changeset | 382 | by (rule tendsto_le[OF assms(2) tendsto_vec_nth tendsto_const, OF assms(1, 3)]) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 383 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 384 | lemma Lim_component_eq_cart: | 
| 49644 | 385 | fixes f :: "'a \<Rightarrow> real^'n" | 
| 69508 | 386 | assumes net: "(f \<longlongrightarrow> l) net" "\<not> trivial_limit net" and ev:"eventually (\<lambda>x. f(x)$i = b) net" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 387 | shows "l$i = b" | 
| 49644 | 388 | using ev[unfolded order_eq_iff eventually_conj_iff] and | 
| 389 | Lim_component_ge_cart[OF net, of b i] and | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 390 | Lim_component_le_cart[OF net, of i b] by auto | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 391 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 392 | lemma connected_ivt_component_cart: | 
| 49644 | 393 | fixes x :: "real^'n" | 
| 394 | shows "connected s \<Longrightarrow> x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> x$k \<le> a \<Longrightarrow> a \<le> y$k \<Longrightarrow> (\<exists>z\<in>s. z$k = a)" | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49962diff
changeset | 395 | using connected_ivt_hyperplane[of s x y "axis k 1" a] | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49962diff
changeset | 396 | by (auto simp add: inner_axis inner_commute) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 397 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 398 | lemma subspace_substandard_cart: "vec.subspace {x. (\<forall>i. P i \<longrightarrow> x$i = 0)}"
 | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 399 | unfolding vec.subspace_def by auto | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 400 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 401 | lemma closed_substandard_cart: | 
| 44213 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 huffman parents: 
44211diff
changeset | 402 |   "closed {x::'a::real_normed_vector ^ 'n. \<forall>i. P i \<longrightarrow> x$i = 0}"
 | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 403 | proof - | 
| 44213 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 huffman parents: 
44211diff
changeset | 404 |   { fix i::'n
 | 
| 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 huffman parents: 
44211diff
changeset | 405 |     have "closed {x::'a ^ 'n. P i \<longrightarrow> x$i = 0}"
 | 
| 71172 | 406 | by (cases "P i") (simp_all add: closed_Collect_eq continuous_on_component) } | 
| 44213 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 huffman parents: 
44211diff
changeset | 407 | thus ?thesis | 
| 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 huffman parents: 
44211diff
changeset | 408 | unfolding Collect_all_eq by (simp add: closed_INT) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 409 | qed | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 410 | |
| 69683 | 411 | subsection "Convex Euclidean Space" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 412 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 413 | lemma Cart_1:"(1::real^'n) = \<Sum>Basis" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49962diff
changeset | 414 | using const_vector_cart[of 1] by (simp add: one_vec_def) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 415 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 416 | declare vector_add_ldistrib[simp] vector_ssub_ldistrib[simp] vector_smult_assoc[simp] vector_smult_rneg[simp] | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 417 | declare vector_sadd_rdistrib[simp] vector_sub_rdistrib[simp] | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 418 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 419 | lemmas vector_component_simps = vector_minus_component vector_smult_component vector_add_component less_eq_vec_def vec_lambda_beta vector_uminus_component | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 420 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 421 | lemma convex_box_cart: | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 422 |   assumes "\<And>i. convex {x. P i x}"
 | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 423 |   shows "convex {x. \<forall>i. P i (x$i)}"
 | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 424 | using assms unfolding convex_def by auto | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 425 | |
| 71175 
a1e94be66bd6
reduced imports; deleted unusewd minor lemmas for that purpose
 nipkow parents: 
71172diff
changeset | 426 | (* Unused | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 427 | lemma convex_positive_orthant_cart: "convex {x::real^'n. (\<forall>i. 0 \<le> x$i)}"
 | 
| 63334 
bd37a72a940a
Multivariate_Analysis: add continuous_on_vec_lambda
 hoelzl parents: 
63332diff
changeset | 428 | by (rule convex_box_cart) (simp add: atLeast_def[symmetric]) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 429 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 430 | lemma unit_interval_convex_hull_cart: | 
| 56188 | 431 |   "cbox (0::real^'n) 1 = convex hull {x. \<forall>i. (x$i = 0) \<or> (x$i = 1)}"
 | 
| 432 | unfolding Cart_1 unit_interval_convex_hull[where 'a="real^'n"] box_real[symmetric] | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49962diff
changeset | 433 | by (rule arg_cong[where f="\<lambda>x. convex hull x"]) (simp add: Basis_vec_def inner_axis) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 434 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 435 | proposition cube_convex_hull_cart: | 
| 49644 | 436 | assumes "0 < d" | 
| 437 | obtains s::"(real^'n) set" | |
| 56188 | 438 | where "finite s" "cbox (x - (\<chi> i. d)) (x + (\<chi> i. d)) = convex hull s" | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 439 | proof - | 
| 55522 | 440 | from assms obtain s where "finite s" | 
| 71175 
a1e94be66bd6
reduced imports; deleted unusewd minor lemmas for that purpose
 nipkow parents: 
71172diff
changeset | 441 | and "cbox (x - sum (( *\<^sub>R) d) Basis) (x + sum (( *\<^sub>R) d) Basis) = convex hull s" | 
| 55522 | 442 | by (rule cube_convex_hull) | 
| 443 | with that[of s] show thesis | |
| 444 | by (simp add: const_vector_cart) | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 445 | qed | 
| 71175 
a1e94be66bd6
reduced imports; deleted unusewd minor lemmas for that purpose
 nipkow parents: 
71172diff
changeset | 446 | *) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 447 | |
| 77943 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 448 | subsection\<open>Arbitrarily good rational approximations\<close> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 449 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 450 | lemma rational_approximation: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 451 | assumes "e > 0" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 452 | obtains r::real where "r \<in> \<rat>" "\<bar>r - x\<bar> < e" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 453 | using Rats_dense_in_real [of "x - e/2" "x + e/2"] assms by auto | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 454 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 455 | lemma Rats_closure_real: "closure \<rat> = (UNIV::real set)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 456 | proof - | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 457 | have "\<And>x::real. x \<in> closure \<rat>" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 458 | by (metis closure_approachable dist_real_def rational_approximation) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 459 | then show ?thesis by auto | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 460 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 461 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 462 | proposition matrix_rational_approximation: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 463 | fixes A :: "real^'n^'m" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 464 | assumes "e > 0" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 465 | obtains B where "\<And>i j. B$i$j \<in> \<rat>" "onorm(\<lambda>x. (A - B) *v x) < e" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 466 | proof - | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 467 |   have "\<forall>i j. \<exists>q \<in> \<rat>. \<bar>q - A $ i $ j\<bar> < e / (2 * CARD('m) * CARD('n))"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 468 |     using assms by (force intro: rational_approximation [of "e / (2 * CARD('m) * CARD('n))"])
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 469 |   then obtain B where B: "\<And>i j. B$i$j \<in> \<rat>" and Bclo: "\<And>i j. \<bar>B$i$j - A $ i $ j\<bar> < e / (2 * CARD('m) * CARD('n))"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 470 | by (auto simp: lambda_skolem Bex_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 471 | show ?thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 472 | proof | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 473 |     have "onorm ((*v) (A - B)) \<le> real CARD('m) * real CARD('n) *
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 474 |     (e / (2 * real CARD('m) * real CARD('n)))"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 475 | apply (rule onorm_le_matrix_component) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 476 | using Bclo by (simp add: abs_minus_commute less_imp_le) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 477 | also have "\<dots> < e" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 478 | using \<open>0 < e\<close> by (simp add: field_split_simps) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 479 | finally show "onorm ((*v) (A - B)) < e" . | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 480 | qed (use B in auto) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 481 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 482 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 483 | |
| 69683 | 484 | subsection "Derivative" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 485 | |
| 70136 | 486 | definition\<^marker>\<open>tag important\<close> "jacobian f net = matrix(frechet_derivative f net)" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 487 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 488 | proposition jacobian_works: | 
| 49644 | 489 | "(f::(real^'a) \<Rightarrow> (real^'b)) differentiable net \<longleftrightarrow> | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 490 | (f has_derivative (\<lambda>h. (jacobian f net) *v h)) net" (is "?lhs = ?rhs") | 
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 491 | proof | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 492 | assume ?lhs then show ?rhs | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 493 | by (simp add: frechet_derivative_works has_derivative_linear jacobian_def) | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 494 | next | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 495 | assume ?rhs then show ?lhs | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 496 | by (rule differentiableI) | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 497 | qed | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 498 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 499 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 500 | text \<open>Component of the differential must be zero if it exists at a local | 
| 67968 | 501 | maximum or minimum for that corresponding component\<close> | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 502 | |
| 69723 
9b9f203e0ba3
tagged 2 theories  ie Cartesian_Euclidean_Space Cartesian_Space
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69686diff
changeset | 503 | proposition differential_zero_maxmin_cart: | 
| 49644 | 504 | fixes f::"real^'a \<Rightarrow> real^'b" | 
| 505 | assumes "0 < e" "((\<forall>y \<in> ball x e. (f y)$k \<le> (f x)$k) \<or> (\<forall>y\<in>ball x e. (f x)$k \<le> (f y)$k))" | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49962diff
changeset | 506 | "f differentiable (at x)" | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49962diff
changeset | 507 | shows "jacobian f (at x) $ k = 0" | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49962diff
changeset | 508 | using differential_zero_maxmin_component[of "axis k 1" e x f] assms | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49962diff
changeset | 509 | vector_cart[of "\<lambda>j. frechet_derivative f (at x) j $ k"] | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49962diff
changeset | 510 | by (simp add: Basis_vec_def axis_eq_axis inner_axis jacobian_def matrix_def) | 
| 49644 | 511 | |
| 70136 | 512 | subsection\<^marker>\<open>tag unimportant\<close>\<open>Routine results connecting the types \<^typ>\<open>real^1\<close> and \<^typ>\<open>real\<close>\<close> | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 513 | |
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 514 | lemma vec_cbox_1_eq [simp]: | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 515 | shows "vec ` cbox u v = cbox (vec u) (vec v ::real^1)" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 516 | by (force simp: Basis_vec_def cart_eq_inner_axis [symmetric] mem_box) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 517 | |
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 518 | lemma vec_nth_cbox_1_eq [simp]: | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 519 | fixes u v :: "'a::euclidean_space^1" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 520 | shows "(\<lambda>x. x $ 1) ` cbox u v = cbox (u$1) (v$1)" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 521 | by (auto simp: Basis_vec_def cart_eq_inner_axis [symmetric] mem_box image_iff Bex_def inner_axis) (metis vec_component) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 522 | |
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 523 | lemma vec_nth_1_iff_cbox [simp]: | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 524 | fixes a b :: "'a::euclidean_space" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 525 | shows "(\<lambda>x::'a^1. x $ 1) ` S = cbox a b \<longleftrightarrow> S = cbox (vec a) (vec b)" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 526 | (is "?lhs = ?rhs") | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 527 | proof | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 528 | assume L: ?lhs show ?rhs | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 529 | proof (intro equalityI subsetI) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 530 | fix x | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 531 | assume "x \<in> S" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 532 | then have "x $ 1 \<in> (\<lambda>v. v $ (1::1)) ` cbox (vec a) (vec b)" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 533 | using L by auto | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 534 | then show "x \<in> cbox (vec a) (vec b)" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 535 | by (metis (no_types, lifting) imageE vector_one_nth) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 536 | next | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 537 | fix x :: "'a^1" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 538 | assume "x \<in> cbox (vec a) (vec b)" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 539 | then show "x \<in> S" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 540 | by (metis (no_types, lifting) L imageE imageI vec_component vec_nth_cbox_1_eq vector_one_nth) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 541 | qed | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 542 | qed simp | 
| 49644 | 543 | |
| 71025 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70817diff
changeset | 544 | lemma vec_nth_real_1_iff_cbox [simp]: | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70817diff
changeset | 545 | fixes a b :: real | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70817diff
changeset | 546 |   shows "(\<lambda>x::real^1. x $ 1) ` S = {a..b} \<longleftrightarrow> S = cbox (vec a) (vec b)"
 | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70817diff
changeset | 547 | using vec_nth_1_iff_cbox[of S a b] | 
| 
be8cec1abcbb
reduce dependencies of Ordered_Euclidean_Space; move more general material from Cartesian_Euclidean_Space
 immler parents: 
70817diff
changeset | 548 | by simp | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 549 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 550 | lemma interval_split_cart: | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 551 |   "{a..b::real^'n} \<inter> {x. x$k \<le> c} = {a .. (\<chi> i. if i = k then min (b$k) c else b$i)}"
 | 
| 56188 | 552 |   "cbox a b \<inter> {x. x$k \<ge> c} = {(\<chi> i. if i = k then max (a$k) c else a$i) .. b}"
 | 
| 77943 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
75078diff
changeset | 553 | unfolding Int_iff mem_box_cart mem_Collect_eq interval_cbox_cart set_eq_iff | 
| 49644 | 554 | unfolding vec_lambda_beta | 
| 555 | by auto | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 556 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 557 | lemmas cartesian_euclidean_space_uniform_limit_intros[uniform_limit_intros] = | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 558 | bounded_linear.uniform_limit[OF blinfun.bounded_linear_right] | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 559 | bounded_linear.uniform_limit[OF bounded_linear_vec_nth] | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 560 | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 561 | end |