| author | wenzelm |
| Thu, 01 Sep 2016 14:49:36 +0200 | |
| changeset 63745 | dde79b7faddf |
| parent 63627 | 6ddb43c6b711 |
| child 63938 | f6ce08859d4c |
| permissions | -rw-r--r-- |
| 63627 | 1 |
(* Title: HOL/Analysis/Euclidean_Space.thy |
|
44133
691c52e900ca
split Linear_Algebra.thy from Euclidean_Space.thy
huffman
parents:
44129
diff
changeset
|
2 |
Author: Johannes Hölzl, TU München |
|
691c52e900ca
split Linear_Algebra.thy from Euclidean_Space.thy
huffman
parents:
44129
diff
changeset
|
3 |
Author: Brian Huffman, Portland State University |
| 33175 | 4 |
*) |
5 |
||
| 60420 | 6 |
section \<open>Finite-Dimensional Inner Product Spaces\<close> |
| 33175 | 7 |
|
8 |
theory Euclidean_Space |
|
9 |
imports |
|
|
44628
bd17b7543af1
move lemmas from Topology_Euclidean_Space to Euclidean_Space
huffman
parents:
44571
diff
changeset
|
10 |
L2_Norm |
|
41413
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
wenzelm
parents:
40786
diff
changeset
|
11 |
"~~/src/HOL/Library/Inner_Product" |
|
44133
691c52e900ca
split Linear_Algebra.thy from Euclidean_Space.thy
huffman
parents:
44129
diff
changeset
|
12 |
"~~/src/HOL/Library/Product_Vector" |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36778
diff
changeset
|
13 |
begin |
|
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36778
diff
changeset
|
14 |
|
| 60420 | 15 |
subsection \<open>Type class of Euclidean spaces\<close> |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36778
diff
changeset
|
16 |
|
| 44129 | 17 |
class euclidean_space = real_inner + |
|
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
18 |
fixes Basis :: "'a set" |
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
19 |
assumes nonempty_Basis [simp]: "Basis \<noteq> {}"
|
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
20 |
assumes finite_Basis [simp]: "finite Basis" |
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
21 |
assumes inner_Basis: |
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
22 |
"\<lbrakk>u \<in> Basis; v \<in> Basis\<rbrakk> \<Longrightarrow> inner u v = (if u = v then 1 else 0)" |
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
23 |
assumes euclidean_all_zero_iff: |
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
24 |
"(\<forall>u\<in>Basis. inner x u = 0) \<longleftrightarrow> (x = 0)" |
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
25 |
|
|
63141
7e5084ad95aa
recovered printing of DIM('a) (cf. 899c9c4e4a4c);
wenzelm
parents:
63114
diff
changeset
|
26 |
syntax "_type_dimension" :: "type \<Rightarrow> nat" ("(1DIM/(1'(_')))")
|
|
7e5084ad95aa
recovered printing of DIM('a) (cf. 899c9c4e4a4c);
wenzelm
parents:
63114
diff
changeset
|
27 |
translations "DIM('a)" \<rightharpoonup> "CONST card (CONST Basis :: 'a set)"
|
|
7e5084ad95aa
recovered printing of DIM('a) (cf. 899c9c4e4a4c);
wenzelm
parents:
63114
diff
changeset
|
28 |
typed_print_translation \<open> |
|
7e5084ad95aa
recovered printing of DIM('a) (cf. 899c9c4e4a4c);
wenzelm
parents:
63114
diff
changeset
|
29 |
[(@{const_syntax card},
|
|
7e5084ad95aa
recovered printing of DIM('a) (cf. 899c9c4e4a4c);
wenzelm
parents:
63114
diff
changeset
|
30 |
fn ctxt => fn _ => fn [Const (@{const_syntax Basis}, Type (@{type_name set}, [T]))] =>
|
|
7e5084ad95aa
recovered printing of DIM('a) (cf. 899c9c4e4a4c);
wenzelm
parents:
63114
diff
changeset
|
31 |
Syntax.const @{syntax_const "_type_dimension"} $ Syntax_Phases.term_of_typ ctxt T)]
|
|
7e5084ad95aa
recovered printing of DIM('a) (cf. 899c9c4e4a4c);
wenzelm
parents:
63114
diff
changeset
|
32 |
\<close> |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36778
diff
changeset
|
33 |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
34 |
lemma (in euclidean_space) norm_Basis[simp]: "u \<in> Basis \<Longrightarrow> norm u = 1" |
|
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
35 |
unfolding norm_eq_sqrt_inner by (simp add: inner_Basis) |
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
36 |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
37 |
lemma (in euclidean_space) inner_same_Basis[simp]: "u \<in> Basis \<Longrightarrow> inner u u = 1" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
38 |
by (simp add: inner_Basis) |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
39 |
|
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
40 |
lemma (in euclidean_space) inner_not_same_Basis: "u \<in> Basis \<Longrightarrow> v \<in> Basis \<Longrightarrow> u \<noteq> v \<Longrightarrow> inner u v = 0" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
41 |
by (simp add: inner_Basis) |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
42 |
|
|
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
43 |
lemma (in euclidean_space) sgn_Basis: "u \<in> Basis \<Longrightarrow> sgn u = u" |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
44 |
unfolding sgn_div_norm by (simp add: scaleR_one) |
|
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
45 |
|
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
46 |
lemma (in euclidean_space) Basis_zero [simp]: "0 \<notin> Basis" |
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
47 |
proof |
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
48 |
assume "0 \<in> Basis" thus "False" |
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
49 |
using inner_Basis [of 0 0] by simp |
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
50 |
qed |
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
51 |
|
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
52 |
lemma (in euclidean_space) nonzero_Basis: "u \<in> Basis \<Longrightarrow> u \<noteq> 0" |
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
53 |
by clarsimp |
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
54 |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
55 |
lemma (in euclidean_space) SOME_Basis: "(SOME i. i \<in> Basis) \<in> Basis" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
56 |
by (metis ex_in_conv nonempty_Basis someI_ex) |
|
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
57 |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
58 |
lemma (in euclidean_space) inner_setsum_left_Basis[simp]: |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
59 |
"b \<in> Basis \<Longrightarrow> inner (\<Sum>i\<in>Basis. f i *\<^sub>R i) b = f b" |
| 57418 | 60 |
by (simp add: inner_setsum_left inner_Basis if_distrib comm_monoid_add_class.setsum.If_cases) |
|
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
61 |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
62 |
lemma (in euclidean_space) euclidean_eqI: |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
63 |
assumes b: "\<And>b. b \<in> Basis \<Longrightarrow> inner x b = inner y b" shows "x = y" |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36778
diff
changeset
|
64 |
proof - |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
65 |
from b have "\<forall>b\<in>Basis. inner (x - y) b = 0" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
66 |
by (simp add: inner_diff_left) |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
67 |
then show "x = y" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
68 |
by (simp add: euclidean_all_zero_iff) |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36778
diff
changeset
|
69 |
qed |
|
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36778
diff
changeset
|
70 |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
71 |
lemma (in euclidean_space) euclidean_eq_iff: |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
72 |
"x = y \<longleftrightarrow> (\<forall>b\<in>Basis. inner x b = inner y b)" |
| 44129 | 73 |
by (auto intro: euclidean_eqI) |
74 |
||
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
75 |
lemma (in euclidean_space) euclidean_representation_setsum: |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
76 |
"(\<Sum>i\<in>Basis. f i *\<^sub>R i) = b \<longleftrightarrow> (\<forall>i\<in>Basis. f i = inner b i)" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
77 |
by (subst euclidean_eq_iff) simp |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36778
diff
changeset
|
78 |
|
|
54776
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents:
53939
diff
changeset
|
79 |
lemma (in euclidean_space) euclidean_representation_setsum': |
|
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents:
53939
diff
changeset
|
80 |
"b = (\<Sum>i\<in>Basis. f i *\<^sub>R i) \<longleftrightarrow> (\<forall>i\<in>Basis. f i = inner b i)" |
|
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents:
53939
diff
changeset
|
81 |
by (auto simp add: euclidean_representation_setsum[symmetric]) |
|
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents:
53939
diff
changeset
|
82 |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
83 |
lemma (in euclidean_space) euclidean_representation: "(\<Sum>b\<in>Basis. inner x b *\<^sub>R b) = x" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
84 |
unfolding euclidean_representation_setsum by simp |
| 44129 | 85 |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
86 |
lemma (in euclidean_space) choice_Basis_iff: |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
87 |
fixes P :: "'a \<Rightarrow> real \<Rightarrow> bool" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
88 |
shows "(\<forall>i\<in>Basis. \<exists>x. P i x) \<longleftrightarrow> (\<exists>x. \<forall>i\<in>Basis. P i (inner x i))" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
89 |
unfolding bchoice_iff |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
90 |
proof safe |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
91 |
fix f assume "\<forall>i\<in>Basis. P i (f i)" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
92 |
then show "\<exists>x. \<forall>i\<in>Basis. P i (inner x i)" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
93 |
by (auto intro!: exI[of _ "\<Sum>i\<in>Basis. f i *\<^sub>R i"]) |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
94 |
qed auto |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36778
diff
changeset
|
95 |
|
| 63627 | 96 |
lemma (in euclidean_space) euclidean_representation_setsum_fun: |
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60420
diff
changeset
|
97 |
"(\<lambda>x. \<Sum>b\<in>Basis. inner (f x) b *\<^sub>R b) = f" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60420
diff
changeset
|
98 |
by (rule ext) (simp add: euclidean_representation_setsum) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60420
diff
changeset
|
99 |
|
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60420
diff
changeset
|
100 |
lemma euclidean_isCont: |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60420
diff
changeset
|
101 |
assumes "\<And>b. b \<in> Basis \<Longrightarrow> isCont (\<lambda>x. (inner (f x) b) *\<^sub>R b) x" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60420
diff
changeset
|
102 |
shows "isCont f x" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60420
diff
changeset
|
103 |
apply (subst euclidean_representation_setsum_fun [symmetric]) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60420
diff
changeset
|
104 |
apply (rule isCont_setsum) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60420
diff
changeset
|
105 |
apply (blast intro: assms) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60420
diff
changeset
|
106 |
done |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60420
diff
changeset
|
107 |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
108 |
lemma DIM_positive: "0 < DIM('a::euclidean_space)"
|
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
109 |
by (simp add: card_gt_0_iff) |
|
44628
bd17b7543af1
move lemmas from Topology_Euclidean_Space to Euclidean_Space
huffman
parents:
44571
diff
changeset
|
110 |
|
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62390
diff
changeset
|
111 |
lemma DIM_ge_Suc0 [iff]: "Suc 0 \<le> card Basis" |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62390
diff
changeset
|
112 |
by (meson DIM_positive Suc_leI) |
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62390
diff
changeset
|
113 |
|
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
114 |
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
115 |
lemma setsum_inner_Basis_scaleR [simp]: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
116 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_vector" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
117 |
assumes "b \<in> Basis" shows "(\<Sum>i\<in>Basis. (inner i b) *\<^sub>R f i) = f b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
118 |
by (simp add: comm_monoid_add_class.setsum.remove [OF finite_Basis assms] |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
119 |
assms inner_not_same_Basis comm_monoid_add_class.setsum.neutral) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
120 |
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
121 |
lemma setsum_inner_Basis_eq [simp]: |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
122 |
assumes "b \<in> Basis" shows "(\<Sum>i\<in>Basis. (inner i b) * f i) = f b" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
123 |
by (simp add: comm_monoid_add_class.setsum.remove [OF finite_Basis assms] |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
124 |
assms inner_not_same_Basis comm_monoid_add_class.setsum.neutral) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
125 |
|
| 60420 | 126 |
subsection \<open>Subclass relationships\<close> |
| 44571 | 127 |
|
128 |
instance euclidean_space \<subseteq> perfect_space |
|
129 |
proof |
|
130 |
fix x :: 'a show "\<not> open {x}"
|
|
131 |
proof |
|
132 |
assume "open {x}"
|
|
133 |
then obtain e where "0 < e" and e: "\<forall>y. dist y x < e \<longrightarrow> y = x" |
|
134 |
unfolding open_dist by fast |
|
| 63040 | 135 |
define y where "y = x + scaleR (e/2) (SOME b. b \<in> Basis)" |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
136 |
have [simp]: "(SOME b. b \<in> Basis) \<in> Basis" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
137 |
by (rule someI_ex) (auto simp: ex_in_conv) |
| 60420 | 138 |
from \<open>0 < e\<close> have "y \<noteq> x" |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
139 |
unfolding y_def by (auto intro!: nonzero_Basis) |
| 60420 | 140 |
from \<open>0 < e\<close> have "dist y x < e" |
| 53939 | 141 |
unfolding y_def by (simp add: dist_norm) |
| 60420 | 142 |
from \<open>y \<noteq> x\<close> and \<open>dist y x < e\<close> show "False" |
| 44571 | 143 |
using e by simp |
144 |
qed |
|
145 |
qed |
|
146 |
||
| 60420 | 147 |
subsection \<open>Class instances\<close> |
| 33175 | 148 |
|
| 60420 | 149 |
subsubsection \<open>Type @{typ real}\<close>
|
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36778
diff
changeset
|
150 |
|
| 44129 | 151 |
instantiation real :: euclidean_space |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36778
diff
changeset
|
152 |
begin |
| 44129 | 153 |
|
| 63627 | 154 |
definition |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
155 |
[simp]: "Basis = {1::real}"
|
| 44129 | 156 |
|
157 |
instance |
|
| 61169 | 158 |
by standard auto |
| 44129 | 159 |
|
160 |
end |
|
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36778
diff
changeset
|
161 |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
162 |
lemma DIM_real[simp]: "DIM(real) = 1" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
163 |
by simp |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
164 |
|
| 60420 | 165 |
subsubsection \<open>Type @{typ complex}\<close>
|
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36778
diff
changeset
|
166 |
|
| 44129 | 167 |
instantiation complex :: euclidean_space |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36778
diff
changeset
|
168 |
begin |
| 44129 | 169 |
|
| 63589 | 170 |
definition Basis_complex_def: "Basis = {1, \<i>}"
|
|
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
171 |
|
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
172 |
instance |
| 62390 | 173 |
by standard (auto simp add: Basis_complex_def intro: complex_eqI split: if_split_asm) |
| 44129 | 174 |
|
175 |
end |
|
176 |
||
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36778
diff
changeset
|
177 |
lemma DIM_complex[simp]: "DIM(complex) = 2" |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
178 |
unfolding Basis_complex_def by simp |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36778
diff
changeset
|
179 |
|
| 60420 | 180 |
subsubsection \<open>Type @{typ "'a \<times> 'b"}\<close>
|
| 38656 | 181 |
|
| 44129 | 182 |
instantiation prod :: (euclidean_space, euclidean_space) euclidean_space |
| 38656 | 183 |
begin |
184 |
||
| 44129 | 185 |
definition |
|
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
186 |
"Basis = (\<lambda>u. (u, 0)) ` Basis \<union> (\<lambda>v. (0, v)) ` Basis" |
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
187 |
|
| 54781 | 188 |
lemma setsum_Basis_prod_eq: |
189 |
fixes f::"('a*'b)\<Rightarrow>('a*'b)"
|
|
190 |
shows "setsum f Basis = setsum (\<lambda>i. f (i, 0)) Basis + setsum (\<lambda>i. f (0, i)) Basis" |
|
191 |
proof - |
|
192 |
have "inj_on (\<lambda>u. (u::'a, 0::'b)) Basis" "inj_on (\<lambda>u. (0::'a, u::'b)) Basis" |
|
193 |
by (auto intro!: inj_onI Pair_inject) |
|
194 |
thus ?thesis |
|
195 |
unfolding Basis_prod_def |
|
| 57418 | 196 |
by (subst setsum.union_disjoint) (auto simp: Basis_prod_def setsum.reindex) |
| 54781 | 197 |
qed |
198 |
||
| 44129 | 199 |
instance proof |
|
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
200 |
show "(Basis :: ('a \<times> 'b) set) \<noteq> {}"
|
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
201 |
unfolding Basis_prod_def by simp |
| 44129 | 202 |
next |
|
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
203 |
show "finite (Basis :: ('a \<times> 'b) set)"
|
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
204 |
unfolding Basis_prod_def by simp |
| 44129 | 205 |
next |
|
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
206 |
fix u v :: "'a \<times> 'b" |
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
207 |
assume "u \<in> Basis" and "v \<in> Basis" |
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
208 |
thus "inner u v = (if u = v then 1 else 0)" |
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
209 |
unfolding Basis_prod_def inner_prod_def |
| 62390 | 210 |
by (auto simp add: inner_Basis split: if_split_asm) |
| 44129 | 211 |
next |
212 |
fix x :: "'a \<times> 'b" |
|
|
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
213 |
show "(\<forall>u\<in>Basis. inner x u = 0) \<longleftrightarrow> x = 0" |
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
214 |
unfolding Basis_prod_def ball_Un ball_simps |
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
215 |
by (simp add: inner_prod_def prod_eq_iff euclidean_all_zero_iff) |
| 38656 | 216 |
qed |
| 44129 | 217 |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
218 |
lemma DIM_prod[simp]: "DIM('a \<times> 'b) = DIM('a) + DIM('b)"
|
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
219 |
unfolding Basis_prod_def |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
220 |
by (subst card_Un_disjoint) (auto intro!: card_image arg_cong2[where f="op +"] inj_onI) |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
221 |
|
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36778
diff
changeset
|
222 |
end |
| 38656 | 223 |
|
224 |
end |