| author | haftmann | 
| Wed, 28 Apr 2010 16:56:18 +0200 | |
| changeset 36513 | 70096cbdd4e0 | 
| parent 36435 | bbe2730e6db6 | 
| child 36587 | 534418d8d494 | 
| permissions | -rw-r--r-- | 
| 36431 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 1 | (* Title: Multivariate_Analysis/Vec1.thy | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 2 | Author: Amine Chaieb, University of Cambridge | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 3 | Author: Robert Himmelmann, TU Muenchen | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 4 | *) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 5 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 6 | header {* Vectors of size 1, 2, or 3 *}
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 7 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 8 | theory Vec1 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 9 | imports Topology_Euclidean_Space | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 10 | begin | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 11 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 12 | text{* Some common special cases.*}
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 13 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 14 | lemma forall_1[simp]: "(\<forall>i::1. P i) \<longleftrightarrow> P 1" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 15 | by (metis num1_eq_iff) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 16 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 17 | lemma ex_1[simp]: "(\<exists>x::1. P x) \<longleftrightarrow> P 1" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 18 | by auto (metis num1_eq_iff) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 19 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 20 | lemma exhaust_2: | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 21 | fixes x :: 2 shows "x = 1 \<or> x = 2" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 22 | proof (induct x) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 23 | case (of_int z) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 24 | then have "0 <= z" and "z < 2" by simp_all | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 25 | then have "z = 0 | z = 1" by arith | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 26 | then show ?case by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 27 | qed | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 28 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 29 | lemma forall_2: "(\<forall>i::2. P i) \<longleftrightarrow> P 1 \<and> P 2" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 30 | by (metis exhaust_2) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 31 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 32 | lemma exhaust_3: | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 33 | fixes x :: 3 shows "x = 1 \<or> x = 2 \<or> x = 3" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 34 | proof (induct x) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 35 | case (of_int z) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 36 | then have "0 <= z" and "z < 3" by simp_all | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 37 | then have "z = 0 \<or> z = 1 \<or> z = 2" by arith | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 38 | then show ?case by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 39 | qed | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 40 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 41 | lemma forall_3: "(\<forall>i::3. P i) \<longleftrightarrow> P 1 \<and> P 2 \<and> P 3" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 42 | by (metis exhaust_3) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 43 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 44 | lemma UNIV_1 [simp]: "UNIV = {1::1}"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 45 | by (auto simp add: num1_eq_iff) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 46 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 47 | lemma UNIV_2: "UNIV = {1::2, 2::2}"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 48 | using exhaust_2 by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 49 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 50 | lemma UNIV_3: "UNIV = {1::3, 2::3, 3::3}"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 51 | using exhaust_3 by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 52 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 53 | lemma setsum_1: "setsum f (UNIV::1 set) = f 1" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 54 | unfolding UNIV_1 by simp | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 55 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 56 | lemma setsum_2: "setsum f (UNIV::2 set) = f 1 + f 2" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 57 | unfolding UNIV_2 by simp | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 58 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 59 | lemma setsum_3: "setsum f (UNIV::3 set) = f 1 + f 2 + f 3" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 60 | unfolding UNIV_3 by (simp add: add_ac) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 61 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 62 | instantiation num1 :: cart_one begin | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 63 | instance proof | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 64 | show "CARD(1) = Suc 0" by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 65 | qed end | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 66 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 67 | (* "lift" from 'a to 'a^1 and "drop" from 'a^1 to 'a -- FIXME: potential use of transfer *) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 68 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 69 | abbreviation vec1:: "'a \<Rightarrow> 'a ^ 1" where "vec1 x \<equiv> vec x" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 70 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 71 | abbreviation dest_vec1:: "'a ^1 \<Rightarrow> 'a" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 72 | where "dest_vec1 x \<equiv> (x$1)" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 73 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 74 | lemma vec1_component[simp]: "(vec1 x)$1 = x" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 75 | by simp | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 76 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 77 | lemma vec1_dest_vec1: "vec1(dest_vec1 x) = x" "dest_vec1(vec1 y) = y" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 78 | by (simp_all add: Cart_eq) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 79 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 80 | declare vec1_dest_vec1(1) [simp] | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 81 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 82 | lemma forall_vec1: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P (vec1 x))" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 83 | by (metis vec1_dest_vec1(1)) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 84 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 85 | lemma exists_vec1: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>x. P(vec1 x))" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 86 | by (metis vec1_dest_vec1(1)) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 87 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 88 | lemma vec1_eq[simp]: "vec1 x = vec1 y \<longleftrightarrow> x = y" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 89 | by (metis vec1_dest_vec1(2)) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 90 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 91 | lemma dest_vec1_eq[simp]: "dest_vec1 x = dest_vec1 y \<longleftrightarrow> x = y" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 92 | by (metis vec1_dest_vec1(1)) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 93 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 94 | subsection{* The collapse of the general concepts to dimension one. *}
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 95 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 96 | lemma vector_one: "(x::'a ^1) = (\<chi> i. (x$1))" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 97 | by (simp add: Cart_eq) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 98 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 99 | lemma forall_one: "(\<forall>(x::'a ^1). P x) \<longleftrightarrow> (\<forall>x. P(\<chi> i. x))" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 100 | apply auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 101 | apply (erule_tac x= "x$1" in allE) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 102 | apply (simp only: vector_one[symmetric]) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 103 | done | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 104 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 105 | lemma norm_vector_1: "norm (x :: _^1) = norm (x$1)" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 106 | by (simp add: norm_vector_def) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 107 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 108 | lemma norm_real: "norm(x::real ^ 1) = abs(x$1)" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 109 | by (simp add: norm_vector_1) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 110 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 111 | lemma dist_real: "dist(x::real ^ 1) y = abs((x$1) - (y$1))" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 112 | by (auto simp add: norm_real dist_norm) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 113 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 114 | subsection{* Explicit vector construction from lists. *}
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 115 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 116 | primrec from_nat :: "nat \<Rightarrow> 'a::{monoid_add,one}"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 117 | where "from_nat 0 = 0" | "from_nat (Suc n) = 1 + from_nat n" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 118 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 119 | lemma from_nat [simp]: "from_nat = of_nat" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 120 | by (rule ext, induct_tac x, simp_all) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 121 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 122 | primrec | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 123 | list_fun :: "nat \<Rightarrow> _ list \<Rightarrow> _ \<Rightarrow> _" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 124 | where | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 125 | "list_fun n [] = (\<lambda>x. 0)" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 126 | | "list_fun n (x # xs) = fun_upd (list_fun (Suc n) xs) (from_nat n) x" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 127 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 128 | definition "vector l = (\<chi> i. list_fun 1 l i)" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 129 | (*definition "vector l = (\<chi> i. if i <= length l then l ! (i - 1) else 0)"*) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 130 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 131 | lemma vector_1: "(vector[x]) $1 = x" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 132 | unfolding vector_def by simp | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 133 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 134 | lemma vector_2: | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 135 | "(vector[x,y]) $1 = x" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 136 | "(vector[x,y] :: 'a^2)$2 = (y::'a::zero)" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 137 | unfolding vector_def by simp_all | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 138 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 139 | lemma vector_3: | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 140 |  "(vector [x,y,z] ::('a::zero)^3)$1 = x"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 141 |  "(vector [x,y,z] ::('a::zero)^3)$2 = y"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 142 |  "(vector [x,y,z] ::('a::zero)^3)$3 = z"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 143 | unfolding vector_def by simp_all | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 144 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 145 | lemma forall_vector_1: "(\<forall>v::'a::zero^1. P v) \<longleftrightarrow> (\<forall>x. P(vector[x]))" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 146 | apply auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 147 | apply (erule_tac x="v$1" in allE) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 148 | apply (subgoal_tac "vector [v$1] = v") | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 149 | apply simp | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 150 | apply (vector vector_def) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 151 | apply simp | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 152 | done | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 153 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 154 | lemma forall_vector_2: "(\<forall>v::'a::zero^2. P v) \<longleftrightarrow> (\<forall>x y. P(vector[x, y]))" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 155 | apply auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 156 | apply (erule_tac x="v$1" in allE) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 157 | apply (erule_tac x="v$2" in allE) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 158 | apply (subgoal_tac "vector [v$1, v$2] = v") | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 159 | apply simp | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 160 | apply (vector vector_def) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 161 | apply (simp add: forall_2) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 162 | done | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 163 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 164 | lemma forall_vector_3: "(\<forall>v::'a::zero^3. P v) \<longleftrightarrow> (\<forall>x y z. P(vector[x, y, z]))" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 165 | apply auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 166 | apply (erule_tac x="v$1" in allE) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 167 | apply (erule_tac x="v$2" in allE) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 168 | apply (erule_tac x="v$3" in allE) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 169 | apply (subgoal_tac "vector [v$1, v$2, v$3] = v") | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 170 | apply simp | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 171 | apply (vector vector_def) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 172 | apply (simp add: forall_3) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 173 | done | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 174 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 175 | lemma range_vec1[simp]:"range vec1 = UNIV" apply(rule set_ext,rule) unfolding image_iff defer | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 176 | apply(rule_tac x="dest_vec1 x" in bexI) by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 177 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 178 | lemma dest_vec1_lambda: "dest_vec1(\<chi> i. x i) = x 1" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 179 | by (simp) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 180 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 181 | lemma dest_vec1_vec: "dest_vec1(vec x) = x" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 182 | by (simp) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 183 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 184 | lemma dest_vec1_sum: assumes fS: "finite S" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 185 | shows "dest_vec1(setsum f S) = setsum (dest_vec1 o f) S" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 186 | apply (induct rule: finite_induct[OF fS]) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 187 | apply simp | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 188 | apply auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 189 | done | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 190 | |
| 36433 | 191 | lemma norm_vec1 [simp]: "norm(vec1 x) = abs(x)" | 
| 36431 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 192 | by (simp add: vec_def norm_real) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 193 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 194 | lemma dist_vec1: "dist(vec1 x) (vec1 y) = abs(x - y)" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 195 | by (simp only: dist_real vec1_component) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 196 | lemma abs_dest_vec1: "norm x = \<bar>dest_vec1 x\<bar>" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 197 | by (metis vec1_dest_vec1(1) norm_vec1) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 198 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 199 | lemmas vec1_dest_vec1_simps = forall_vec1 vec_add[THEN sym] dist_vec1 vec_sub[THEN sym] vec1_dest_vec1 norm_vec1 vector_smult_component | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 200 | vec1_eq vec_cmul[THEN sym] smult_conv_scaleR[THEN sym] o_def dist_real_def norm_vec1 real_norm_def | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 201 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 202 | lemma bounded_linear_vec1:"bounded_linear (vec1::real\<Rightarrow>real^1)" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 203 | unfolding bounded_linear_def additive_def bounded_linear_axioms_def | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 204 | unfolding smult_conv_scaleR[THEN sym] unfolding vec1_dest_vec1_simps | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 205 | apply(rule conjI) defer apply(rule conjI) defer apply(rule_tac x=1 in exI) by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 206 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 207 | lemma linear_vmul_dest_vec1: | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 208 | fixes f:: "'a::semiring_1^_ \<Rightarrow> 'a^1" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 209 | shows "linear f \<Longrightarrow> linear (\<lambda>x. dest_vec1(f x) *s v)" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 210 | apply (rule linear_vmul_component) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 211 | by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 212 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 213 | lemma linear_from_scalars: | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 214 | assumes lf: "linear (f::'a::comm_ring_1 ^1 \<Rightarrow> 'a^_)" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 215 | shows "f = (\<lambda>x. dest_vec1 x *s column 1 (matrix f))" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 216 | apply (rule ext) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 217 | apply (subst matrix_works[OF lf, symmetric]) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 218 | apply (auto simp add: Cart_eq matrix_vector_mult_def column_def mult_commute) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 219 | done | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 220 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 221 | lemma linear_to_scalars: assumes lf: "linear (f::real ^'n \<Rightarrow> real^1)" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 222 | shows "f = (\<lambda>x. vec1(row 1 (matrix f) \<bullet> x))" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 223 | apply (rule ext) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 224 | apply (subst matrix_works[OF lf, symmetric]) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 225 | apply (simp add: Cart_eq matrix_vector_mult_def row_def inner_vector_def mult_commute) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 226 | done | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 227 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 228 | lemma dest_vec1_eq_0: "dest_vec1 x = 0 \<longleftrightarrow> x = 0" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 229 | by (simp add: dest_vec1_eq[symmetric]) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 230 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 231 | lemma setsum_scalars: assumes fS: "finite S" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 232 | shows "setsum f S = vec1 (setsum (dest_vec1 o f) S)" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 233 | unfolding vec_setsum[OF fS] by simp | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 234 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 235 | lemma dest_vec1_wlog_le: "(\<And>(x::'a::linorder ^ 1) y. P x y \<longleftrightarrow> P y x) \<Longrightarrow> (\<And>x y. dest_vec1 x <= dest_vec1 y ==> P x y) \<Longrightarrow> P x y" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 236 | apply (cases "dest_vec1 x \<le> dest_vec1 y") | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 237 | apply simp | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 238 | apply (subgoal_tac "dest_vec1 y \<le> dest_vec1 x") | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 239 | apply (auto) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 240 | done | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 241 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 242 | text{* Lifting and dropping *}
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 243 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 244 | lemma continuous_on_o_dest_vec1: fixes f::"real \<Rightarrow> 'a::real_normed_vector" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 245 |   assumes "continuous_on {a..b::real} f" shows "continuous_on {vec1 a..vec1 b} (f o dest_vec1)"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 246 | using assms unfolding continuous_on_iff apply safe | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 247 | apply(erule_tac x="x$1" in ballE,erule_tac x=e in allE) apply safe | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 248 | apply(rule_tac x=d in exI) apply safe unfolding o_def dist_real_def dist_real | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 249 | apply(erule_tac x="dest_vec1 x'" in ballE) by(auto simp add:vector_le_def) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 250 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 251 | lemma continuous_on_o_vec1: fixes f::"real^1 \<Rightarrow> 'a::real_normed_vector" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 252 |   assumes "continuous_on {a..b} f" shows "continuous_on {dest_vec1 a..dest_vec1 b} (f o vec1)"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 253 | using assms unfolding continuous_on_iff apply safe | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 254 | apply(erule_tac x="vec x" in ballE,erule_tac x=e in allE) apply safe | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 255 | apply(rule_tac x=d in exI) apply safe unfolding o_def dist_real_def dist_real | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 256 | apply(erule_tac x="vec1 x'" in ballE) by(auto simp add:vector_le_def) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 257 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 258 | lemma continuous_on_vec1:"continuous_on A (vec1::real\<Rightarrow>real^1)" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 259 | by(rule linear_continuous_on[OF bounded_linear_vec1]) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 260 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 261 | lemma mem_interval_1: fixes x :: "real^1" shows | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 262 |  "(x \<in> {a .. b} \<longleftrightarrow> dest_vec1 a \<le> dest_vec1 x \<and> dest_vec1 x \<le> dest_vec1 b)"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 263 |  "(x \<in> {a<..<b} \<longleftrightarrow> dest_vec1 a < dest_vec1 x \<and> dest_vec1 x < dest_vec1 b)"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 264 | by(simp_all add: Cart_eq vector_less_def vector_le_def) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 265 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 266 | lemma vec1_interval:fixes a::"real" shows | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 267 |   "vec1 ` {a .. b} = {vec1 a .. vec1 b}"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 268 |   "vec1 ` {a<..<b} = {vec1 a<..<vec1 b}"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 269 | apply(rule_tac[!] set_ext) unfolding image_iff vector_less_def unfolding mem_interval | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 270 | unfolding forall_1 unfolding vec1_dest_vec1_simps | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 271 | apply rule defer apply(rule_tac x="dest_vec1 x" in bexI) prefer 3 apply rule defer | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 272 | apply(rule_tac x="dest_vec1 x" in bexI) by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 273 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 274 | (* Some special cases for intervals in R^1. *) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 275 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 276 | lemma interval_cases_1: fixes x :: "real^1" shows | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 277 |  "x \<in> {a .. b} ==> x \<in> {a<..<b} \<or> (x = a) \<or> (x = b)"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 278 | unfolding Cart_eq vector_less_def vector_le_def mem_interval by(auto simp del:dest_vec1_eq) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 279 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 280 | lemma in_interval_1: fixes x :: "real^1" shows | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 281 |  "(x \<in> {a .. b} \<longleftrightarrow> dest_vec1 a \<le> dest_vec1 x \<and> dest_vec1 x \<le> dest_vec1 b) \<and>
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 282 |   (x \<in> {a<..<b} \<longleftrightarrow> dest_vec1 a < dest_vec1 x \<and> dest_vec1 x < dest_vec1 b)"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 283 | unfolding Cart_eq vector_less_def vector_le_def mem_interval by(auto simp del:dest_vec1_eq) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 284 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 285 | lemma interval_eq_empty_1: fixes a :: "real^1" shows | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 286 |   "{a .. b} = {} \<longleftrightarrow> dest_vec1 b < dest_vec1 a"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 287 |   "{a<..<b} = {} \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 288 | unfolding interval_eq_empty and ex_1 by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 289 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 290 | lemma subset_interval_1: fixes a :: "real^1" shows | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 291 |  "({a .. b} \<subseteq> {c .. d} \<longleftrightarrow>  dest_vec1 b < dest_vec1 a \<or>
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 292 | dest_vec1 c \<le> dest_vec1 a \<and> dest_vec1 a \<le> dest_vec1 b \<and> dest_vec1 b \<le> dest_vec1 d)" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 293 |  "({a .. b} \<subseteq> {c<..<d} \<longleftrightarrow>  dest_vec1 b < dest_vec1 a \<or>
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 294 | dest_vec1 c < dest_vec1 a \<and> dest_vec1 a \<le> dest_vec1 b \<and> dest_vec1 b < dest_vec1 d)" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 295 |  "({a<..<b} \<subseteq> {c .. d} \<longleftrightarrow>  dest_vec1 b \<le> dest_vec1 a \<or>
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 296 | dest_vec1 c \<le> dest_vec1 a \<and> dest_vec1 a < dest_vec1 b \<and> dest_vec1 b \<le> dest_vec1 d)" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 297 |  "({a<..<b} \<subseteq> {c<..<d} \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a \<or>
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 298 | dest_vec1 c \<le> dest_vec1 a \<and> dest_vec1 a < dest_vec1 b \<and> dest_vec1 b \<le> dest_vec1 d)" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 299 | unfolding subset_interval[of a b c d] unfolding forall_1 by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 300 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 301 | lemma eq_interval_1: fixes a :: "real^1" shows | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 302 |  "{a .. b} = {c .. d} \<longleftrightarrow>
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 303 | dest_vec1 b < dest_vec1 a \<and> dest_vec1 d < dest_vec1 c \<or> | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 304 | dest_vec1 a = dest_vec1 c \<and> dest_vec1 b = dest_vec1 d" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 305 | unfolding set_eq_subset[of "{a .. b}" "{c .. d}"]
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 306 | unfolding subset_interval_1(1)[of a b c d] | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 307 | unfolding subset_interval_1(1)[of c d a b] | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 308 | by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 309 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 310 | lemma disjoint_interval_1: fixes a :: "real^1" shows | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 311 |   "{a .. b} \<inter> {c .. d} = {} \<longleftrightarrow> dest_vec1 b < dest_vec1 a \<or> dest_vec1 d < dest_vec1 c  \<or>  dest_vec1 b < dest_vec1 c \<or> dest_vec1 d < dest_vec1 a"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 312 |   "{a .. b} \<inter> {c<..<d} = {} \<longleftrightarrow> dest_vec1 b < dest_vec1 a \<or> dest_vec1 d \<le> dest_vec1 c  \<or>  dest_vec1 b \<le> dest_vec1 c \<or> dest_vec1 d \<le> dest_vec1 a"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 313 |   "{a<..<b} \<inter> {c .. d} = {} \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a \<or> dest_vec1 d < dest_vec1 c  \<or>  dest_vec1 b \<le> dest_vec1 c \<or> dest_vec1 d \<le> dest_vec1 a"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 314 |   "{a<..<b} \<inter> {c<..<d} = {} \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a \<or> dest_vec1 d \<le> dest_vec1 c  \<or>  dest_vec1 b \<le> dest_vec1 c \<or> dest_vec1 d \<le> dest_vec1 a"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 315 | unfolding disjoint_interval and ex_1 by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 316 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 317 | lemma open_closed_interval_1: fixes a :: "real^1" shows | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 318 |  "{a<..<b} = {a .. b} - {a, b}"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 319 | unfolding expand_set_eq apply simp unfolding vector_less_def and vector_le_def and forall_1 and dest_vec1_eq[THEN sym] by(auto simp del:dest_vec1_eq) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 320 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 321 | lemma closed_open_interval_1: "dest_vec1 (a::real^1) \<le> dest_vec1 b ==> {a .. b} = {a<..<b} \<union> {a,b}"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 322 | unfolding expand_set_eq apply simp unfolding vector_less_def and vector_le_def and forall_1 and dest_vec1_eq[THEN sym] by(auto simp del:dest_vec1_eq) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 323 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 324 | lemma Lim_drop_le: fixes f :: "'a \<Rightarrow> real^1" shows | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 325 | "(f ---> l) net \<Longrightarrow> ~(trivial_limit net) \<Longrightarrow> eventually (\<lambda>x. dest_vec1 (f x) \<le> b) net ==> dest_vec1 l \<le> b" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 326 | using Lim_component_le[of f l net 1 b] by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 327 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 328 | lemma Lim_drop_ge: fixes f :: "'a \<Rightarrow> real^1" shows | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 329 | "(f ---> l) net \<Longrightarrow> ~(trivial_limit net) \<Longrightarrow> eventually (\<lambda>x. b \<le> dest_vec1 (f x)) net ==> b \<le> dest_vec1 l" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 330 | using Lim_component_ge[of f l net b 1] by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 331 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 332 | text{* Also more convenient formulations of monotone convergence.                *}
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 333 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 334 | lemma bounded_increasing_convergent: fixes s::"nat \<Rightarrow> real^1" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 335 |   assumes "bounded {s n| n::nat. True}"  "\<forall>n. dest_vec1(s n) \<le> dest_vec1(s(Suc n))"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 336 | shows "\<exists>l. (s ---> l) sequentially" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 337 | proof- | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 338 | obtain a where a:"\<forall>n. \<bar>dest_vec1 (s n)\<bar> \<le> a" using assms(1)[unfolded bounded_iff abs_dest_vec1] by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 339 |   { fix m::nat
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 340 | have "\<And> n. n\<ge>m \<longrightarrow> dest_vec1 (s m) \<le> dest_vec1 (s n)" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 341 | apply(induct_tac n) apply simp using assms(2) apply(erule_tac x="na" in allE) by(auto simp add: not_less_eq_eq) } | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 342 | hence "\<forall>m n. m \<le> n \<longrightarrow> dest_vec1 (s m) \<le> dest_vec1 (s n)" by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 343 | then obtain l where "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<bar>dest_vec1 (s n) - l\<bar> < e" using convergent_bounded_monotone[OF a] unfolding monoseq_def by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 344 | thus ?thesis unfolding Lim_sequentially apply(rule_tac x="vec1 l" in exI) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 345 | unfolding dist_norm unfolding abs_dest_vec1 by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 346 | qed | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 347 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 348 | lemma dest_vec1_simps[simp]: fixes a::"real^1" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 349 | shows "a$1 = 0 \<longleftrightarrow> a = 0" (*"a \<le> 1 \<longleftrightarrow> dest_vec1 a \<le> 1" "0 \<le> a \<longleftrightarrow> 0 \<le> dest_vec1 a"*) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 350 | "a \<le> b \<longleftrightarrow> dest_vec1 a \<le> dest_vec1 b" "dest_vec1 (1::real^1) = 1" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 351 | by(auto simp add: vector_le_def Cart_eq) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 352 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 353 | lemma dest_vec1_inverval: | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 354 |   "dest_vec1 ` {a .. b} = {dest_vec1 a .. dest_vec1 b}"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 355 |   "dest_vec1 ` {a<.. b} = {dest_vec1 a<.. dest_vec1 b}"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 356 |   "dest_vec1 ` {a ..<b} = {dest_vec1 a ..<dest_vec1 b}"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 357 |   "dest_vec1 ` {a<..<b} = {dest_vec1 a<..<dest_vec1 b}"
 | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 358 | apply(rule_tac [!] equalityI) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 359 | unfolding subset_eq Ball_def Bex_def mem_interval_1 image_iff | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 360 | apply(rule_tac [!] allI)apply(rule_tac [!] impI) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 361 | apply(rule_tac[2] x="vec1 x" in exI)apply(rule_tac[4] x="vec1 x" in exI) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 362 | apply(rule_tac[6] x="vec1 x" in exI)apply(rule_tac[8] x="vec1 x" in exI) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 363 | by (auto simp add: vector_less_def vector_le_def) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 364 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 365 | lemma dest_vec1_setsum: assumes "finite S" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 366 | shows " dest_vec1 (setsum f S) = setsum (\<lambda>x. dest_vec1 (f x)) S" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 367 | using dest_vec1_sum[OF assms] by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 368 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 369 | lemma open_dest_vec1_vimage: "open S \<Longrightarrow> open (dest_vec1 -` S)" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 370 | unfolding open_vector_def forall_1 by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 371 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 372 | lemma tendsto_dest_vec1 [tendsto_intros]: | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 373 | "(f ---> l) net \<Longrightarrow> ((\<lambda>x. dest_vec1 (f x)) ---> dest_vec1 l) net" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 374 | by(rule tendsto_Cart_nth) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 375 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 376 | lemma continuous_dest_vec1: "continuous net f \<Longrightarrow> continuous net (\<lambda>x. dest_vec1 (f x))" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 377 | unfolding continuous_def by (rule tendsto_dest_vec1) | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 378 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 379 | lemma forall_dest_vec1: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P(dest_vec1 x))" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 380 | apply safe defer apply(erule_tac x="vec1 x" in allE) by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 381 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 382 | lemma forall_of_dest_vec1: "(\<forall>v. P (\<lambda>x. dest_vec1 (v x))) \<longleftrightarrow> (\<forall>x. P x)" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 383 | apply rule apply rule apply(erule_tac x="(vec1 \<circ> x)" in allE) unfolding o_def vec1_dest_vec1 by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 384 | |
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 385 | lemma forall_of_dest_vec1': "(\<forall>v. P (dest_vec1 v)) \<longleftrightarrow> (\<forall>x. P x)" | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 386 | apply rule apply rule apply(erule_tac x="(vec1 x)" in allE) defer apply rule | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 387 | apply(erule_tac x="dest_vec1 v" in allE) unfolding o_def vec1_dest_vec1 by auto | 
| 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 388 | |
| 36433 | 389 | lemma dist_vec1_0[simp]: "dist(vec1 (x::real)) 0 = norm x" unfolding vector_dist_norm by auto | 
| 390 | ||
| 391 | lemma bounded_linear_vec1_dest_vec1: fixes f::"real \<Rightarrow> real" | |
| 392 | shows "linear (vec1 \<circ> f \<circ> dest_vec1) = bounded_linear f" (is "?l = ?r") proof- | |
| 393 |   { assume ?l guess K using linear_bounded[OF `?l`] ..
 | |
| 394 | hence "\<exists>K. \<forall>x. \<bar>f x\<bar> \<le> \<bar>x\<bar> * K" apply(rule_tac x=K in exI) | |
| 395 | unfolding vec1_dest_vec1_simps by (auto simp add:field_simps) } | |
| 396 | thus ?thesis unfolding linear_def bounded_linear_def additive_def bounded_linear_axioms_def o_def | |
| 36435 | 397 | unfolding vec1_dest_vec1_simps by auto qed | 
| 398 | ||
| 399 | lemma vec1_le[simp]:fixes a::real shows "vec1 a \<le> vec1 b \<longleftrightarrow> a \<le> b" | |
| 400 | unfolding vector_le_def by auto | |
| 401 | lemma vec1_less[simp]:fixes a::real shows "vec1 a < vec1 b \<longleftrightarrow> a < b" | |
| 402 | unfolding vector_less_def by auto | |
| 36433 | 403 | |
| 36431 
340755027840
move definitions and theorems for type real^1 to separate theory file
 huffman parents: diff
changeset | 404 | end |