| author | haftmann | 
| Sat, 20 Oct 2012 09:12:16 +0200 | |
| changeset 49948 | 744934b818c7 | 
| parent 49834 | b27bbb021df1 | 
| child 50252 | 4aa34bd43228 | 
| permissions | -rw-r--r-- | 
| 35253 | 1 | (* Title: HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy | 
| 2 | Author: Amine Chaieb, University of Cambridge | |
| 33175 | 3 | *) | 
| 4 | ||
| 5 | header {* Definition of finite Cartesian product types. *}
 | |
| 6 | ||
| 7 | theory Finite_Cartesian_Product | |
| 41413 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
39302diff
changeset | 8 | imports | 
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 9 | Euclidean_Space | 
| 41413 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
39302diff
changeset | 10 | L2_Norm | 
| 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
39302diff
changeset | 11 | "~~/src/HOL/Library/Numeral_Type" | 
| 33175 | 12 | begin | 
| 13 | ||
| 14 | subsection {* Finite Cartesian products, with indexing and lambdas. *}
 | |
| 15 | ||
| 49834 | 16 | typedef ('a, 'b) vec = "UNIV :: (('b::finite) \<Rightarrow> 'a) set"
 | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 17 | morphisms vec_nth vec_lambda .. | 
| 33175 | 18 | |
| 35254 
0f17eda72e60
binder notation for default print_mode -- to avoid strange output if "xsymbols" is not active;
 wenzelm parents: 
35253diff
changeset | 19 | notation | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 20 | vec_nth (infixl "$" 90) and | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 21 | vec_lambda (binder "\<chi>" 10) | 
| 33175 | 22 | |
| 34290 
1edf0f223c6e
added syntax translation to automatically add finite typeclass to index type of cartesian product type
 hoelzl parents: 
34289diff
changeset | 23 | (* | 
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34290diff
changeset | 24 |   Translate "'b ^ 'n" into "'b ^ ('n :: finite)". When 'n has already more than
 | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 25 | the finite type class write "vec 'b 'n" | 
| 34290 
1edf0f223c6e
added syntax translation to automatically add finite typeclass to index type of cartesian product type
 hoelzl parents: 
34289diff
changeset | 26 | *) | 
| 
1edf0f223c6e
added syntax translation to automatically add finite typeclass to index type of cartesian product type
 hoelzl parents: 
34289diff
changeset | 27 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 28 | syntax "_finite_vec" :: "type \<Rightarrow> type \<Rightarrow> type" ("(_ ^/ _)" [15, 16] 15)
 | 
| 34290 
1edf0f223c6e
added syntax translation to automatically add finite typeclass to index type of cartesian product type
 hoelzl parents: 
34289diff
changeset | 29 | |
| 
1edf0f223c6e
added syntax translation to automatically add finite typeclass to index type of cartesian product type
 hoelzl parents: 
34289diff
changeset | 30 | parse_translation {*
 | 
| 
1edf0f223c6e
added syntax translation to automatically add finite typeclass to index type of cartesian product type
 hoelzl parents: 
34289diff
changeset | 31 | let | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 32 |   fun vec t u = Syntax.const @{type_syntax vec} $ t $ u;
 | 
| 49674 
dbadb4d03cbc
report sort assignment of visible type variables;
 wenzelm parents: 
45694diff
changeset | 33 | fun finite_vec_tr [t, u] = | 
| 
dbadb4d03cbc
report sort assignment of visible type variables;
 wenzelm parents: 
45694diff
changeset | 34 | (case Term_Position.strip_positions u of | 
| 
dbadb4d03cbc
report sort assignment of visible type variables;
 wenzelm parents: 
45694diff
changeset | 35 | v as Free (x, _) => | 
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
41413diff
changeset | 36 | if Lexicon.is_tid x then | 
| 49674 
dbadb4d03cbc
report sort assignment of visible type variables;
 wenzelm parents: 
45694diff
changeset | 37 |           vec t (Syntax.const @{syntax_const "_ofsort"} $ v $ Syntax.const @{class_syntax finite})
 | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 38 | else vec t u | 
| 49674 
dbadb4d03cbc
report sort assignment of visible type variables;
 wenzelm parents: 
45694diff
changeset | 39 | | _ => vec t u) | 
| 34290 
1edf0f223c6e
added syntax translation to automatically add finite typeclass to index type of cartesian product type
 hoelzl parents: 
34289diff
changeset | 40 | in | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 41 |   [(@{syntax_const "_finite_vec"}, finite_vec_tr)]
 | 
| 34290 
1edf0f223c6e
added syntax translation to automatically add finite typeclass to index type of cartesian product type
 hoelzl parents: 
34289diff
changeset | 42 | end | 
| 
1edf0f223c6e
added syntax translation to automatically add finite typeclass to index type of cartesian product type
 hoelzl parents: 
34289diff
changeset | 43 | *} | 
| 
1edf0f223c6e
added syntax translation to automatically add finite typeclass to index type of cartesian product type
 hoelzl parents: 
34289diff
changeset | 44 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 45 | lemma vec_eq_iff: "(x = y) \<longleftrightarrow> (\<forall>i. x$i = y$i)" | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 46 | by (simp add: vec_nth_inject [symmetric] fun_eq_iff) | 
| 33175 | 47 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 48 | lemma vec_lambda_beta [simp]: "vec_lambda g $ i = g i" | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 49 | by (simp add: vec_lambda_inverse) | 
| 33175 | 50 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 51 | lemma vec_lambda_unique: "(\<forall>i. f$i = g i) \<longleftrightarrow> vec_lambda g = f" | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 52 | by (auto simp add: vec_eq_iff) | 
| 33175 | 53 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 54 | lemma vec_lambda_eta: "(\<chi> i. (g$i)) = g" | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 55 | by (simp add: vec_eq_iff) | 
| 33175 | 56 | |
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 57 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 58 | subsection {* Group operations and class instances *}
 | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 59 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 60 | instantiation vec :: (zero, finite) zero | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 61 | begin | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 62 | definition "0 \<equiv> (\<chi> i. 0)" | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 63 | instance .. | 
| 33175 | 64 | end | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 65 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 66 | instantiation vec :: (plus, finite) plus | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 67 | begin | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 68 | definition "op + \<equiv> (\<lambda> x y. (\<chi> i. x$i + y$i))" | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 69 | instance .. | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 70 | end | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 71 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 72 | instantiation vec :: (minus, finite) minus | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 73 | begin | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 74 | definition "op - \<equiv> (\<lambda> x y. (\<chi> i. x$i - y$i))" | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 75 | instance .. | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 76 | end | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 77 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 78 | instantiation vec :: (uminus, finite) uminus | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 79 | begin | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 80 | definition "uminus \<equiv> (\<lambda> x. (\<chi> i. - (x$i)))" | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 81 | instance .. | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 82 | end | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 83 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 84 | lemma zero_index [simp]: "0 $ i = 0" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 85 | unfolding zero_vec_def by simp | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 86 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 87 | lemma vector_add_component [simp]: "(x + y)$i = x$i + y$i" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 88 | unfolding plus_vec_def by simp | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 89 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 90 | lemma vector_minus_component [simp]: "(x - y)$i = x$i - y$i" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 91 | unfolding minus_vec_def by simp | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 92 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 93 | lemma vector_uminus_component [simp]: "(- x)$i = - (x$i)" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 94 | unfolding uminus_vec_def by simp | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 95 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 96 | instance vec :: (semigroup_add, finite) semigroup_add | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 97 | by default (simp add: vec_eq_iff add_assoc) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 98 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 99 | instance vec :: (ab_semigroup_add, finite) ab_semigroup_add | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 100 | by default (simp add: vec_eq_iff add_commute) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 101 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 102 | instance vec :: (monoid_add, finite) monoid_add | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 103 | by default (simp_all add: vec_eq_iff) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 104 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 105 | instance vec :: (comm_monoid_add, finite) comm_monoid_add | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 106 | by default (simp add: vec_eq_iff) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 107 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 108 | instance vec :: (cancel_semigroup_add, finite) cancel_semigroup_add | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 109 | by default (simp_all add: vec_eq_iff) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 110 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 111 | instance vec :: (cancel_ab_semigroup_add, finite) cancel_ab_semigroup_add | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 112 | by default (simp add: vec_eq_iff) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 113 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 114 | instance vec :: (cancel_comm_monoid_add, finite) cancel_comm_monoid_add .. | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 115 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 116 | instance vec :: (group_add, finite) group_add | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 117 | by default (simp_all add: vec_eq_iff diff_minus) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 118 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 119 | instance vec :: (ab_group_add, finite) ab_group_add | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 120 | by default (simp_all add: vec_eq_iff) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 121 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 122 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 123 | subsection {* Real vector space *}
 | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 124 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 125 | instantiation vec :: (real_vector, finite) real_vector | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 126 | begin | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 127 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 128 | definition "scaleR \<equiv> (\<lambda> r x. (\<chi> i. scaleR r (x$i)))" | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 129 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 130 | lemma vector_scaleR_component [simp]: "(scaleR r x)$i = scaleR r (x$i)" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 131 | unfolding scaleR_vec_def by simp | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 132 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 133 | instance | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 134 | by default (simp_all add: vec_eq_iff scaleR_left_distrib scaleR_right_distrib) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 135 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 136 | end | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 137 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 138 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 139 | subsection {* Topological space *}
 | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 140 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 141 | instantiation vec :: (topological_space, finite) topological_space | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 142 | begin | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 143 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 144 | definition | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 145 |   "open (S :: ('a ^ 'b) set) \<longleftrightarrow>
 | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 146 | (\<forall>x\<in>S. \<exists>A. (\<forall>i. open (A i) \<and> x$i \<in> A i) \<and> | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 147 | (\<forall>y. (\<forall>i. y$i \<in> A i) \<longrightarrow> y \<in> S))" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 148 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 149 | instance proof | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 150 |   show "open (UNIV :: ('a ^ 'b) set)"
 | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 151 | unfolding open_vec_def by auto | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 152 | next | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 153 |   fix S T :: "('a ^ 'b) set"
 | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 154 | assume "open S" "open T" thus "open (S \<inter> T)" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 155 | unfolding open_vec_def | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 156 | apply clarify | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 157 | apply (drule (1) bspec)+ | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 158 | apply (clarify, rename_tac Sa Ta) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 159 | apply (rule_tac x="\<lambda>i. Sa i \<inter> Ta i" in exI) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 160 | apply (simp add: open_Int) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 161 | done | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 162 | next | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 163 |   fix K :: "('a ^ 'b) set set"
 | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 164 | assume "\<forall>S\<in>K. open S" thus "open (\<Union>K)" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 165 | unfolding open_vec_def | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 166 | apply clarify | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 167 | apply (drule (1) bspec) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 168 | apply (drule (1) bspec) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 169 | apply clarify | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 170 | apply (rule_tac x=A in exI) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 171 | apply fast | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 172 | done | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 173 | qed | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 174 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 175 | end | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 176 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 177 | lemma open_vector_box: "\<forall>i. open (S i) \<Longrightarrow> open {x. \<forall>i. x $ i \<in> S i}"
 | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 178 | unfolding open_vec_def by auto | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 179 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 180 | lemma open_vimage_vec_nth: "open S \<Longrightarrow> open ((\<lambda>x. x $ i) -` S)" | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 181 | unfolding open_vec_def | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 182 | apply clarify | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 183 | apply (rule_tac x="\<lambda>k. if k = i then S else UNIV" in exI, simp) | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 184 | done | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 185 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 186 | lemma closed_vimage_vec_nth: "closed S \<Longrightarrow> closed ((\<lambda>x. x $ i) -` S)" | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 187 | unfolding closed_open vimage_Compl [symmetric] | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 188 | by (rule open_vimage_vec_nth) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 189 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 190 | lemma closed_vector_box: "\<forall>i. closed (S i) \<Longrightarrow> closed {x. \<forall>i. x $ i \<in> S i}"
 | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 191 | proof - | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 192 |   have "{x. \<forall>i. x $ i \<in> S i} = (\<Inter>i. (\<lambda>x. x $ i) -` S i)" by auto
 | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 193 |   thus "\<forall>i. closed (S i) \<Longrightarrow> closed {x. \<forall>i. x $ i \<in> S i}"
 | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 194 | by (simp add: closed_INT closed_vimage_vec_nth) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 195 | qed | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 196 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 197 | lemma tendsto_vec_nth [tendsto_intros]: | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 198 | assumes "((\<lambda>x. f x) ---> a) net" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 199 | shows "((\<lambda>x. f x $ i) ---> a $ i) net" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 200 | proof (rule topological_tendstoI) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 201 | fix S assume "open S" "a $ i \<in> S" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 202 | then have "open ((\<lambda>y. y $ i) -` S)" "a \<in> ((\<lambda>y. y $ i) -` S)" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 203 | by (simp_all add: open_vimage_vec_nth) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 204 | with assms have "eventually (\<lambda>x. f x \<in> (\<lambda>y. y $ i) -` S) net" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 205 | by (rule topological_tendstoD) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 206 | then show "eventually (\<lambda>x. f x $ i \<in> S) net" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 207 | by simp | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 208 | qed | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 209 | |
| 44631 | 210 | lemma isCont_vec_nth [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. f x $ i) a" | 
| 211 | unfolding isCont_def by (rule tendsto_vec_nth) | |
| 212 | ||
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 213 | lemma eventually_Ball_finite: (* TODO: move *) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 214 | assumes "finite A" and "\<forall>y\<in>A. eventually (\<lambda>x. P x y) net" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 215 | shows "eventually (\<lambda>x. \<forall>y\<in>A. P x y) net" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 216 | using assms by (induct set: finite, simp, simp add: eventually_conj) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 217 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 218 | lemma eventually_all_finite: (* TODO: move *) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 219 | fixes P :: "'a \<Rightarrow> 'b::finite \<Rightarrow> bool" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 220 | assumes "\<And>y. eventually (\<lambda>x. P x y) net" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 221 | shows "eventually (\<lambda>x. \<forall>y. P x y) net" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 222 | using eventually_Ball_finite [of UNIV P] assms by simp | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 223 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 224 | lemma vec_tendstoI: | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 225 | assumes "\<And>i. ((\<lambda>x. f x $ i) ---> a $ i) net" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 226 | shows "((\<lambda>x. f x) ---> a) net" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 227 | proof (rule topological_tendstoI) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 228 | fix S assume "open S" and "a \<in> S" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 229 | then obtain A where A: "\<And>i. open (A i)" "\<And>i. a $ i \<in> A i" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 230 | and S: "\<And>y. \<forall>i. y $ i \<in> A i \<Longrightarrow> y \<in> S" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 231 | unfolding open_vec_def by metis | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 232 | have "\<And>i. eventually (\<lambda>x. f x $ i \<in> A i) net" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 233 | using assms A by (rule topological_tendstoD) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 234 | hence "eventually (\<lambda>x. \<forall>i. f x $ i \<in> A i) net" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 235 | by (rule eventually_all_finite) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 236 | thus "eventually (\<lambda>x. f x \<in> S) net" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 237 | by (rule eventually_elim1, simp add: S) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 238 | qed | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 239 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 240 | lemma tendsto_vec_lambda [tendsto_intros]: | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 241 | assumes "\<And>i. ((\<lambda>x. f x i) ---> a i) net" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 242 | shows "((\<lambda>x. \<chi> i. f x i) ---> (\<chi> i. a i)) net" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 243 | using assms by (simp add: vec_tendstoI) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 244 | |
| 44571 | 245 | lemma open_image_vec_nth: assumes "open S" shows "open ((\<lambda>x. x $ i) ` S)" | 
| 246 | proof (rule openI) | |
| 247 | fix a assume "a \<in> (\<lambda>x. x $ i) ` S" | |
| 248 | then obtain z where "a = z $ i" and "z \<in> S" .. | |
| 249 | then obtain A where A: "\<forall>i. open (A i) \<and> z $ i \<in> A i" | |
| 250 | and S: "\<forall>y. (\<forall>i. y $ i \<in> A i) \<longrightarrow> y \<in> S" | |
| 251 | using `open S` unfolding open_vec_def by auto | |
| 252 | hence "A i \<subseteq> (\<lambda>x. x $ i) ` S" | |
| 253 | by (clarsimp, rule_tac x="\<chi> j. if j = i then x else z $ j" in image_eqI, | |
| 254 | simp_all) | |
| 255 | hence "open (A i) \<and> a \<in> A i \<and> A i \<subseteq> (\<lambda>x. x $ i) ` S" | |
| 256 | using A `a = z $ i` by simp | |
| 257 | then show "\<exists>T. open T \<and> a \<in> T \<and> T \<subseteq> (\<lambda>x. x $ i) ` S" by - (rule exI) | |
| 258 | qed | |
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 259 | |
| 44571 | 260 | instance vec :: (perfect_space, finite) perfect_space | 
| 261 | proof | |
| 262 |   fix x :: "'a ^ 'b" show "\<not> open {x}"
 | |
| 263 | proof | |
| 264 |     assume "open {x}"
 | |
| 265 |     hence "\<forall>i. open ((\<lambda>x. x $ i) ` {x})" by (fast intro: open_image_vec_nth)   
 | |
| 266 |     hence "\<forall>i. open {x $ i}" by simp
 | |
| 267 | thus "False" by (simp add: not_open_singleton) | |
| 268 | qed | |
| 269 | qed | |
| 270 | ||
| 271 | ||
| 272 | subsection {* Metric space *}
 | |
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 273 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 274 | instantiation vec :: (metric_space, finite) metric_space | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 275 | begin | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 276 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 277 | definition | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 278 | "dist x y = setL2 (\<lambda>i. dist (x$i) (y$i)) UNIV" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 279 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 280 | lemma dist_vec_nth_le: "dist (x $ i) (y $ i) \<le> dist x y" | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 281 | unfolding dist_vec_def by (rule member_le_setL2) simp_all | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 282 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 283 | instance proof | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 284 | fix x y :: "'a ^ 'b" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 285 | show "dist x y = 0 \<longleftrightarrow> x = y" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 286 | unfolding dist_vec_def | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 287 | by (simp add: setL2_eq_0_iff vec_eq_iff) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 288 | next | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 289 | fix x y z :: "'a ^ 'b" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 290 | show "dist x y \<le> dist x z + dist y z" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 291 | unfolding dist_vec_def | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 292 | apply (rule order_trans [OF _ setL2_triangle_ineq]) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 293 | apply (simp add: setL2_mono dist_triangle2) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 294 | done | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 295 | next | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 296 |   fix S :: "('a ^ 'b) set"
 | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 297 | show "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)" | 
| 44630 | 298 | proof | 
| 299 | assume "open S" show "\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S" | |
| 300 | proof | |
| 301 | fix x assume "x \<in> S" | |
| 302 | obtain A where A: "\<forall>i. open (A i)" "\<forall>i. x $ i \<in> A i" | |
| 303 | and S: "\<forall>y. (\<forall>i. y $ i \<in> A i) \<longrightarrow> y \<in> S" | |
| 304 | using `open S` and `x \<in> S` unfolding open_vec_def by metis | |
| 305 | have "\<forall>i\<in>UNIV. \<exists>r>0. \<forall>y. dist y (x $ i) < r \<longrightarrow> y \<in> A i" | |
| 306 | using A unfolding open_dist by simp | |
| 307 | hence "\<exists>r. \<forall>i\<in>UNIV. 0 < r i \<and> (\<forall>y. dist y (x $ i) < r i \<longrightarrow> y \<in> A i)" | |
| 44681 
49ef76b4a634
remove duplicate lemma finite_choice in favor of finite_set_choice
 huffman parents: 
44631diff
changeset | 308 | by (rule finite_set_choice [OF finite]) | 
| 44630 | 309 | then obtain r where r1: "\<forall>i. 0 < r i" | 
| 310 | and r2: "\<forall>i y. dist y (x $ i) < r i \<longrightarrow> y \<in> A i" by fast | |
| 311 | have "0 < Min (range r) \<and> (\<forall>y. dist y x < Min (range r) \<longrightarrow> y \<in> S)" | |
| 312 | by (simp add: r1 r2 S le_less_trans [OF dist_vec_nth_le]) | |
| 313 | thus "\<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S" .. | |
| 314 | qed | |
| 315 | next | |
| 316 | assume *: "\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S" show "open S" | |
| 317 | proof (unfold open_vec_def, rule) | |
| 318 | fix x assume "x \<in> S" | |
| 319 | then obtain e where "0 < e" and S: "\<forall>y. dist y x < e \<longrightarrow> y \<in> S" | |
| 320 | using * by fast | |
| 321 |       def r \<equiv> "\<lambda>i::'b. e / sqrt (of_nat CARD('b))"
 | |
| 322 | from `0 < e` have r: "\<forall>i. 0 < r i" | |
| 323 | unfolding r_def by (simp_all add: divide_pos_pos) | |
| 324 | from `0 < e` have e: "e = setL2 r UNIV" | |
| 325 | unfolding r_def by (simp add: setL2_constant) | |
| 326 |       def A \<equiv> "\<lambda>i. {y. dist (x $ i) y < r i}"
 | |
| 327 | have "\<forall>i. open (A i) \<and> x $ i \<in> A i" | |
| 328 | unfolding A_def by (simp add: open_ball r) | |
| 329 | moreover have "\<forall>y. (\<forall>i. y $ i \<in> A i) \<longrightarrow> y \<in> S" | |
| 330 | by (simp add: A_def S dist_vec_def e setL2_strict_mono dist_commute) | |
| 331 | ultimately show "\<exists>A. (\<forall>i. open (A i) \<and> x $ i \<in> A i) \<and> | |
| 332 | (\<forall>y. (\<forall>i. y $ i \<in> A i) \<longrightarrow> y \<in> S)" by metis | |
| 333 | qed | |
| 334 | qed | |
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 335 | qed | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 336 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 337 | end | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 338 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 339 | lemma Cauchy_vec_nth: | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 340 | "Cauchy (\<lambda>n. X n) \<Longrightarrow> Cauchy (\<lambda>n. X n $ i)" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 341 | unfolding Cauchy_def by (fast intro: le_less_trans [OF dist_vec_nth_le]) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 342 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 343 | lemma vec_CauchyI: | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 344 | fixes X :: "nat \<Rightarrow> 'a::metric_space ^ 'n" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 345 | assumes X: "\<And>i. Cauchy (\<lambda>n. X n $ i)" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 346 | shows "Cauchy (\<lambda>n. X n)" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 347 | proof (rule metric_CauchyI) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 348 | fix r :: real assume "0 < r" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 349 |   then have "0 < r / of_nat CARD('n)" (is "0 < ?s")
 | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 350 | by (simp add: divide_pos_pos) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 351 | def N \<equiv> "\<lambda>i. LEAST N. \<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m $ i) (X n $ i) < ?s" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 352 | def M \<equiv> "Max (range N)" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 353 | have "\<And>i. \<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. dist (X m $ i) (X n $ i) < ?s" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 354 | using X `0 < ?s` by (rule metric_CauchyD) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 355 | hence "\<And>i. \<forall>m\<ge>N i. \<forall>n\<ge>N i. dist (X m $ i) (X n $ i) < ?s" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 356 | unfolding N_def by (rule LeastI_ex) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 357 | hence M: "\<And>i. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m $ i) (X n $ i) < ?s" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 358 | unfolding M_def by simp | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 359 |   {
 | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 360 | fix m n :: nat | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 361 | assume "M \<le> m" "M \<le> n" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 362 | have "dist (X m) (X n) = setL2 (\<lambda>i. dist (X m $ i) (X n $ i)) UNIV" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 363 | unfolding dist_vec_def .. | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 364 | also have "\<dots> \<le> setsum (\<lambda>i. dist (X m $ i) (X n $ i)) UNIV" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 365 | by (rule setL2_le_setsum [OF zero_le_dist]) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 366 | also have "\<dots> < setsum (\<lambda>i::'n. ?s) UNIV" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 367 | by (rule setsum_strict_mono, simp_all add: M `M \<le> m` `M \<le> n`) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 368 | also have "\<dots> = r" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 369 | by simp | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 370 | finally have "dist (X m) (X n) < r" . | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 371 | } | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 372 | hence "\<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < r" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 373 | by simp | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 374 | then show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (X m) (X n) < r" .. | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 375 | qed | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 376 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 377 | instance vec :: (complete_space, finite) complete_space | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 378 | proof | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 379 | fix X :: "nat \<Rightarrow> 'a ^ 'b" assume "Cauchy X" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 380 | have "\<And>i. (\<lambda>n. X n $ i) ----> lim (\<lambda>n. X n $ i)" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 381 | using Cauchy_vec_nth [OF `Cauchy X`] | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 382 | by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 383 | hence "X ----> vec_lambda (\<lambda>i. lim (\<lambda>n. X n $ i))" | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 384 | by (simp add: vec_tendstoI) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 385 | then show "convergent X" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 386 | by (rule convergentI) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 387 | qed | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 388 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 389 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 390 | subsection {* Normed vector space *}
 | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 391 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 392 | instantiation vec :: (real_normed_vector, finite) real_normed_vector | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 393 | begin | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 394 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 395 | definition "norm x = setL2 (\<lambda>i. norm (x$i)) UNIV" | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 396 | |
| 44141 | 397 | definition "sgn (x::'a^'b) = scaleR (inverse (norm x)) x" | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 398 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 399 | instance proof | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 400 | fix a :: real and x y :: "'a ^ 'b" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 401 | show "0 \<le> norm x" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 402 | unfolding norm_vec_def | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 403 | by (rule setL2_nonneg) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 404 | show "norm x = 0 \<longleftrightarrow> x = 0" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 405 | unfolding norm_vec_def | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 406 | by (simp add: setL2_eq_0_iff vec_eq_iff) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 407 | show "norm (x + y) \<le> norm x + norm y" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 408 | unfolding norm_vec_def | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 409 | apply (rule order_trans [OF _ setL2_triangle_ineq]) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 410 | apply (simp add: setL2_mono norm_triangle_ineq) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 411 | done | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 412 | show "norm (scaleR a x) = \<bar>a\<bar> * norm x" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 413 | unfolding norm_vec_def | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 414 | by (simp add: setL2_right_distrib) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 415 | show "sgn x = scaleR (inverse (norm x)) x" | 
| 44141 | 416 | by (rule sgn_vec_def) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 417 | show "dist x y = norm (x - y)" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 418 | unfolding dist_vec_def norm_vec_def | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 419 | by (simp add: dist_norm) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 420 | qed | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 421 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 422 | end | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 423 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 424 | lemma norm_nth_le: "norm (x $ i) \<le> norm x" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 425 | unfolding norm_vec_def | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 426 | by (rule member_le_setL2) simp_all | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 427 | |
| 44282 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 huffman parents: 
44233diff
changeset | 428 | lemma bounded_linear_vec_nth: "bounded_linear (\<lambda>x. x $ i)" | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 429 | apply default | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 430 | apply (rule vector_add_component) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 431 | apply (rule vector_scaleR_component) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 432 | apply (rule_tac x="1" in exI, simp add: norm_nth_le) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 433 | done | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 434 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 435 | instance vec :: (banach, finite) banach .. | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 436 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 437 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 438 | subsection {* Inner product space *}
 | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 439 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 440 | instantiation vec :: (real_inner, finite) real_inner | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 441 | begin | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 442 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 443 | definition "inner x y = setsum (\<lambda>i. inner (x$i) (y$i)) UNIV" | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 444 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 445 | instance proof | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 446 | fix r :: real and x y z :: "'a ^ 'b" | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 447 | show "inner x y = inner y x" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 448 | unfolding inner_vec_def | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 449 | by (simp add: inner_commute) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 450 | show "inner (x + y) z = inner x z + inner y z" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 451 | unfolding inner_vec_def | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 452 | by (simp add: inner_add_left setsum_addf) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 453 | show "inner (scaleR r x) y = r * inner x y" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 454 | unfolding inner_vec_def | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 455 | by (simp add: setsum_right_distrib) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 456 | show "0 \<le> inner x x" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 457 | unfolding inner_vec_def | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 458 | by (simp add: setsum_nonneg) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 459 | show "inner x x = 0 \<longleftrightarrow> x = 0" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 460 | unfolding inner_vec_def | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 461 | by (simp add: vec_eq_iff setsum_nonneg_eq_0_iff) | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 462 | show "norm x = sqrt (inner x x)" | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 463 | unfolding inner_vec_def norm_vec_def setL2_def | 
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 464 | by (simp add: power2_norm_eq_inner) | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 465 | qed | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 466 | |
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 467 | end | 
| 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 468 | |
| 44166 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 469 | |
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 470 | subsection {* Euclidean space *}
 | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 471 | |
| 44166 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 472 | text {* Vectors pointing along a single axis. *}
 | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 473 | |
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 474 | definition "axis k x = (\<chi> i. if i = k then x else 0)" | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 475 | |
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 476 | lemma axis_nth [simp]: "axis i x $ i = x" | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 477 | unfolding axis_def by simp | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 478 | |
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 479 | lemma axis_eq_axis: "axis i x = axis j y \<longleftrightarrow> x = y \<and> i = j \<or> x = 0 \<and> y = 0" | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 480 | unfolding axis_def vec_eq_iff by auto | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 481 | |
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 482 | lemma inner_axis_axis: | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 483 | "inner (axis i x) (axis j y) = (if i = j then inner x y else 0)" | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 484 | unfolding inner_vec_def | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 485 | apply (cases "i = j") | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 486 | apply clarsimp | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 487 | apply (subst setsum_diff1' [where a=j], simp_all) | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 488 | apply (rule setsum_0', simp add: axis_def) | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 489 | apply (rule setsum_0', simp add: axis_def) | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 490 | done | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 491 | |
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 492 | lemma setsum_single: | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 493 | assumes "finite A" and "k \<in> A" and "f k = y" | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 494 | assumes "\<And>i. i \<in> A \<Longrightarrow> i \<noteq> k \<Longrightarrow> f i = 0" | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 495 | shows "(\<Sum>i\<in>A. f i) = y" | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 496 | apply (subst setsum_diff1' [OF assms(1,2)]) | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 497 | apply (simp add: setsum_0' assms(3,4)) | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 498 | done | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 499 | |
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 500 | lemma inner_axis: "inner x (axis i y) = inner (x $ i) y" | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 501 | unfolding inner_vec_def | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 502 | apply (rule_tac k=i in setsum_single) | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 503 | apply simp_all | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 504 | apply (simp add: axis_def) | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 505 | done | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 506 | |
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 507 | text {* A bijection between @{text "'n::finite"} and @{text "{..<CARD('n)}"} *}
 | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 508 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 509 | definition vec_bij_nat :: "nat \<Rightarrow> ('n::finite)" where
 | 
| 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 510 |   "vec_bij_nat = (SOME p. bij_betw p {..<CARD('n)} (UNIV::'n set) )"
 | 
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 511 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 512 | abbreviation "\<pi> \<equiv> vec_bij_nat" | 
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 513 | definition "\<pi>' = inv_into {..<CARD('n)} (\<pi>::nat \<Rightarrow> ('n::finite))"
 | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 514 | |
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 515 | lemma bij_betw_pi: | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 516 |   "bij_betw \<pi> {..<CARD('n::finite)} (UNIV::('n::finite) set)"
 | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 517 | using ex_bij_betw_nat_finite[of "UNIV::'n set"] | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 518 | by (auto simp: vec_bij_nat_def atLeast0LessThan | 
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 519 |     intro!: someI_ex[of "\<lambda>x. bij_betw x {..<CARD('n)} (UNIV::'n set)"])
 | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 520 | |
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 521 | lemma bij_betw_pi'[intro]: "bij_betw \<pi>' (UNIV::'n set) {..<CARD('n::finite)}"
 | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 522 | using bij_betw_inv_into[OF bij_betw_pi] unfolding \<pi>'_def by auto | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 523 | |
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 524 | lemma pi'_inj[intro]: "inj \<pi>'" | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 525 | using bij_betw_pi' unfolding bij_betw_def by auto | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 526 | |
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 527 | lemma pi'_range[intro]: "\<And>i::'n. \<pi>' i < CARD('n::finite)"
 | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 528 | using bij_betw_pi' unfolding bij_betw_def by auto | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 529 | |
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 530 | lemma \<pi>\<pi>'[simp]: "\<And>i::'n::finite. \<pi> (\<pi>' i) = i" | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 531 | using bij_betw_pi by (auto intro!: f_inv_into_f simp: \<pi>'_def bij_betw_def) | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 532 | |
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 533 | lemma \<pi>'\<pi>[simp]: "\<And>i. i\<in>{..<CARD('n::finite)} \<Longrightarrow> \<pi>' (\<pi> i::'n) = i"
 | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 534 | using bij_betw_pi by (auto intro!: inv_into_f_eq simp: \<pi>'_def bij_betw_def) | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 535 | |
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 536 | lemma \<pi>\<pi>'_alt[simp]: "\<And>i. i<CARD('n::finite) \<Longrightarrow> \<pi>' (\<pi> i::'n) = i"
 | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 537 | by auto | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 538 | |
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 539 | lemma \<pi>_inj_on: "inj_on (\<pi>::nat\<Rightarrow>'n::finite) {..<CARD('n)}"
 | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 540 | using bij_betw_pi[where 'n='n] by (simp add: bij_betw_def) | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 541 | |
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 542 | instantiation vec :: (euclidean_space, finite) euclidean_space | 
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 543 | begin | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 544 | |
| 44166 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 545 | definition "Basis = (\<Union>i. \<Union>u\<in>Basis. {axis i u})"
 | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 546 | |
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 547 | definition "dimension (t :: ('a ^ 'b) itself) = CARD('b) * DIM('a)"
 | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 548 | |
| 44166 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 549 | definition "basis i = | 
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 550 |   (if i < (CARD('b) * DIM('a))
 | 
| 44166 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 551 |   then axis (\<pi>(i div DIM('a))) (basis (i mod DIM('a)))
 | 
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 552 | else 0)" | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 553 | |
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 554 | lemma basis_eq: | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 555 |   assumes "i < CARD('b)" and "j < DIM('a)"
 | 
| 44166 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 556 |   shows "basis (j + i * DIM('a)) = axis (\<pi> i) (basis j)"
 | 
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 557 | proof - | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 558 |   have "j + i * DIM('a) <  DIM('a) * (i + 1)" using assms by (auto simp: field_simps)
 | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 559 |   also have "\<dots> \<le> DIM('a) * CARD('b)" using assms unfolding mult_le_cancel1 by auto
 | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 560 | finally show ?thesis | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 561 | unfolding basis_vec_def using assms by (auto simp: vec_eq_iff not_less field_simps) | 
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 562 | qed | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 563 | |
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 564 | lemma basis_eq_pi': | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 565 |   assumes "j < DIM('a)"
 | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 566 |   shows "basis (j + \<pi>' i * DIM('a)) $ k = (if k = i then basis j else 0)"
 | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 567 | apply (subst basis_eq) | 
| 44166 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 568 | using pi'_range assms by (simp_all add: axis_def) | 
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 569 | |
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 570 | lemma split_times_into_modulo[consumes 1]: | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 571 | fixes k :: nat | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 572 | assumes "k < A * B" | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 573 | obtains i j where "i < A" and "j < B" and "k = j + i * B" | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 574 | proof | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 575 | have "A * B \<noteq> 0" | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 576 | proof assume "A * B = 0" with assms show False by simp qed | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 577 | hence "0 < B" by auto | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 578 | thus "k mod B < B" using `0 < B` by auto | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 579 | next | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 580 | have "k div B * B \<le> k div B * B + k mod B" by (rule le_add1) | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 581 | also have "... < A * B" using assms by simp | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 582 | finally show "k div B < A" by auto | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 583 | qed simp | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 584 | |
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 585 | lemma linear_less_than_times: | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 586 | fixes i j A B :: nat assumes "i < B" "j < A" | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 587 | shows "j + i * A < B * A" | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 588 | proof - | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 589 | have "i * A + j < (Suc i)*A" using `j < A` by simp | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 590 | also have "\<dots> \<le> B * A" using `i < B` unfolding mult_le_cancel2 by simp | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 591 | finally show ?thesis by simp | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 592 | qed | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 593 | |
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 594 | lemma DIM_cart[simp]: "DIM('a^'b) = CARD('b) * DIM('a)"
 | 
| 44136 
e63ad7d5158d
more uniform naming scheme for finite cartesian product type and related theorems
 huffman parents: 
44135diff
changeset | 595 | by (rule dimension_vec_def) | 
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 596 | |
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 597 | instance proof | 
| 44166 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 598 |   show "(Basis :: ('a ^ 'b) set) \<noteq> {}"
 | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 599 | unfolding Basis_vec_def by simp | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 600 | next | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 601 |   show "finite (Basis :: ('a ^ 'b) set)"
 | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 602 | unfolding Basis_vec_def by simp | 
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 603 | next | 
| 44166 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 604 | fix u v :: "'a ^ 'b" | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 605 | assume "u \<in> Basis" and "v \<in> Basis" | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 606 | thus "inner u v = (if u = v then 1 else 0)" | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 607 | unfolding Basis_vec_def | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 608 | by (auto simp add: inner_axis_axis axis_eq_axis inner_Basis) | 
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 609 | next | 
| 44166 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 610 | fix x :: "'a ^ 'b" | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 611 | show "(\<forall>u\<in>Basis. inner x u = 0) \<longleftrightarrow> x = 0" | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 612 | unfolding Basis_vec_def | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 613 | by (simp add: inner_axis euclidean_all_zero_iff vec_eq_iff) | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 614 | next | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 615 |   show "DIM('a ^ 'b) = card (Basis :: ('a ^ 'b) set)"
 | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 616 | unfolding Basis_vec_def dimension_vec_def dimension_def | 
| 44215 | 617 | by (simp add: card_UN_disjoint [unfolded disjoint_iff_not_equal] | 
| 44166 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 618 | axis_eq_axis nonzero_Basis) | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 619 | next | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 620 |   show "basis ` {..<DIM('a ^ 'b)} = (Basis :: ('a ^ 'b) set)"
 | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 621 | unfolding Basis_vec_def | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 622 | apply auto | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 623 | apply (erule split_times_into_modulo) | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 624 | apply (simp add: basis_eq axis_eq_axis) | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 625 | apply (erule Basis_elim) | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 626 | apply (simp add: image_def basis_vec_def axis_eq_axis) | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 627 | apply (rule rev_bexI, simp) | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 628 | apply (erule linear_less_than_times [OF pi'_range]) | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 629 | apply simp | 
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 630 | done | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 631 | next | 
| 44166 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 632 |   show "basis ` {DIM('a ^ 'b)..} = {0::'a ^ 'b}"
 | 
| 
d12d89a66742
modify euclidean_space class to include basis set
 huffman parents: 
44165diff
changeset | 633 | by (auto simp add: image_def basis_vec_def) | 
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 634 | qed | 
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 635 | |
| 36591 
df38e0c5c123
move class instantiations from Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
36590diff
changeset | 636 | end | 
| 44135 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 637 | |
| 
18b4ab6854f1
move euclidean_space instance from Cartesian_Euclidean_Space.thy to Finite_Cartesian_Product.thy
 huffman parents: 
42290diff
changeset | 638 | end |