| author | wenzelm | 
| Sat, 08 Sep 2018 22:43:25 +0200 | |
| changeset 68955 | 0851db8cde12 | 
| parent 68833 | fde093888c16 | 
| child 69064 | 5840724b1d71 | 
| 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 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 5 | section%important \<open>Instantiates the finite Cartesian product of Euclidean spaces as a Euclidean space\<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 | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 8 | imports Cartesian_Space 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 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 11 | lemma%unimportant 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 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 14 | lemma%important sum_mult_product: | 
| 64267 | 15 |   "sum h {..<A * B :: nat} = (\<Sum>i\<in>{..<A}. \<Sum>j\<in>{..<B}. h (j + i * B))"
 | 
| 16 | unfolding sum_nat_group[of h B A, unfolded atLeast0LessThan, symmetric] | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 17 | proof%unimportant (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 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 28 | lemma%unimportant 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 | ||
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 31 | lemma%unimportant 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 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 36 | lemma%unimportant 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 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 43 | lemma%unimportant 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 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 48 | subsection%important\<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 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 50 | lemma%important interior_halfspace_le [simp]: | 
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 51 | assumes "a \<noteq> 0" | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 52 |     shows "interior {x. a \<bullet> x \<le> b} = {x. a \<bullet> x < b}"
 | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 53 | proof%unimportant - | 
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 54 |   have *: "a \<bullet> x < b" if x: "x \<in> S" and S: "S \<subseteq> {x. a \<bullet> x \<le> b}" and "open S" for S x
 | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 55 | proof - | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 56 | obtain e where "e>0" and e: "cball x e \<subseteq> S" | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 57 | using \<open>open S\<close> open_contains_cball x by blast | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 58 | then have "x + (e / norm a) *\<^sub>R a \<in> cball x e" | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 59 | by (simp add: dist_norm) | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 60 | then have "x + (e / norm a) *\<^sub>R a \<in> S" | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 61 | using e by blast | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 62 |     then have "x + (e / norm a) *\<^sub>R a \<in> {x. a \<bullet> x \<le> b}"
 | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 63 | using S by blast | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 64 | moreover have "e * (a \<bullet> a) / norm a > 0" | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 65 | by (simp add: \<open>0 < e\<close> assms) | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 66 | ultimately show ?thesis | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 67 | by (simp add: algebra_simps) | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 68 | qed | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 69 | show ?thesis | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 70 | by (rule interior_unique) (auto simp: open_halfspace_lt *) | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 71 | qed | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 72 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 73 | lemma%unimportant interior_halfspace_ge [simp]: | 
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 74 |    "a \<noteq> 0 \<Longrightarrow> interior {x. a \<bullet> x \<ge> b} = {x. a \<bullet> x > b}"
 | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 75 | using interior_halfspace_le [of "-a" "-b"] by simp | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 76 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 77 | lemma%important interior_halfspace_component_le [simp]: | 
| 67731 | 78 |      "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 | 79 | and interior_halfspace_component_ge [simp]: | 
| 67731 | 80 |      "interior {x. x$k \<ge> a} = {x :: (real^'n). x$k > a}" (is "?GE")
 | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 81 | proof%unimportant - | 
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 82 | 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 | 83 | 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 | 84 | 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 | 85 | 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 | 86 | ultimately show ?LE ?GE | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 87 | 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 | 88 | 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 | 89 | qed | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 90 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 91 | lemma%unimportant closure_halfspace_lt [simp]: | 
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 92 | assumes "a \<noteq> 0" | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 93 |     shows "closure {x. a \<bullet> x < b} = {x. a \<bullet> x \<le> b}"
 | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 94 | proof - | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 95 |   have [simp]: "-{x. a \<bullet> x < b} = {x. a \<bullet> x \<ge> b}"
 | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 96 | by (force simp:) | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 97 | then show ?thesis | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 98 | using interior_halfspace_ge [of a b] assms | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 99 | by (force simp: closure_interior) | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 100 | qed | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 101 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 102 | lemma%unimportant closure_halfspace_gt [simp]: | 
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 103 |    "a \<noteq> 0 \<Longrightarrow> closure {x. a \<bullet> x > b} = {x. a \<bullet> x \<ge> b}"
 | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 104 | using closure_halfspace_lt [of "-a" "-b"] by simp | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 105 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 106 | lemma%important closure_halfspace_component_lt [simp]: | 
| 67731 | 107 |      "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 | 108 | and closure_halfspace_component_gt [simp]: | 
| 67731 | 109 |      "closure {x. x$k > a} = {x :: (real^'n). x$k \<ge> a}" (is "?GE")
 | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 110 | proof%unimportant - | 
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 111 | 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 | 112 | 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 | 113 | 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 | 114 | 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 | 115 | ultimately show ?LE ?GE | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 116 | 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 | 117 | 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 | 118 | qed | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 119 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 120 | lemma%unimportant interior_hyperplane [simp]: | 
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 121 | assumes "a \<noteq> 0" | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 122 |     shows "interior {x. a \<bullet> x = b} = {}"
 | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 123 | proof%unimportant - | 
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 124 |   have [simp]: "{x. a \<bullet> x = b} = {x. a \<bullet> x \<le> b} \<inter> {x. a \<bullet> x \<ge> b}"
 | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 125 | by (force simp:) | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 126 | then show ?thesis | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 127 | by (auto simp: assms) | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 128 | qed | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 129 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 130 | lemma%unimportant frontier_halfspace_le: | 
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 131 | assumes "a \<noteq> 0 \<or> b \<noteq> 0" | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 132 |     shows "frontier {x. a \<bullet> x \<le> b} = {x. a \<bullet> x = b}"
 | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 133 | proof (cases "a = 0") | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 134 | case True with assms show ?thesis by simp | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 135 | next | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 136 | case False then show ?thesis | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 137 | by (force simp: frontier_def closed_halfspace_le) | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 138 | qed | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 139 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 140 | lemma%unimportant frontier_halfspace_ge: | 
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 141 | assumes "a \<noteq> 0 \<or> b \<noteq> 0" | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 142 |     shows "frontier {x. a \<bullet> x \<ge> b} = {x. a \<bullet> x = b}"
 | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 143 | proof (cases "a = 0") | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 144 | case True with assms show ?thesis by simp | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 145 | next | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 146 | case False then show ?thesis | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 147 | by (force simp: frontier_def closed_halfspace_ge) | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 148 | qed | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 149 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 150 | lemma%unimportant frontier_halfspace_lt: | 
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 151 | assumes "a \<noteq> 0 \<or> b \<noteq> 0" | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 152 |     shows "frontier {x. a \<bullet> x < b} = {x. a \<bullet> x = b}"
 | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 153 | proof (cases "a = 0") | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 154 | case True with assms show ?thesis by simp | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 155 | next | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 156 | case False then show ?thesis | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 157 | by (force simp: frontier_def interior_open open_halfspace_lt) | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 158 | qed | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 159 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 160 | lemma%important frontier_halfspace_gt: | 
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 161 | assumes "a \<noteq> 0 \<or> b \<noteq> 0" | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 162 |     shows "frontier {x. a \<bullet> x > b} = {x. a \<bullet> x = b}"
 | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 163 | proof%unimportant (cases "a = 0") | 
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 164 | case True with assms show ?thesis by simp | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 165 | next | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 166 | case False then show ?thesis | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 167 | by (force simp: frontier_def interior_open open_halfspace_gt) | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 168 | qed | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 169 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 170 | lemma%important interior_standard_hyperplane: | 
| 67731 | 171 |    "interior {x :: (real^'n). x$k = a} = {}"
 | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 172 | proof%unimportant - | 
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 173 | 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 | 174 | 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 | 175 | 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 | 176 | 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 | 177 | ultimately show ?thesis | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 178 | 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 | 179 | by force | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 180 | qed | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62127diff
changeset | 181 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 182 | lemma%unimportant matrix_mult_transpose_dot_column: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 183 | shows "transpose A ** A = (\<chi> i j. inner (column i A) (column j A))" | 
| 67673 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 184 | by (simp add: matrix_matrix_mult_def vec_eq_iff transpose_def column_def inner_vec_def) | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 185 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 186 | lemma%unimportant matrix_mult_transpose_dot_row: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 187 | shows "A ** transpose A = (\<chi> i j. inner (row i A) (row j A))" | 
| 67673 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 188 | by (simp add: matrix_matrix_mult_def vec_eq_iff transpose_def row_def inner_vec_def) | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 189 | |
| 60420 | 190 | text\<open>Two sometimes fruitful ways of looking at matrix-vector multiplication.\<close> | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 191 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 192 | lemma%important matrix_mult_dot: "A *v x = (\<chi> i. inner (A$i) x)" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 193 | by (simp add: matrix_vector_mult_def inner_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 | 194 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 195 | lemma%unimportant adjoint_matrix: "adjoint(\<lambda>x. (A::real^'n^'m) *v x) = (\<lambda>x. transpose A *v x)" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 196 | apply (rule adjoint_unique) | 
| 49644 | 197 | apply (simp add: transpose_def inner_vec_def matrix_vector_mult_def | 
| 64267 | 198 | sum_distrib_right sum_distrib_left) | 
| 66804 
3f9bb52082c4
avoid name clashes on interpretation of abstract locales
 haftmann parents: 
66447diff
changeset | 199 | apply (subst sum.swap) | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 200 | apply (simp add: ac_simps) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 201 | done | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 202 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 203 | lemma%important matrix_adjoint: assumes lf: "linear (f :: real^'n \<Rightarrow> real ^'m)" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 204 | shows "matrix(adjoint f) = transpose(matrix f)" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 205 | proof%unimportant - | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 206 | have "matrix(adjoint f) = matrix(adjoint (( *v) (matrix f)))" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 207 | by (simp add: lf) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 208 | also have "\<dots> = transpose(matrix f)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 209 | unfolding adjoint_matrix matrix_of_matrix_vector_mul | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 210 | apply rule | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 211 | done | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 212 | finally show ?thesis . | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 213 | qed | 
| 49644 | 214 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 215 | lemma%unimportant 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 | 216 | 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 | 217 | 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 | 218 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 219 | lemma%unimportant (* FIX ME needs name*) | 
| 68073 
fad29d2a17a5
merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
 immlerdiff
changeset | 220 |   fixes 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 | 221 | shows matrix_vector_mult_linear_continuous_at [continuous_intros]: "isCont (( *v) A) z" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 222 | and matrix_vector_mult_linear_continuous_on [continuous_intros]: "continuous_on S (( *v) A)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 223 | 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 | 224 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 225 | lemma%unimportant scalar_invertible: | 
| 68050 
7eacc812ad1c
minor typeclass generalisations and junk removal
 paulson <lp15@cam.ac.uk> parents: 
68045diff
changeset | 226 |   fixes A :: "('a::real_algebra_1)^'m^'n"
 | 
| 68038 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 227 | assumes "k \<noteq> 0" and "invertible A" | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 228 | shows "invertible (k *\<^sub>R A)" | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 229 | proof - | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 230 | obtain A' where "A ** A' = mat 1" and "A' ** A = mat 1" | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 231 | using assms unfolding invertible_def by auto | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 232 | with `k \<noteq> 0` | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 233 | have "(k *\<^sub>R A) ** ((1/k) *\<^sub>R A') = mat 1" "((1/k) *\<^sub>R A') ** (k *\<^sub>R A) = mat 1" | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 234 | by (simp_all add: assms matrix_scalar_ac) | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 235 | thus "invertible (k *\<^sub>R A)" | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 236 | unfolding invertible_def by auto | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 237 | qed | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 238 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 239 | lemma%unimportant scalar_invertible_iff: | 
| 68050 
7eacc812ad1c
minor typeclass generalisations and junk removal
 paulson <lp15@cam.ac.uk> parents: 
68045diff
changeset | 240 |   fixes A :: "('a::real_algebra_1)^'m^'n"
 | 
| 68038 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 241 | assumes "k \<noteq> 0" and "invertible A" | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 242 | shows "invertible (k *\<^sub>R A) \<longleftrightarrow> k \<noteq> 0 \<and> invertible A" | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 243 | by (simp add: assms scalar_invertible) | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 244 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 245 | lemma%unimportant vector_transpose_matrix [simp]: "x v* transpose A = A *v x" | 
| 68038 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 246 | unfolding transpose_def vector_matrix_mult_def matrix_vector_mult_def | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 247 | by simp | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 248 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 249 | lemma%unimportant transpose_matrix_vector [simp]: "transpose A *v x = x v* A" | 
| 68038 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 250 | unfolding transpose_def vector_matrix_mult_def matrix_vector_mult_def | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 251 | by simp | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 252 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 253 | lemma%unimportant vector_scalar_commute: | 
| 68043 
d345e9c35ae1
some of Jose Divasón's material from Rank_Nullity_Theorem/Miscellaneous
 paulson <lp15@cam.ac.uk> parents: 
68041diff
changeset | 254 |   fixes A :: "'a::{field}^'m^'n"
 | 
| 
d345e9c35ae1
some of Jose Divasón's material from Rank_Nullity_Theorem/Miscellaneous
 paulson <lp15@cam.ac.uk> parents: 
68041diff
changeset | 255 | shows "A *v (c *s x) = c *s (A *v x)" | 
| 
d345e9c35ae1
some of Jose Divasón's material from Rank_Nullity_Theorem/Miscellaneous
 paulson <lp15@cam.ac.uk> parents: 
68041diff
changeset | 256 | by (simp add: vector_scalar_mult_def matrix_vector_mult_def mult_ac sum_distrib_left) | 
| 
d345e9c35ae1
some of Jose Divasón's material from Rank_Nullity_Theorem/Miscellaneous
 paulson <lp15@cam.ac.uk> parents: 
68041diff
changeset | 257 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 258 | lemma%unimportant scalar_vector_matrix_assoc: | 
| 68043 
d345e9c35ae1
some of Jose Divasón's material from Rank_Nullity_Theorem/Miscellaneous
 paulson <lp15@cam.ac.uk> parents: 
68041diff
changeset | 259 |   fixes k :: "'a::{field}" and x :: "'a::{field}^'n" and A :: "'a^'m^'n"
 | 
| 
d345e9c35ae1
some of Jose Divasón's material from Rank_Nullity_Theorem/Miscellaneous
 paulson <lp15@cam.ac.uk> parents: 
68041diff
changeset | 260 | shows "(k *s x) v* A = k *s (x v* A)" | 
| 
d345e9c35ae1
some of Jose Divasón's material from Rank_Nullity_Theorem/Miscellaneous
 paulson <lp15@cam.ac.uk> parents: 
68041diff
changeset | 261 | by (metis transpose_matrix_vector vector_scalar_commute) | 
| 
d345e9c35ae1
some of Jose Divasón's material from Rank_Nullity_Theorem/Miscellaneous
 paulson <lp15@cam.ac.uk> parents: 
68041diff
changeset | 262 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 263 | lemma%unimportant vector_matrix_mult_0 [simp]: "0 v* A = 0" | 
| 68043 
d345e9c35ae1
some of Jose Divasón's material from Rank_Nullity_Theorem/Miscellaneous
 paulson <lp15@cam.ac.uk> parents: 
68041diff
changeset | 264 | unfolding vector_matrix_mult_def by (simp add: zero_vec_def) | 
| 
d345e9c35ae1
some of Jose Divasón's material from Rank_Nullity_Theorem/Miscellaneous
 paulson <lp15@cam.ac.uk> parents: 
68041diff
changeset | 265 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 266 | lemma%unimportant vector_matrix_mult_0_right [simp]: "x v* 0 = 0" | 
| 68043 
d345e9c35ae1
some of Jose Divasón's material from Rank_Nullity_Theorem/Miscellaneous
 paulson <lp15@cam.ac.uk> parents: 
68041diff
changeset | 267 | unfolding vector_matrix_mult_def by (simp add: zero_vec_def) | 
| 
d345e9c35ae1
some of Jose Divasón's material from Rank_Nullity_Theorem/Miscellaneous
 paulson <lp15@cam.ac.uk> parents: 
68041diff
changeset | 268 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 269 | lemma%unimportant vector_matrix_mul_rid [simp]: | 
| 68038 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 270 |   fixes v :: "('a::semiring_1)^'n"
 | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 271 | shows "v v* mat 1 = v" | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 272 | by (metis matrix_vector_mul_lid transpose_mat vector_transpose_matrix) | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 273 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 274 | lemma%unimportant scaleR_vector_matrix_assoc: | 
| 68038 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 275 | fixes k :: real and x :: "real^'n" and A :: "real^'m^'n" | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 276 | shows "(k *\<^sub>R x) v* A = k *\<^sub>R (x v* A)" | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 277 | by (metis matrix_vector_mult_scaleR transpose_matrix_vector) | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 278 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 279 | lemma%important vector_scaleR_matrix_ac: | 
| 68038 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 280 | fixes k :: real and x :: "real^'n" and A :: "real^'m^'n" | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 281 | shows "x v* (k *\<^sub>R A) = k *\<^sub>R (x v* A)" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 282 | proof%unimportant - | 
| 68038 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 283 | have "x v* (k *\<^sub>R A) = (k *\<^sub>R x) v* A" | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 284 | unfolding vector_matrix_mult_def | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 285 | by (simp add: algebra_simps) | 
| 68043 
d345e9c35ae1
some of Jose Divasón's material from Rank_Nullity_Theorem/Miscellaneous
 paulson <lp15@cam.ac.uk> parents: 
68041diff
changeset | 286 | with scaleR_vector_matrix_assoc | 
| 68038 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 287 | show "x v* (k *\<^sub>R A) = k *\<^sub>R (x v* A)" | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 288 | by auto | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 289 | qed | 
| 
20b713cff87a
new material on matricies by Tim Makarios (from Tarskis_Geometry in the AFP)
 paulson <lp15@cam.ac.uk> parents: 
67986diff
changeset | 290 | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 291 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 292 | subsection%important\<open>Some 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 | 293 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 294 | lemma%important 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 | 295 | 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 | 296 | shows "norm(column i A) \<le> onorm(( *v) A)" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 297 | proof%unimportant - | 
| 67719 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 298 | 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 | 299 | by (simp add: matrix_mult_dot cart_eq_inner_axis) | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 300 | 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 | 301 | using onorm [OF matrix_vector_mul_bounded_linear, of A "axis i 1"] by auto | 
| 67719 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 302 | finally have "norm (\<chi> j. A $ j $ i) \<le> onorm (( *v) A)" . | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 303 | then show ?thesis | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 304 | unfolding column_def . | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 305 | qed | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 306 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 307 | lemma%important 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 | 308 | 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 | 309 | shows "\<bar>A $ i $ j\<bar> \<le> onorm(( *v) A)" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 310 | proof%unimportant - | 
| 67719 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 311 | 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 | 312 | by (metis (full_types, lifting) component_le_norm_cart vec_lambda_beta) | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 313 | also have "\<dots> \<le> onorm (( *v) A)" | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 314 | 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 | 315 | finally show ?thesis . | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 316 | qed | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 317 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 318 | lemma%unimportant 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 | 319 | 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 | 320 | shows "linear f \<Longrightarrow> \<bar>matrix f $ i $ j\<bar> \<le> onorm f" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 321 | by (metis linear_matrix_vector_mul_eq matrix_component_le_onorm matrix_vector_mul) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 322 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 323 | lemma%important 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 | 324 | 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 | 325 | shows "onorm(( *v) A) \<le> (\<Sum>i\<in>UNIV. \<Sum>j\<in>UNIV. \<bar>A $ i $ j\<bar>)" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 326 | proof%unimportant (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 | 327 | fix x | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 328 | 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 | 329 | 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 | 330 | 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 | 331 | 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 | 332 | fix i | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 333 | 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 | 334 | 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 | 335 | 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 | 336 | 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 | 337 | 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 | 338 | 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 | 339 | 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 | 340 | qed | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 341 | 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 | 342 | 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 | 343 | qed | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 344 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 345 | lemma%important 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 | 346 | 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 | 347 | assumes "\<And>i j. abs(A$i$j) \<le> B" | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 348 |   shows "onorm(( *v) A) \<le> real (CARD('m)) * real (CARD('n)) * B"
 | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 349 | proof%unimportant (rule onorm_le) | 
| 67731 | 350 | 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 | 351 | 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 | 352 | 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 | 353 |   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 | 354 | 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 | 355 | fix i | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 356 | 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 | 357 | 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 | 358 | 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 | 359 | 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 | 360 |     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 | 361 | 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 | 362 |     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 | 363 | qed | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 364 |   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 | 365 | by simp | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 366 |   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 | 367 | qed | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 368 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 369 | subsection%important \<open>lambda skolemization on cartesian products\<close> | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 370 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 371 | lemma%important lambda_skolem: "(\<forall>i. \<exists>x. P i x) \<longleftrightarrow> | 
| 37494 | 372 | (\<exists>x::'a ^ 'n. \<forall>i. P i (x $ i))" (is "?lhs \<longleftrightarrow> ?rhs") | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 373 | proof%unimportant - | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 374 | let ?S = "(UNIV :: 'n set)" | 
| 49644 | 375 |   { assume H: "?rhs"
 | 
| 376 | then have ?lhs 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 | 377 | moreover | 
| 49644 | 378 |   { assume H: "?lhs"
 | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 379 | then obtain f where f:"\<forall>i. P i (f i)" unfolding choice_iff by metis | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 380 | let ?x = "(\<chi> i. (f i)) :: 'a ^ 'n" | 
| 49644 | 381 |     { fix i
 | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 382 | from f have "P i (f i)" by metis | 
| 37494 | 383 | then have "P i (?x $ i)" 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 | 384 | } | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 385 | hence "\<forall>i. P i (?x$i)" by metis | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 386 | hence ?rhs by metis } | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 387 | ultimately show ?thesis by metis | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 388 | qed | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 389 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 390 | lemma%unimportant rational_approximation: | 
| 67719 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 391 | assumes "e > 0" | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 392 | obtains r::real where "r \<in> \<rat>" "\<bar>r - x\<bar> < e" | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 393 | using Rats_dense_in_real [of "x - e/2" "x + e/2"] assms by auto | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 394 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 395 | lemma%important matrix_rational_approximation: | 
| 67719 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 396 | 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 | 397 | assumes "e > 0" | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 398 | obtains B where "\<And>i j. B$i$j \<in> \<rat>" "onorm(\<lambda>x. (A - B) *v x) < e" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 399 | proof%unimportant - | 
| 67719 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 400 |   have "\<forall>i j. \<exists>q \<in> \<rat>. \<bar>q - A $ i $ j\<bar> < e / (2 * CARD('m) * CARD('n))"
 | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 401 |     using assms by (force intro: rational_approximation [of "e / (2 * CARD('m) * CARD('n))"])
 | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 402 |   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))"
 | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 403 | by (auto simp: lambda_skolem Bex_def) | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 404 | show ?thesis | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 405 | proof | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 406 |     have "onorm (( *v) (A - B)) \<le> real CARD('m) * real CARD('n) *
 | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 407 |     (e / (2 * real CARD('m) * real CARD('n)))"
 | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 408 | apply (rule onorm_le_matrix_component) | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 409 | using Bclo by (simp add: abs_minus_commute less_imp_le) | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 410 | also have "\<dots> < e" | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 411 | using \<open>0 < e\<close> by (simp add: divide_simps) | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 412 | finally show "onorm (( *v) (A - B)) < e" . | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 413 | qed (use B in auto) | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 414 | qed | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 415 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 416 | lemma%unimportant 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 | 417 | 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 | 418 | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 419 | |
| 60420 | 420 | text \<open>The same result in terms of square matrices.\<close> | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 421 | |
| 68041 
d45b78cb86cf
more messy proofs redone, and new material
 paulson <lp15@cam.ac.uk> parents: 
68038diff
changeset | 422 | |
| 60420 | 423 | text \<open>Considering an n-element vector as an n-by-1 or 1-by-n matrix.\<close> | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 424 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 425 | definition%unimportant "rowvector v = (\<chi> i j. (v$j))" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 426 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 427 | definition%unimportant "columnvector v = (\<chi> i j. (v$i))" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 428 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 429 | lemma%unimportant transpose_columnvector: "transpose(columnvector v) = rowvector v" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 430 | by (simp add: transpose_def rowvector_def columnvector_def vec_eq_iff) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 431 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 432 | lemma%unimportant transpose_rowvector: "transpose(rowvector v) = columnvector v" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 433 | by (simp add: transpose_def columnvector_def rowvector_def vec_eq_iff) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 434 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 435 | lemma%unimportant dot_rowvector_columnvector: "columnvector (A *v v) = A ** columnvector v" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 436 | by (vector columnvector_def matrix_matrix_mult_def matrix_vector_mult_def) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 437 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 438 | lemma%unimportant dot_matrix_product: | 
| 49644 | 439 | "(x::real^'n) \<bullet> y = (((rowvector x ::real^'n^1) ** (columnvector y :: real^1^'n))$1)$1" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 440 | by (vector matrix_matrix_mult_def rowvector_def columnvector_def inner_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 | 441 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 442 | lemma%unimportant dot_matrix_vector_mul: | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 443 | fixes A B :: "real ^'n ^'n" and x y :: "real ^'n" | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 444 | shows "(A *v x) \<bullet> (B *v y) = | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 445 | (((rowvector x :: real^'n^1) ** ((transpose A ** B) ** (columnvector y :: real ^1^'n)))$1)$1" | 
| 49644 | 446 | unfolding dot_matrix_product transpose_columnvector[symmetric] | 
| 447 | dot_rowvector_columnvector matrix_transpose_mul matrix_mul_assoc .. | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 448 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 449 | lemma%unimportant 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 | 450 | 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 | 451 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 452 | lemma%unimportant 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 | 453 | 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 | 454 | 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 | 455 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 456 | lemma%unimportant 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 | 457 | 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 | 458 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 459 | lemma%unimportant 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 | 460 | 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 | 461 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 462 | lemma%unimportant continuous_on_vec_lambda[continuous_intros]: | 
| 63334 
bd37a72a940a
Multivariate_Analysis: add continuous_on_vec_lambda
 hoelzl parents: 
63332diff
changeset | 463 | "(\<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 | 464 | unfolding continuous_on_def by (auto intro: tendsto_vec_lambda) | 
| 
bd37a72a940a
Multivariate_Analysis: add continuous_on_vec_lambda
 hoelzl parents: 
63332diff
changeset | 465 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 466 | lemma%unimportant closed_positive_orthant: "closed {x::real^'n. \<forall>i. 0 \<le>x$i}"
 | 
| 63332 | 467 | by (simp add: Collect_all_eq closed_INT closed_Collect_le continuous_on_const continuous_on_id continuous_on_component) | 
| 44213 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 huffman parents: 
44211diff
changeset | 468 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 469 | lemma%unimportant bounded_component_cart: "bounded s \<Longrightarrow> bounded ((\<lambda>x. x $ i) ` s)" | 
| 49644 | 470 | unfolding bounded_def | 
| 471 | apply clarify | |
| 472 | apply (rule_tac x="x $ i" in exI) | |
| 473 | apply (rule_tac x="e" in exI) | |
| 474 | apply clarify | |
| 475 | apply (rule order_trans [OF dist_vec_nth_le], simp) | |
| 476 | done | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 477 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 478 | lemma%important 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 | 479 | fixes f :: "nat \<Rightarrow> 'a::heine_borel ^ 'n" | 
| 50998 | 480 | assumes f: "bounded (range f)" | 
| 66447 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 481 | 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 | 482 | (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) $ i) (l $ i) < e) sequentially)" | 
| 62127 | 483 | (is "?th d") | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 484 | proof%unimportant - | 
| 62127 | 485 | have "\<forall>d' \<subseteq> d. ?th d'" | 
| 486 | by (rule compact_lemma_general[where unproj=vec_lambda]) | |
| 487 | (auto intro!: f bounded_component_cart simp: vec_lambda_eta) | |
| 488 | 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 | 489 | qed | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 490 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 491 | 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 | 492 | proof | 
| 50998 | 493 | fix f :: "nat \<Rightarrow> 'a ^ 'b" | 
| 494 | assume f: "bounded (range f)" | |
| 66447 
a1f5c5c26fa6
Replaced subseq with strict_mono
 eberlm <eberlm@in.tum.de> parents: 
64267diff
changeset | 495 | then obtain l r where r: "strict_mono r" | 
| 49644 | 496 | and l: "\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>UNIV. dist (f (r n) $ i) (l $ i) < e) sequentially" | 
| 50998 | 497 | 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 | 498 | 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 | 499 |   { 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 | 500 | hence "0 < e / (real_of_nat (card ?d))" | 
| 49644 | 501 | 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 | 502 | 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 | 503 | by simp | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 504 | moreover | 
| 49644 | 505 |     { fix n
 | 
| 506 | 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 | 507 | have "dist (f (r n)) l \<le> (\<Sum>i\<in>?d. dist (f (r n) $ i) (l $ i))" | 
| 67155 | 508 | 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 | 509 | also have "\<dots> < (\<Sum>i\<in>?d. e / (real_of_nat (card ?d)))" | 
| 64267 | 510 | 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 | 511 | 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 | 512 | } | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 513 | ultimately have "eventually (\<lambda>n. dist (f (r n)) l < e) sequentially" | 
| 61810 | 514 | 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 | 515 | } | 
| 61973 | 516 | 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 | 517 | 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 | 518 | qed | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 519 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 520 | lemma%unimportant interval_cart: | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 521 | fixes a :: "real^'n" | 
| 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 522 |   shows "box a b = {x::real^'n. \<forall>i. a$i < x$i \<and> x$i < b$i}"
 | 
| 56188 | 523 |     and "cbox a b = {x::real^'n. \<forall>i. a$i \<le> x$i \<and> x$i \<le> b$i}"
 | 
| 524 | 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 | 525 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 526 | lemma%unimportant mem_box_cart: | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 527 | fixes a :: "real^'n" | 
| 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 528 | shows "x \<in> box a b \<longleftrightarrow> (\<forall>i. a$i < x$i \<and> x$i < b$i)" | 
| 56188 | 529 | and "x \<in> cbox a b \<longleftrightarrow> (\<forall>i. a$i \<le> x$i \<and> x$i \<le> b$i)" | 
| 49644 | 530 | 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 | 531 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 532 | lemma%unimportant interval_eq_empty_cart: | 
| 49644 | 533 | fixes a :: "real^'n" | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 534 |   shows "(box a b = {} \<longleftrightarrow> (\<exists>i. b$i \<le> a$i))" (is ?th1)
 | 
| 56188 | 535 |     and "(cbox a b = {} \<longleftrightarrow> (\<exists>i. b$i < a$i))" (is ?th2)
 | 
| 49644 | 536 | proof - | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 537 |   { 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 | 538 | 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 | 539 | hence "a$i < b$i" by auto | 
| 49644 | 540 | 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 | 541 | moreover | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 542 |   { 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 | 543 | 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 | 544 |     { fix i
 | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 545 | 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 | 546 | 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 | 547 | unfolding vector_smult_component and vector_add_component | 
| 49644 | 548 | by auto } | 
| 67673 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 549 |     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 | 550 | 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 | 551 | |
| 56188 | 552 |   { 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 | 553 | 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 | 554 | hence "a$i \<le> b$i" by auto | 
| 49644 | 555 | 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 | 556 | moreover | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 557 |   { 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 | 558 | 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 | 559 |     { fix i
 | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 560 | 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 | 561 | 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 | 562 | unfolding vector_smult_component and vector_add_component | 
| 49644 | 563 | by auto } | 
| 67673 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 564 |     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 | 565 | 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 | 566 | qed | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 567 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 568 | lemma%unimportant interval_ne_empty_cart: | 
| 49644 | 569 | fixes a :: "real^'n" | 
| 56188 | 570 |   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 | 571 |     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 | 572 | 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 | 573 | (* 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 | 574 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 575 | lemma%unimportant subset_interval_imp_cart: | 
| 49644 | 576 | fixes a :: "real^'n" | 
| 56188 | 577 | shows "(\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> cbox c d \<subseteq> cbox a b" | 
| 578 | and "(\<forall>i. a$i < c$i \<and> d$i < b$i) \<Longrightarrow> cbox c d \<subseteq> box a b" | |
| 579 | 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 | 580 | 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 | 581 | 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 | 582 | 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 | 583 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 584 | lemma%unimportant interval_sing: | 
| 49644 | 585 | fixes a :: "'a::linorder^'n" | 
| 586 |   shows "{a .. a} = {a} \<and> {a<..<a} = {}"
 | |
| 587 | apply (auto simp add: set_eq_iff less_vec_def less_eq_vec_def vec_eq_iff) | |
| 588 | done | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 589 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 590 | lemma%unimportant subset_interval_cart: | 
| 49644 | 591 | fixes a :: "real^'n" | 
| 56188 | 592 | 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) | 
| 593 | 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) | |
| 594 | 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 | 595 | 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 | 596 | 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 | 597 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 598 | lemma%unimportant disjoint_interval_cart: | 
| 49644 | 599 | fixes a::"real^'n" | 
| 56188 | 600 |   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)
 | 
| 601 |     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)
 | |
| 602 |     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 | 603 |     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 | 604 | 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 | 605 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 606 | lemma%unimportant Int_interval_cart: | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54489diff
changeset | 607 | fixes a :: "real^'n" | 
| 56188 | 608 |   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 | 609 | unfolding Int_interval | 
| 56188 | 610 | by (auto simp: mem_box less_eq_vec_def) | 
| 611 | (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 | 612 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 613 | lemma%unimportant closed_interval_left_cart: | 
| 49644 | 614 | 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 | 615 |   shows "closed {x::real^'n. \<forall>i. x$i \<le> b$i}"
 | 
| 63332 | 616 | by (simp add: Collect_all_eq closed_INT closed_Collect_le continuous_on_const continuous_on_id 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 | 617 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 618 | lemma%unimportant closed_interval_right_cart: | 
| 49644 | 619 | 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 | 620 |   shows "closed {x::real^'n. \<forall>i. a$i \<le> x$i}"
 | 
| 63332 | 621 | by (simp add: Collect_all_eq closed_INT closed_Collect_le continuous_on_const continuous_on_id 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 | 622 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 623 | lemma%unimportant is_interval_cart: | 
| 49644 | 624 | "is_interval (s::(real^'n) set) \<longleftrightarrow> | 
| 625 | (\<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 | 626 | 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 | 627 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 628 | lemma%unimportant closed_halfspace_component_le_cart: "closed {x::real^'n. x$i \<le> a}"
 | 
| 63332 | 629 | by (simp add: closed_Collect_le continuous_on_const continuous_on_id 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 | 630 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 631 | lemma%unimportant closed_halfspace_component_ge_cart: "closed {x::real^'n. x$i \<ge> a}"
 | 
| 63332 | 632 | by (simp add: closed_Collect_le continuous_on_const continuous_on_id 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 | 633 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 634 | lemma%unimportant open_halfspace_component_lt_cart: "open {x::real^'n. x$i < a}"
 | 
| 63332 | 635 | by (simp add: open_Collect_less continuous_on_const continuous_on_id continuous_on_component) | 
| 49644 | 636 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 637 | lemma%unimportant open_halfspace_component_gt_cart: "open {x::real^'n. x$i  > a}"
 | 
| 63332 | 638 | by (simp add: open_Collect_less continuous_on_const continuous_on_id 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 | 639 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 640 | lemma%unimportant Lim_component_le_cart: | 
| 49644 | 641 | fixes f :: "'a \<Rightarrow> real^'n" | 
| 61973 | 642 | 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 | 643 | 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 | 644 | 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 | 645 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 646 | lemma%unimportant Lim_component_ge_cart: | 
| 49644 | 647 | fixes f :: "'a \<Rightarrow> real^'n" | 
| 61973 | 648 | 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 | 649 | 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 | 650 | 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 | 651 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 652 | lemma%unimportant Lim_component_eq_cart: | 
| 49644 | 653 | fixes f :: "'a \<Rightarrow> real^'n" | 
| 61973 | 654 | assumes net: "(f \<longlongrightarrow> l) net" "~(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 | 655 | shows "l$i = b" | 
| 49644 | 656 | using ev[unfolded order_eq_iff eventually_conj_iff] and | 
| 657 | 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 | 658 | 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 | 659 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 660 | lemma%unimportant connected_ivt_component_cart: | 
| 49644 | 661 | fixes x :: "real^'n" | 
| 662 | 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 | 663 | 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 | 664 | 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 | 665 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 666 | lemma%unimportant 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 | 667 | 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 | 668 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 669 | lemma%important closed_substandard_cart: | 
| 44213 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 huffman parents: 
44211diff
changeset | 670 |   "closed {x::'a::real_normed_vector ^ 'n. \<forall>i. P i \<longrightarrow> x$i = 0}"
 | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 671 | proof%unimportant - | 
| 44213 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 huffman parents: 
44211diff
changeset | 672 |   { fix i::'n
 | 
| 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 huffman parents: 
44211diff
changeset | 673 |     have "closed {x::'a ^ 'n. P i \<longrightarrow> x$i = 0}"
 | 
| 63332 | 674 | by (cases "P i") (simp_all add: closed_Collect_eq continuous_on_const continuous_on_id continuous_on_component) } | 
| 44213 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 huffman parents: 
44211diff
changeset | 675 | thus ?thesis | 
| 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 huffman parents: 
44211diff
changeset | 676 | 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 | 677 | qed | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 678 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 679 | lemma%important dim_substandard_cart: "vec.dim {x::'a::field^'n. \<forall>i. i \<notin> d \<longrightarrow> x$i = 0} = card d"
 | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 680 | (is "vec.dim ?A = _") | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 681 | proof%unimportant (rule vec.dim_unique) | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 682 | let ?B = "((\<lambda>x. axis x 1) ` d)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 683 | have subset_basis: "?B \<subseteq> cart_basis" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 684 | by (auto simp: cart_basis_def) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 685 | show "?B \<subseteq> ?A" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 686 | by (auto simp: axis_def) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 687 | show "vec.independent ((\<lambda>x. axis x 1) ` d)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 688 | using subset_basis | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 689 | by (rule vec.independent_mono[OF vec.independent_Basis]) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 690 | have "x \<in> vec.span ?B" if "\<forall>i. i \<notin> d \<longrightarrow> x $ i = 0" for x::"'a^'n" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 691 | proof - | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 692 | have "finite ?B" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 693 | using subset_basis finite_cart_basis | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 694 | by (rule finite_subset) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 695 | have "x = (\<Sum>i\<in>UNIV. x $ i *s axis i 1)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 696 | by (rule basis_expansion[symmetric]) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 697 | also have "\<dots> = (\<Sum>i\<in>d. (x $ i) *s axis i 1)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 698 | by (rule sum.mono_neutral_cong_right) (auto simp: that) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 699 | also have "\<dots> \<in> vec.span ?B" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 700 | by (simp add: vec.span_sum vec.span_clauses) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 701 | finally show "x \<in> vec.span ?B" . | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 702 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 703 | then show "?A \<subseteq> vec.span ?B" by auto | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 704 | qed (simp add: card_image inj_on_def axis_eq_axis) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 705 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 706 | lemma%unimportant dim_subset_UNIV_cart_gen: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 707 |   fixes S :: "('a::field^'n) set"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 708 |   shows "vec.dim S \<le> CARD('n)"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 709 | by (metis vec.dim_eq_full vec.dim_subset_UNIV vec.span_UNIV vec_dim_card) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 710 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 711 | lemma%unimportant dim_subset_UNIV_cart: | 
| 67719 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 712 | fixes S :: "(real^'n) set" | 
| 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 713 |   shows "dim S \<le> CARD('n)"
 | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 714 | using dim_subset_UNIV_cart_gen[of S] by (simp add: dim_vec_eq) | 
| 67719 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 715 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 716 | lemma%unimportant affinity_inverses: | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 717 | assumes m0: "m \<noteq> (0::'a::field)" | 
| 61736 | 718 | shows "(\<lambda>x. m *s x + c) \<circ> (\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) = id" | 
| 719 | "(\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) \<circ> (\<lambda>x. m *s x + c) = id" | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 720 | using m0 | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 721 | by (auto simp add: fun_eq_iff vector_add_ldistrib diff_conv_add_uminus simp del: add_uminus_conv_diff) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 722 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 723 | lemma%important vector_affinity_eq: | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 724 | assumes m0: "(m::'a::field) \<noteq> 0" | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 725 | shows "m *s x + c = y \<longleftrightarrow> x = inverse m *s y + -(inverse m *s c)" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 726 | proof%unimportant | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 727 | assume h: "m *s x + c = y" | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 728 | hence "m *s x = y - c" by (simp add: field_simps) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 729 | hence "inverse m *s (m *s x) = inverse m *s (y - c)" by simp | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 730 | then show "x = inverse m *s y + - (inverse m *s c)" | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 731 | using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 732 | next | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 733 | assume h: "x = inverse m *s y + - (inverse m *s c)" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53600diff
changeset | 734 | show "m *s x + c = y" unfolding h | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 735 | using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 736 | qed | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 737 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 738 | lemma%unimportant vector_eq_affinity: | 
| 49644 | 739 | "(m::'a::field) \<noteq> 0 ==> (y = m *s x + c \<longleftrightarrow> inverse(m) *s y + -(inverse(m) *s c) = x)" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 740 | using vector_affinity_eq[where m=m and x=x and y=y and c=c] | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 741 | by metis | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 742 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 743 | lemma%unimportant vector_cart: | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49962diff
changeset | 744 | fixes f :: "real^'n \<Rightarrow> real" | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49962diff
changeset | 745 | shows "(\<chi> i. f (axis i 1)) = (\<Sum>i\<in>Basis. f i *\<^sub>R i)" | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49962diff
changeset | 746 | unfolding euclidean_eq_iff[where 'a="real^'n"] | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
49962diff
changeset | 747 | by simp (simp add: Basis_vec_def inner_axis) | 
| 63332 | 748 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 749 | lemma%unimportant const_vector_cart:"((\<chi> i. d)::real^'n) = (\<Sum>i\<in>Basis. d *\<^sub>R 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 | 750 | by (rule vector_cart) | 
| 49644 | 751 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 752 | subsection%important "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 | 753 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 754 | lemma%unimportant 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 | 755 | 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 | 756 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 757 | 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 | 758 | 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 | 759 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 760 | lemmas%unimportant 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 | 761 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 762 | lemma%unimportant 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 | 763 |   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 | 764 |   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 | 765 | 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 | 766 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 767 | lemma%unimportant 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 | 768 | 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 | 769 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 770 | lemma%unimportant unit_interval_convex_hull_cart: | 
| 56188 | 771 |   "cbox (0::real^'n) 1 = convex hull {x. \<forall>i. (x$i = 0) \<or> (x$i = 1)}"
 | 
| 772 | 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 | 773 | 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 | 774 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 775 | lemma%important cube_convex_hull_cart: | 
| 49644 | 776 | assumes "0 < d" | 
| 777 | obtains s::"(real^'n) set" | |
| 56188 | 778 | where "finite s" "cbox (x - (\<chi> i. d)) (x + (\<chi> i. d)) = convex hull s" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 779 | proof%unimportant - | 
| 55522 | 780 | from assms obtain s where "finite s" | 
| 67399 | 781 | and "cbox (x - sum (( *\<^sub>R) d) Basis) (x + sum (( *\<^sub>R) d) Basis) = convex hull s" | 
| 55522 | 782 | by (rule cube_convex_hull) | 
| 783 | with that[of s] show thesis | |
| 784 | 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 | 785 | qed | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 786 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 787 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 788 | subsection%important "Derivative" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 789 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 790 | definition%important "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 | 791 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 792 | lemma%important jacobian_works: | 
| 49644 | 793 | "(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 | 794 | (f has_derivative (\<lambda>h. (jacobian f net) *v h)) net" (is "?lhs = ?rhs") | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 795 | proof%unimportant | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 796 | 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 | 797 | 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 | 798 | next | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 799 | 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 | 800 | 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 | 801 | qed | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 802 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 803 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 804 | subsection%important \<open>Component of the differential must be zero if it exists at a local | 
| 67968 | 805 | 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 | 806 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 807 | lemma%important differential_zero_maxmin_cart: | 
| 49644 | 808 | fixes f::"real^'a \<Rightarrow> real^'b" | 
| 809 | 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 | 810 | "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 | 811 | 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 | 812 | 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 | 813 | 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 | 814 | by (simp add: Basis_vec_def axis_eq_axis inner_axis jacobian_def matrix_def) | 
| 49644 | 815 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 816 | subsection%unimportant \<open>Lemmas for working on @{typ "real^1"}\<close>
 | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 817 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 818 | lemma forall_1[simp]: "(\<forall>i::1. P i) \<longleftrightarrow> P 1" | 
| 49644 | 819 | by (metis (full_types) num1_eq_iff) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 820 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 821 | lemma ex_1[simp]: "(\<exists>x::1. P x) \<longleftrightarrow> P 1" | 
| 49644 | 822 | by auto (metis (full_types) num1_eq_iff) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 823 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 824 | lemma exhaust_2: | 
| 49644 | 825 | fixes x :: 2 | 
| 826 | shows "x = 1 \<or> x = 2" | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 827 | proof (induct x) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 828 | case (of_int z) | 
| 67979 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 829 | then have "0 \<le> z" and "z < 2" by simp_all | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 830 | then have "z = 0 | z = 1" by arith | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 831 | then show ?case by auto | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 832 | qed | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 833 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 834 | lemma forall_2: "(\<forall>i::2. P i) \<longleftrightarrow> P 1 \<and> P 2" | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 835 | by (metis exhaust_2) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 836 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 837 | lemma exhaust_3: | 
| 49644 | 838 | fixes x :: 3 | 
| 839 | shows "x = 1 \<or> x = 2 \<or> x = 3" | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 840 | proof (induct x) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 841 | case (of_int z) | 
| 67979 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 842 | then have "0 \<le> z" and "z < 3" by simp_all | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 843 | then have "z = 0 \<or> z = 1 \<or> z = 2" by arith | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 844 | then show ?case by auto | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 845 | qed | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 846 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 847 | lemma forall_3: "(\<forall>i::3. P i) \<longleftrightarrow> P 1 \<and> P 2 \<and> P 3" | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 848 | by (metis exhaust_3) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 849 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 850 | lemma UNIV_1 [simp]: "UNIV = {1::1}"
 | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 851 | by (auto simp add: num1_eq_iff) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 852 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 853 | lemma UNIV_2: "UNIV = {1::2, 2::2}"
 | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 854 | using exhaust_2 by auto | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 855 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 856 | lemma UNIV_3: "UNIV = {1::3, 2::3, 3::3}"
 | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 857 | using exhaust_3 by auto | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 858 | |
| 64267 | 859 | lemma sum_1: "sum f (UNIV::1 set) = f 1" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 860 | unfolding UNIV_1 by simp | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 861 | |
| 64267 | 862 | lemma sum_2: "sum f (UNIV::2 set) = f 1 + f 2" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 863 | unfolding UNIV_2 by simp | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 864 | |
| 64267 | 865 | lemma sum_3: "sum f (UNIV::3 set) = f 1 + f 2 + f 3" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 866 | unfolding UNIV_3 by (simp add: ac_simps) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 867 | |
| 67979 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 868 | lemma num1_eqI: | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 869 | fixes a::num1 shows "a = b" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 870 | by (metis (full_types) UNIV_1 UNIV_I empty_iff insert_iff) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 871 | |
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 872 | lemma num1_eq1 [simp]: | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 873 | fixes a::num1 shows "a = 1" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 874 | by (rule num1_eqI) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 875 | |
| 49644 | 876 | instantiation num1 :: cart_one | 
| 877 | begin | |
| 878 | ||
| 879 | instance | |
| 880 | proof | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 881 | show "CARD(1) = Suc 0" by auto | 
| 49644 | 882 | qed | 
| 883 | ||
| 884 | end | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 885 | |
| 67979 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 886 | instantiation num1 :: linorder begin | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 887 | definition "a < b \<longleftrightarrow> Rep_num1 a < Rep_num1 b" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 888 | definition "a \<le> b \<longleftrightarrow> Rep_num1 a \<le> Rep_num1 b" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 889 | instance | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 890 | by intro_classes (auto simp: less_eq_num1_def less_num1_def intro: num1_eqI) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 891 | end | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 892 | |
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 893 | instance num1 :: wellorder | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 894 | by intro_classes (auto simp: less_eq_num1_def less_num1_def) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 895 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 896 | subsection%unimportant\<open>The collapse of the general concepts to dimension one\<close> | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 897 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 898 | lemma vector_one: "(x::'a ^1) = (\<chi> i. (x$1))" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 899 | by (simp add: vec_eq_iff) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 900 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 901 | lemma forall_one: "(\<forall>(x::'a ^1). P x) \<longleftrightarrow> (\<forall>x. P(\<chi> i. x))" | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 902 | apply auto | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 903 | apply (erule_tac x= "x$1" in allE) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 904 | apply (simp only: vector_one[symmetric]) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 905 | done | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 906 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 907 | lemma norm_vector_1: "norm (x :: _^1) = norm (x$1)" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 908 | by (simp add: norm_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 | 909 | |
| 67979 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 910 | lemma dist_vector_1: | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 911 | fixes x :: "'a::real_normed_vector^1" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 912 | shows "dist x y = dist (x$1) (y$1)" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 913 | by (simp add: dist_norm norm_vector_1) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 914 | |
| 61945 | 915 | lemma norm_real: "norm(x::real ^ 1) = \<bar>x$1\<bar>" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 916 | by (simp add: norm_vector_1) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 917 | |
| 61945 | 918 | lemma dist_real: "dist(x::real ^ 1) y = \<bar>(x$1) - (y$1)\<bar>" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 919 | by (auto simp add: norm_real dist_norm) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 920 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 921 | subsection%important\<open> Rank of a matrix\<close> | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 922 | |
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 923 | text\<open>Equivalence of row and column rank is taken from George Mackiw's paper, Mathematics Magazine 1995, p. 285.\<close> | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 924 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 925 | lemma%unimportant matrix_vector_mult_in_columnspace_gen: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 926 | fixes A :: "'a::field^'n^'m" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 927 | shows "(A *v x) \<in> vec.span(columns A)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 928 | apply (simp add: matrix_vector_column columns_def transpose_def column_def) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 929 | apply (intro vec.span_sum vec.span_scale) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 930 | apply (force intro: vec.span_base) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 931 | done | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 932 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 933 | lemma%unimportant matrix_vector_mult_in_columnspace: | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 934 | fixes A :: "real^'n^'m" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 935 | shows "(A *v x) \<in> span(columns A)" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 936 | using matrix_vector_mult_in_columnspace_gen[of A x] by (simp add: span_vec_eq) | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 937 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 938 | lemma%important orthogonal_nullspace_rowspace: | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 939 | fixes A :: "real^'n^'m" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 940 | assumes 0: "A *v x = 0" and y: "y \<in> span(rows A)" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 941 | shows "orthogonal x y" | 
| 68077 
ee8c13ae81e9
Some tidying up (mostly regarding summations from 0)
 paulson <lp15@cam.ac.uk> parents: 
68074diff
changeset | 942 | using y | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 943 | proof%unimportant (induction rule: span_induct) | 
| 68077 
ee8c13ae81e9
Some tidying up (mostly regarding summations from 0)
 paulson <lp15@cam.ac.uk> parents: 
68074diff
changeset | 944 | case base | 
| 
ee8c13ae81e9
Some tidying up (mostly regarding summations from 0)
 paulson <lp15@cam.ac.uk> parents: 
68074diff
changeset | 945 | then show ?case | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 946 | by (simp add: subspace_orthogonal_to_vector) | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 947 | next | 
| 68077 
ee8c13ae81e9
Some tidying up (mostly regarding summations from 0)
 paulson <lp15@cam.ac.uk> parents: 
68074diff
changeset | 948 | case (step v) | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 949 | then obtain i where "v = row i A" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 950 | by (auto simp: rows_def) | 
| 68077 
ee8c13ae81e9
Some tidying up (mostly regarding summations from 0)
 paulson <lp15@cam.ac.uk> parents: 
68074diff
changeset | 951 | with 0 show ?case | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 952 | unfolding orthogonal_def inner_vec_def matrix_vector_mult_def row_def | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 953 | by (simp add: mult.commute) (metis (no_types) vec_lambda_beta zero_index) | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 954 | qed | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 955 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 956 | lemma%unimportant nullspace_inter_rowspace: | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 957 | fixes A :: "real^'n^'m" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 958 | shows "A *v x = 0 \<and> x \<in> span(rows A) \<longleftrightarrow> x = 0" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 959 | using orthogonal_nullspace_rowspace orthogonal_self span_zero matrix_vector_mult_0_right | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 960 | by blast | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 961 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 962 | lemma%unimportant matrix_vector_mul_injective_on_rowspace: | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 963 | fixes A :: "real^'n^'m" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 964 | shows "\<lbrakk>A *v x = A *v y; x \<in> span(rows A); y \<in> span(rows A)\<rbrakk> \<Longrightarrow> x = y" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 965 | using nullspace_inter_rowspace [of A "x-y"] | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 966 | by (metis diff_eq_diff_eq diff_self matrix_vector_mult_diff_distrib span_diff) | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 967 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 968 | definition%important rank :: "'a::field^'n^'m=>nat" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 969 | where row_rank_def_gen: "rank A \<equiv> vec.dim(rows A)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 970 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 971 | lemma%important row_rank_def: "rank A = dim (rows A)" for A::"real^'n^'m" | 
| 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 972 | by%unimportant (auto simp: row_rank_def_gen dim_vec_eq) | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 973 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 974 | lemma%important dim_rows_le_dim_columns: | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 975 | fixes A :: "real^'n^'m" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 976 | shows "dim(rows A) \<le> dim(columns A)" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 977 | proof%unimportant - | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 978 | have "dim (span (rows A)) \<le> dim (span (columns A))" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 979 | proof - | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 980 | obtain B where "independent B" "span(rows A) \<subseteq> span B" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 981 | and B: "B \<subseteq> span(rows A)""card B = dim (span(rows A))" | 
| 68074 | 982 | using basis_exists [of "span(rows A)"] by metis | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 983 | with span_subspace have eq: "span B = span(rows A)" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 984 | by auto | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 985 | then have inj: "inj_on (( *v) A) (span B)" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 986 | by (simp add: inj_on_def matrix_vector_mul_injective_on_rowspace) | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 987 | then have ind: "independent (( *v) A ` B)" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 988 | by (rule linear_independent_injective_image [OF Finite_Cartesian_Product.matrix_vector_mul_linear \<open>independent B\<close>]) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 989 | have "dim (span (rows A)) \<le> card (( *v) A ` B)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 990 | unfolding B(2)[symmetric] | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 991 | using inj | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 992 | by (auto simp: card_image inj_on_subset span_superset) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 993 | also have "\<dots> \<le> dim (span (columns A))" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 994 | using _ ind | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 995 | by (rule independent_card_le_dim) (auto intro!: matrix_vector_mult_in_columnspace) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 996 | finally show ?thesis . | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 997 | qed | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 998 | then show ?thesis | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 999 | by (simp add: dim_span) | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1000 | qed | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1001 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 1002 | lemma%unimportant column_rank_def: | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1003 | fixes A :: "real^'n^'m" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 1004 | shows "rank A = dim(columns A)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 1005 | unfolding row_rank_def | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 1006 | by (metis columns_transpose dim_rows_le_dim_columns le_antisym rows_transpose) | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1007 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 1008 | lemma%unimportant rank_transpose: | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1009 | fixes A :: "real^'n^'m" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1010 | shows "rank(transpose A) = rank A" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 1011 | by (metis column_rank_def row_rank_def rows_transpose) | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1012 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 1013 | lemma%unimportant matrix_vector_mult_basis: | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1014 | fixes A :: "real^'n^'m" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1015 | shows "A *v (axis k 1) = column k A" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1016 | by (simp add: cart_eq_inner_axis column_def matrix_mult_dot) | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1017 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 1018 | lemma%unimportant columns_image_basis: | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1019 | fixes A :: "real^'n^'m" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1020 | shows "columns A = ( *v) A ` (range (\<lambda>i. axis i 1))" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1021 | by (force simp: columns_def matrix_vector_mult_basis [symmetric]) | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1022 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 1023 | lemma%important rank_dim_range: | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1024 | fixes A :: "real^'n^'m" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1025 | shows "rank A = dim(range (\<lambda>x. A *v x))" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 1026 | unfolding column_rank_def | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 1027 | proof%unimportant (rule span_eq_dim) | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 1028 | have "span (columns A) \<subseteq> span (range (( *v) A))" (is "?l \<subseteq> ?r") | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 1029 | by (simp add: columns_image_basis image_subsetI span_mono) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 1030 | then show "?l = ?r" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 1031 | by (metis (no_types, lifting) image_subset_iff matrix_vector_mult_in_columnspace | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 1032 | span_eq span_span) | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1033 | qed | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1034 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 1035 | lemma%unimportant rank_bound: | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1036 | fixes A :: "real^'n^'m" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1037 |   shows "rank A \<le> min CARD('m) (CARD('n))"
 | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 1038 | by (metis (mono_tags, lifting) dim_subset_UNIV_cart min.bounded_iff | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 1039 | column_rank_def row_rank_def) | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1040 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 1041 | lemma%unimportant full_rank_injective: | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1042 | fixes A :: "real^'n^'m" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1043 |   shows "rank A = CARD('n) \<longleftrightarrow> inj (( *v) A)"
 | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 1044 | by (simp add: matrix_left_invertible_injective [symmetric] matrix_left_invertible_span_rows row_rank_def | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 1045 | dim_eq_full [symmetric] card_cart_basis vec.dimension_def) | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1046 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 1047 | lemma%unimportant full_rank_surjective: | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1048 | fixes A :: "real^'n^'m" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1049 |   shows "rank A = CARD('m) \<longleftrightarrow> surj (( *v) A)"
 | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1050 | by (simp add: matrix_right_invertible_surjective [symmetric] left_invertible_transpose [symmetric] | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1051 | matrix_left_invertible_injective full_rank_injective [symmetric] rank_transpose) | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1052 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 1053 | lemma%unimportant rank_I: "rank(mat 1::real^'n^'n) = CARD('n)"
 | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1054 | by (simp add: full_rank_injective inj_on_def) | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1055 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 1056 | lemma%unimportant less_rank_noninjective: | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1057 | fixes A :: "real^'n^'m" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1058 |   shows "rank A < CARD('n) \<longleftrightarrow> \<not> inj (( *v) A)"
 | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1059 | using less_le rank_bound by (auto simp: full_rank_injective [symmetric]) | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1060 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 1061 | lemma%unimportant matrix_nonfull_linear_equations_eq: | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1062 | fixes A :: "real^'n^'m" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1063 |   shows "(\<exists>x. (x \<noteq> 0) \<and> A *v x = 0) \<longleftrightarrow> ~(rank A = CARD('n))"
 | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1064 | by (meson matrix_left_invertible_injective full_rank_injective matrix_left_invertible_ker) | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1065 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 1066 | lemma%unimportant rank_eq_0: "rank A = 0 \<longleftrightarrow> A = 0" and rank_0 [simp]: "rank (0::real^'n^'m) = 0" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 1067 | for A :: "real^'n^'m" | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1068 | by (auto simp: rank_dim_range matrix_eq) | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1069 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 1070 | lemma%important rank_mul_le_right: | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1071 | fixes A :: "real^'n^'m" and B :: "real^'p^'n" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1072 | shows "rank(A ** B) \<le> rank B" | 
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 1073 | proof%unimportant - | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1074 | have "rank(A ** B) \<le> dim (( *v) A ` range (( *v) B))" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1075 | by (auto simp: rank_dim_range image_comp o_def matrix_vector_mul_assoc) | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1076 | also have "\<dots> \<le> rank B" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67986diff
changeset | 1077 | by (simp add: rank_dim_range dim_image_le) | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1078 | finally show ?thesis . | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1079 | qed | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1080 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 1081 | lemma%unimportant rank_mul_le_left: | 
| 67986 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1082 | fixes A :: "real^'n^'m" and B :: "real^'p^'n" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1083 | shows "rank(A ** B) \<le> rank A" | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1084 | by (metis matrix_transpose_mul rank_mul_le_right rank_transpose) | 
| 
b65c4a6a015e
quite a few more results about negligibility, etc., and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
67982diff
changeset | 1085 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 1086 | subsection%unimportant\<open>Routine results connecting the types @{typ "real^1"} and @{typ real}\<close>
 | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1087 | |
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1088 | lemma vector_one_nth [simp]: | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1089 | fixes x :: "'a^1" shows "vec (x $ 1) = x" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1090 | by (metis vec_def vector_one) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1091 | |
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1092 | lemma vec_cbox_1_eq [simp]: | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1093 | 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 | 1094 | 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 | 1095 | |
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1096 | lemma vec_nth_cbox_1_eq [simp]: | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1097 | fixes u v :: "'a::euclidean_space^1" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1098 | 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 | 1099 | 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 | 1100 | |
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1101 | lemma vec_nth_1_iff_cbox [simp]: | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1102 | fixes a b :: "'a::euclidean_space" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1103 | 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 | 1104 | (is "?lhs = ?rhs") | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1105 | proof | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1106 | assume L: ?lhs show ?rhs | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1107 | proof (intro equalityI subsetI) | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1108 | fix x | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1109 | assume "x \<in> S" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1110 | 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 | 1111 | using L by auto | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1112 | 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 | 1113 | 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 | 1114 | next | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1115 | fix x :: "'a^1" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1116 | 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 | 1117 | then show "x \<in> S" | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1118 | 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 | 1119 | qed | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67979diff
changeset | 1120 | qed simp | 
| 49644 | 1121 | |
| 67979 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1122 | lemma tendsto_at_within_vector_1: | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1123 | fixes S :: "'a :: metric_space set" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1124 | assumes "(f \<longlongrightarrow> fx) (at x within S)" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1125 | shows "((\<lambda>y::'a^1. \<chi> i. f (y $ 1)) \<longlongrightarrow> (vec fx::'a^1)) (at (vec x) within vec ` S)" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1126 | proof (rule topological_tendstoI) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1127 |   fix T :: "('a^1) set"
 | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1128 | assume "open T" "vec fx \<in> T" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1129 | have "\<forall>\<^sub>F x in at x within S. f x \<in> (\<lambda>x. x $ 1) ` T" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1130 | using \<open>open T\<close> \<open>vec fx \<in> T\<close> assms open_image_vec_nth tendsto_def by fastforce | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1131 | then show "\<forall>\<^sub>F x::'a^1 in at (vec x) within vec ` S. (\<chi> i. f (x $ 1)) \<in> T" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1132 | unfolding eventually_at dist_norm [symmetric] | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1133 | by (rule ex_forward) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1134 | (use \<open>open T\<close> in | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1135 | \<open>fastforce simp: dist_norm dist_vec_def L2_set_def image_iff vector_one open_vec_def\<close>) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1136 | qed | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1137 | |
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1138 | lemma has_derivative_vector_1: | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1139 | assumes der_g: "(g has_derivative (\<lambda>x. x * g' a)) (at a within S)" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1140 | shows "((\<lambda>x. vec (g (x $ 1))) has_derivative ( *\<^sub>R) (g' a)) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1141 | (at ((vec a)::real^1) within vec ` S)" | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1142 | using der_g | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1143 | apply (auto simp: Deriv.has_derivative_within bounded_linear_scaleR_right norm_vector_1) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1144 | apply (drule tendsto_at_within_vector_1, vector) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1145 | apply (auto simp: algebra_simps eventually_at tendsto_def) | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1146 | done | 
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1147 | |
| 
53323937ee25
new material about vec, real^1, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 1148 | |
| 68833 
fde093888c16
tagged 21 theories in the Analysis library for the manual
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
68077diff
changeset | 1149 | subsection%unimportant\<open>Explicit vector construction from lists\<close> | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1150 | |
| 43995 
c479836d9048
simplified definition of vector (also removed Cartesian_Euclidean_Space.from_nat which collides with Countable.from_nat)
 hoelzl parents: 
42814diff
changeset | 1151 | definition "vector l = (\<chi> i. foldr (\<lambda>x f n. fun_upd (f (n+1)) n x) l (\<lambda>n x. 0) 1 i)" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1152 | |
| 68054 
ebd179b82e20
getting rid of more "defer", etc.
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 1153 | lemma vector_1 [simp]: "(vector[x]) $1 = x" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1154 | unfolding vector_def by simp | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1155 | |
| 68054 
ebd179b82e20
getting rid of more "defer", etc.
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 1156 | lemma vector_2 [simp]: "(vector[x,y]) $1 = x" "(vector[x,y] :: 'a^2)$2 = (y::'a::zero)" | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1157 | unfolding vector_def by simp_all | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1158 | |
| 68054 
ebd179b82e20
getting rid of more "defer", etc.
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 1159 | lemma vector_3 [simp]: | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1160 |  "(vector [x,y,z] ::('a::zero)^3)$1 = x"
 | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1161 |  "(vector [x,y,z] ::('a::zero)^3)$2 = y"
 | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1162 |  "(vector [x,y,z] ::('a::zero)^3)$3 = z"
 | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1163 | unfolding vector_def by simp_all | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1164 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1165 | lemma forall_vector_1: "(\<forall>v::'a::zero^1. P v) \<longleftrightarrow> (\<forall>x. P(vector[x]))" | 
| 67719 
bffb7482faaa
new material on matrices, etc., and consolidating duplicate results about of_nat
 paulson <lp15@cam.ac.uk> parents: 
67686diff
changeset | 1166 | by (metis vector_1 vector_one) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1167 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1168 | lemma forall_vector_2: "(\<forall>v::'a::zero^2. P v) \<longleftrightarrow> (\<forall>x y. P(vector[x, y]))" | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1169 | apply auto | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1170 | apply (erule_tac x="v$1" in allE) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1171 | apply (erule_tac x="v$2" in allE) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1172 | apply (subgoal_tac "vector [v$1, v$2] = v") | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1173 | apply simp | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1174 | apply (vector vector_def) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1175 | apply (simp add: forall_2) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1176 | done | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1177 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1178 | lemma forall_vector_3: "(\<forall>v::'a::zero^3. P v) \<longleftrightarrow> (\<forall>x y z. P(vector[x, y, z]))" | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1179 | apply auto | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1180 | apply (erule_tac x="v$1" in allE) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1181 | apply (erule_tac x="v$2" in allE) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1182 | apply (erule_tac x="v$3" in allE) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1183 | apply (subgoal_tac "vector [v$1, v$2, v$3] = v") | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1184 | apply simp | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1185 | apply (vector vector_def) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1186 | apply (simp add: forall_3) | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1187 | done | 
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1188 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1189 | lemma bounded_linear_component_cart[intro]: "bounded_linear (\<lambda>x::real^'n. x $ k)" | 
| 68062 | 1190 | apply (rule bounded_linear_intro[where K=1]) | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1191 | using component_le_norm_cart[of _ k] unfolding real_norm_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 | 1192 | |
| 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1193 | 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 | 1194 |   "{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 | 1195 |   "cbox a b \<inter> {x. x$k \<ge> c} = {(\<chi> i. if i = k then max (a$k) c else a$i) .. b}"
 | 
| 49644 | 1196 | apply (rule_tac[!] set_eqI) | 
| 67673 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 1197 | unfolding Int_iff mem_box_cart mem_Collect_eq interval_cbox_cart | 
| 49644 | 1198 | unfolding vec_lambda_beta | 
| 1199 | 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 | 1200 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1201 | 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 | 1202 | 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 | 1203 | 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 | 1204 | bounded_linear.uniform_limit[OF bounded_linear_component_cart] | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67673diff
changeset | 1205 | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: diff
changeset | 1206 | end |