author | wenzelm |
Sat, 04 Mar 2023 12:16:58 +0100 | |
changeset 77501 | 2d8815f98537 |
parent 77490 | 2c86ea8961b5 |
child 80768 | c7723cc15de8 |
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 |
|
69511 | 10 |
L2_Norm |
11 |
Inner_Product |
|
12 |
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 |
|
69516
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
immler
parents:
69511
diff
changeset
|
15 |
|
70136 | 16 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>Interlude: Some properties of real sets\<close> |
69516
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
immler
parents:
69511
diff
changeset
|
17 |
|
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
immler
parents:
69511
diff
changeset
|
18 |
lemma seq_mono_lemma: |
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
immler
parents:
69511
diff
changeset
|
19 |
assumes "\<forall>(n::nat) \<ge> m. (d n :: real) < e n" |
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
immler
parents:
69511
diff
changeset
|
20 |
and "\<forall>n \<ge> m. e n \<le> e m" |
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
immler
parents:
69511
diff
changeset
|
21 |
shows "\<forall>n \<ge> m. d n < e m" |
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
immler
parents:
69511
diff
changeset
|
22 |
using assms by force |
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
immler
parents:
69511
diff
changeset
|
23 |
|
09bb8f470959
most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
immler
parents:
69511
diff
changeset
|
24 |
|
60420 | 25 |
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
|
26 |
|
68617 | 27 |
class euclidean_space = real_inner + |
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
28 |
fixes Basis :: "'a set" |
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
29 |
assumes nonempty_Basis [simp]: "Basis \<noteq> {}" |
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
30 |
assumes finite_Basis [simp]: "finite Basis" |
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
31 |
assumes inner_Basis: |
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
32 |
"\<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
|
33 |
assumes euclidean_all_zero_iff: |
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
34 |
"(\<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
|
35 |
|
63141
7e5084ad95aa
recovered printing of DIM('a) (cf. 899c9c4e4a4c);
wenzelm
parents:
63114
diff
changeset
|
36 |
syntax "_type_dimension" :: "type \<Rightarrow> nat" ("(1DIM/(1'(_')))") |
7e5084ad95aa
recovered printing of DIM('a) (cf. 899c9c4e4a4c);
wenzelm
parents:
63114
diff
changeset
|
37 |
translations "DIM('a)" \<rightharpoonup> "CONST card (CONST Basis :: 'a set)" |
7e5084ad95aa
recovered printing of DIM('a) (cf. 899c9c4e4a4c);
wenzelm
parents:
63114
diff
changeset
|
38 |
typed_print_translation \<open> |
69597 | 39 |
[(\<^const_syntax>\<open>card\<close>, |
40 |
fn ctxt => fn _ => fn [Const (\<^const_syntax>\<open>Basis\<close>, Type (\<^type_name>\<open>set\<close>, [T]))] => |
|
41 |
Syntax.const \<^syntax_const>\<open>_type_dimension\<close> $ Syntax_Phases.term_of_typ ctxt T)] |
|
63141
7e5084ad95aa
recovered printing of DIM('a) (cf. 899c9c4e4a4c);
wenzelm
parents:
63114
diff
changeset
|
42 |
\<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
|
43 |
|
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 |
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
|
45 |
unfolding norm_eq_sqrt_inner by (simp add: inner_Basis) |
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
46 |
|
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
|
47 |
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
|
48 |
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
|
49 |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
44902
diff
changeset
|
50 |
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
|
51 |
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
|
52 |
|
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
53 |
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
|
54 |
unfolding sgn_div_norm by (simp add: scaleR_one) |
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
55 |
|
77490
2c86ea8961b5
Some new lemmas. Some tidying up
paulson <lp15@cam.ac.uk>
parents:
70136
diff
changeset
|
56 |
lemma inner_sum_Basis[simp]: "i \<in> Basis \<Longrightarrow> inner (\<Sum>Basis) i = 1" |
2c86ea8961b5
Some new lemmas. Some tidying up
paulson <lp15@cam.ac.uk>
parents:
70136
diff
changeset
|
57 |
by (simp add: inner_sum_left sum.If_cases inner_Basis) |
2c86ea8961b5
Some new lemmas. Some tidying up
paulson <lp15@cam.ac.uk>
parents:
70136
diff
changeset
|
58 |
|
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
59 |
lemma (in euclidean_space) Basis_zero [simp]: "0 \<notin> Basis" |
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
60 |
proof |
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
61 |
assume "0 \<in> Basis" thus "False" |
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
62 |
using inner_Basis [of 0 0] by simp |
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
63 |
qed |
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
64 |
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
65 |
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
|
66 |
by clarsimp |
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
67 |
|
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
|
68 |
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
|
69 |
by (metis ex_in_conv nonempty_Basis someI_ex) |
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
70 |
|
64773
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64267
diff
changeset
|
71 |
lemma norm_some_Basis [simp]: "norm (SOME i. i \<in> Basis) = 1" |
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64267
diff
changeset
|
72 |
by (simp add: SOME_Basis) |
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64267
diff
changeset
|
73 |
|
64267 | 74 |
lemma (in euclidean_space) inner_sum_left_Basis[simp]: |
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 |
"b \<in> Basis \<Longrightarrow> inner (\<Sum>i\<in>Basis. f i *\<^sub>R i) b = f b" |
64267 | 76 |
by (simp add: inner_sum_left inner_Basis if_distrib comm_monoid_add_class.sum.If_cases) |
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
77 |
|
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
|
78 |
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
|
79 |
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
|
80 |
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
|
81 |
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
|
82 |
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
|
83 |
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
|
84 |
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
|
85 |
qed |
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36778
diff
changeset
|
86 |
|
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
|
87 |
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
|
88 |
"x = y \<longleftrightarrow> (\<forall>b\<in>Basis. inner x b = inner y b)" |
44129 | 89 |
by (auto intro: euclidean_eqI) |
90 |
||
64267 | 91 |
lemma (in euclidean_space) euclidean_representation_sum: |
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
|
92 |
"(\<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
|
93 |
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
|
94 |
|
64267 | 95 |
lemma (in euclidean_space) euclidean_representation_sum': |
54776
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents:
53939
diff
changeset
|
96 |
"b = (\<Sum>i\<in>Basis. f i *\<^sub>R i) \<longleftrightarrow> (\<forall>i\<in>Basis. f i = inner b i)" |
64267 | 97 |
by (auto simp add: euclidean_representation_sum[symmetric]) |
54776
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents:
53939
diff
changeset
|
98 |
|
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
|
99 |
lemma (in euclidean_space) euclidean_representation: "(\<Sum>b\<in>Basis. inner x b *\<^sub>R b) = x" |
64267 | 100 |
unfolding euclidean_representation_sum by simp |
44129 | 101 |
|
67685
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67399
diff
changeset
|
102 |
lemma (in euclidean_space) euclidean_inner: "inner x y = (\<Sum>b\<in>Basis. (inner x b) * (inner y b))" |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67399
diff
changeset
|
103 |
by (subst (1 2) euclidean_representation [symmetric]) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67399
diff
changeset
|
104 |
(simp add: inner_sum_right inner_Basis ac_simps) |
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67399
diff
changeset
|
105 |
|
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
|
106 |
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
|
107 |
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
|
108 |
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
|
109 |
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
|
110 |
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
|
111 |
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
|
112 |
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
|
113 |
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
|
114 |
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
|
115 |
|
63952
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
116 |
lemma (in euclidean_space) bchoice_Basis_iff: |
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
117 |
fixes P :: "'a \<Rightarrow> real \<Rightarrow> bool" |
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
118 |
shows "(\<forall>i\<in>Basis. \<exists>x\<in>A. P i x) \<longleftrightarrow> (\<exists>x. \<forall>i\<in>Basis. inner x i \<in> A \<and> P i (inner x i))" |
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
119 |
by (simp add: choice_Basis_iff Bex_def) |
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
120 |
|
64267 | 121 |
lemma (in euclidean_space) euclidean_representation_sum_fun: |
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60420
diff
changeset
|
122 |
"(\<lambda>x. \<Sum>b\<in>Basis. inner (f x) b *\<^sub>R b) = f" |
64267 | 123 |
by (rule ext) (simp add: euclidean_representation_sum) |
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60420
diff
changeset
|
124 |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60420
diff
changeset
|
125 |
lemma euclidean_isCont: |
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60420
diff
changeset
|
126 |
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
|
127 |
shows "isCont f x" |
64267 | 128 |
apply (subst euclidean_representation_sum_fun [symmetric]) |
129 |
apply (rule isCont_sum) |
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60420
diff
changeset
|
130 |
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
|
131 |
done |
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60420
diff
changeset
|
132 |
|
63938 | 133 |
lemma DIM_positive [simp]: "0 < DIM('a::euclidean_space)" |
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
|
134 |
by (simp add: card_gt_0_iff) |
44628
bd17b7543af1
move lemmas from Topology_Euclidean_Space to Euclidean_Space
huffman
parents:
44571
diff
changeset
|
135 |
|
63938 | 136 |
lemma DIM_ge_Suc0 [simp]: "Suc 0 \<le> card Basis" |
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62390
diff
changeset
|
137 |
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
|
138 |
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
139 |
|
64267 | 140 |
lemma sum_inner_Basis_scaleR [simp]: |
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
141 |
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
|
142 |
assumes "b \<in> Basis" shows "(\<Sum>i\<in>Basis. (inner i b) *\<^sub>R f i) = f b" |
64267 | 143 |
by (simp add: comm_monoid_add_class.sum.remove [OF finite_Basis assms] |
144 |
assms inner_not_same_Basis comm_monoid_add_class.sum.neutral) |
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
145 |
|
64267 | 146 |
lemma sum_inner_Basis_eq [simp]: |
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
147 |
assumes "b \<in> Basis" shows "(\<Sum>i\<in>Basis. (inner i b) * f i) = f b" |
64267 | 148 |
by (simp add: comm_monoid_add_class.sum.remove [OF finite_Basis assms] |
149 |
assms inner_not_same_Basis comm_monoid_add_class.sum.neutral) |
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
150 |
|
66154
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
151 |
lemma sum_if_inner [simp]: |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
152 |
assumes "i \<in> Basis" "j \<in> Basis" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
153 |
shows "inner (\<Sum>k\<in>Basis. if k = i then f i *\<^sub>R i else g k *\<^sub>R k) j = (if j=i then f j else g j)" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
154 |
proof (cases "i=j") |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
155 |
case True |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
156 |
with assms show ?thesis |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
157 |
by (auto simp: inner_sum_left if_distrib [of "\<lambda>x. inner x j"] inner_Basis cong: if_cong) |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
158 |
next |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
159 |
case False |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
160 |
have "(\<Sum>k\<in>Basis. inner (if k = i then f i *\<^sub>R i else g k *\<^sub>R k) j) = |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
161 |
(\<Sum>k\<in>Basis. if k = j then g k else 0)" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
162 |
apply (rule sum.cong) |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
163 |
using False assms by (auto simp: inner_Basis) |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
164 |
also have "... = g j" |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
165 |
using assms by auto |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
166 |
finally show ?thesis |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
167 |
using False by (auto simp: inner_sum_left) |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
168 |
qed |
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
169 |
|
68072
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
170 |
lemma norm_le_componentwise: |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
171 |
"(\<And>b. b \<in> Basis \<Longrightarrow> abs(inner x b) \<le> abs(inner y b)) \<Longrightarrow> norm x \<le> norm y" |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
172 |
by (auto simp: norm_le euclidean_inner [of x x] euclidean_inner [of y y] abs_le_square_iff power2_eq_square intro!: sum_mono) |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
173 |
|
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
174 |
lemma Basis_le_norm: "b \<in> Basis \<Longrightarrow> \<bar>inner x b\<bar> \<le> norm x" |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
175 |
by (rule order_trans [OF Cauchy_Schwarz_ineq2]) simp |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
176 |
|
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
177 |
lemma norm_bound_Basis_le: "b \<in> Basis \<Longrightarrow> norm x \<le> e \<Longrightarrow> \<bar>inner x b\<bar> \<le> e" |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
178 |
by (metis Basis_le_norm order_trans) |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
179 |
|
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
180 |
lemma norm_bound_Basis_lt: "b \<in> Basis \<Longrightarrow> norm x < e \<Longrightarrow> \<bar>inner x b\<bar> < e" |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
181 |
by (metis Basis_le_norm le_less_trans) |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
182 |
|
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
183 |
lemma norm_le_l1: "norm x \<le> (\<Sum>b\<in>Basis. \<bar>inner x b\<bar>)" |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
184 |
apply (subst euclidean_representation[of x, symmetric]) |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
185 |
apply (rule order_trans[OF norm_sum]) |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
186 |
apply (auto intro!: sum_mono) |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
187 |
done |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
188 |
|
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
189 |
lemma sum_norm_allsubsets_bound: |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
190 |
fixes f :: "'a \<Rightarrow> 'n::euclidean_space" |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
191 |
assumes fP: "finite P" |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
192 |
and fPs: "\<And>Q. Q \<subseteq> P \<Longrightarrow> norm (sum f Q) \<le> e" |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
193 |
shows "(\<Sum>x\<in>P. norm (f x)) \<le> 2 * real DIM('n) * e" |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
194 |
proof - |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
195 |
have "(\<Sum>x\<in>P. norm (f x)) \<le> (\<Sum>x\<in>P. \<Sum>b\<in>Basis. \<bar>inner (f x) b\<bar>)" |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
196 |
by (rule sum_mono) (rule norm_le_l1) |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
197 |
also have "(\<Sum>x\<in>P. \<Sum>b\<in>Basis. \<bar>inner (f x) b\<bar>) = (\<Sum>b\<in>Basis. \<Sum>x\<in>P. \<bar>inner (f x) b\<bar>)" |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
198 |
by (rule sum.swap) |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
199 |
also have "\<dots> \<le> of_nat (card (Basis :: 'n set)) * (2 * e)" |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
200 |
proof (rule sum_bounded_above) |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
201 |
fix i :: 'n |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
202 |
assume i: "i \<in> Basis" |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
203 |
have "norm (\<Sum>x\<in>P. \<bar>inner (f x) i\<bar>) \<le> |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
204 |
norm (inner (\<Sum>x\<in>P \<inter> - {x. inner (f x) i < 0}. f x) i) + norm (inner (\<Sum>x\<in>P \<inter> {x. inner (f x) i < 0}. f x) i)" |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
205 |
by (simp add: abs_real_def sum.If_cases[OF fP] sum_negf norm_triangle_ineq4 inner_sum_left |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
206 |
del: real_norm_def) |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
207 |
also have "\<dots> \<le> e + e" |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
208 |
unfolding real_norm_def |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
209 |
by (intro add_mono norm_bound_Basis_le i fPs) auto |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
210 |
finally show "(\<Sum>x\<in>P. \<bar>inner (f x) i\<bar>) \<le> 2*e" by simp |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
211 |
qed |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
212 |
also have "\<dots> = 2 * real DIM('n) * e" by simp |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
213 |
finally show ?thesis . |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
214 |
qed |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
215 |
|
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
216 |
|
70136 | 217 |
subsection\<^marker>\<open>tag unimportant\<close> \<open>Subclass relationships\<close> |
44571 | 218 |
|
219 |
instance euclidean_space \<subseteq> perfect_space |
|
220 |
proof |
|
221 |
fix x :: 'a show "\<not> open {x}" |
|
222 |
proof |
|
223 |
assume "open {x}" |
|
224 |
then obtain e where "0 < e" and e: "\<forall>y. dist y x < e \<longrightarrow> y = x" |
|
225 |
unfolding open_dist by fast |
|
63040 | 226 |
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
|
227 |
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
|
228 |
by (rule someI_ex) (auto simp: ex_in_conv) |
60420 | 229 |
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
|
230 |
unfolding y_def by (auto intro!: nonzero_Basis) |
60420 | 231 |
from \<open>0 < e\<close> have "dist y x < e" |
53939 | 232 |
unfolding y_def by (simp add: dist_norm) |
60420 | 233 |
from \<open>y \<noteq> x\<close> and \<open>dist y x < e\<close> show "False" |
44571 | 234 |
using e by simp |
235 |
qed |
|
236 |
qed |
|
237 |
||
60420 | 238 |
subsection \<open>Class instances\<close> |
33175 | 239 |
|
70136 | 240 |
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Type \<^typ>\<open>real\<close>\<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
|
241 |
|
68617 | 242 |
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
|
243 |
begin |
44129 | 244 |
|
63627 | 245 |
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
|
246 |
[simp]: "Basis = {1::real}" |
44129 | 247 |
|
248 |
instance |
|
61169 | 249 |
by standard auto |
44129 | 250 |
|
251 |
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
|
252 |
|
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
|
253 |
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
|
254 |
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
|
255 |
|
70136 | 256 |
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Type \<^typ>\<open>complex\<close>\<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
|
257 |
|
68617 | 258 |
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
|
259 |
begin |
44129 | 260 |
|
63589 | 261 |
definition Basis_complex_def: "Basis = {1, \<i>}" |
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
262 |
|
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
263 |
instance |
62390 | 264 |
by standard (auto simp add: Basis_complex_def intro: complex_eqI split: if_split_asm) |
44129 | 265 |
|
266 |
end |
|
267 |
||
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36778
diff
changeset
|
268 |
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
|
269 |
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
|
270 |
|
68310 | 271 |
lemma complex_Basis_1 [iff]: "(1::complex) \<in> Basis" |
272 |
by (simp add: Basis_complex_def) |
|
273 |
||
274 |
lemma complex_Basis_i [iff]: "\<i> \<in> Basis" |
|
275 |
by (simp add: Basis_complex_def) |
|
276 |
||
70136 | 277 |
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Type \<^typ>\<open>'a \<times> 'b\<close>\<close> |
38656 | 278 |
|
69511 | 279 |
instantiation prod :: (real_inner, real_inner) real_inner |
280 |
begin |
|
281 |
||
282 |
definition inner_prod_def: |
|
283 |
"inner x y = inner (fst x) (fst y) + inner (snd x) (snd y)" |
|
284 |
||
285 |
lemma inner_Pair [simp]: "inner (a, b) (c, d) = inner a c + inner b d" |
|
286 |
unfolding inner_prod_def by simp |
|
287 |
||
288 |
instance |
|
289 |
proof |
|
290 |
fix r :: real |
|
291 |
fix x y z :: "'a::real_inner \<times> 'b::real_inner" |
|
292 |
show "inner x y = inner y x" |
|
293 |
unfolding inner_prod_def |
|
294 |
by (simp add: inner_commute) |
|
295 |
show "inner (x + y) z = inner x z + inner y z" |
|
296 |
unfolding inner_prod_def |
|
297 |
by (simp add: inner_add_left) |
|
298 |
show "inner (scaleR r x) y = r * inner x y" |
|
299 |
unfolding inner_prod_def |
|
300 |
by (simp add: distrib_left) |
|
301 |
show "0 \<le> inner x x" |
|
302 |
unfolding inner_prod_def |
|
303 |
by (intro add_nonneg_nonneg inner_ge_zero) |
|
304 |
show "inner x x = 0 \<longleftrightarrow> x = 0" |
|
305 |
unfolding inner_prod_def prod_eq_iff |
|
306 |
by (simp add: add_nonneg_eq_0_iff) |
|
307 |
show "norm x = sqrt (inner x x)" |
|
308 |
unfolding norm_prod_def inner_prod_def |
|
309 |
by (simp add: power2_norm_eq_inner) |
|
310 |
qed |
|
311 |
||
312 |
end |
|
313 |
||
314 |
lemma inner_Pair_0: "inner x (0, b) = inner (snd x) b" "inner x (a, 0) = inner (fst x) a" |
|
315 |
by (cases x, simp)+ |
|
316 |
||
68617 | 317 |
instantiation prod :: (euclidean_space, euclidean_space) euclidean_space |
38656 | 318 |
begin |
319 |
||
44129 | 320 |
definition |
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
321 |
"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
|
322 |
|
64267 | 323 |
lemma sum_Basis_prod_eq: |
54781 | 324 |
fixes f::"('a*'b)\<Rightarrow>('a*'b)" |
64267 | 325 |
shows "sum f Basis = sum (\<lambda>i. f (i, 0)) Basis + sum (\<lambda>i. f (0, i)) Basis" |
54781 | 326 |
proof - |
327 |
have "inj_on (\<lambda>u. (u::'a, 0::'b)) Basis" "inj_on (\<lambda>u. (0::'a, u::'b)) Basis" |
|
328 |
by (auto intro!: inj_onI Pair_inject) |
|
329 |
thus ?thesis |
|
330 |
unfolding Basis_prod_def |
|
64267 | 331 |
by (subst sum.union_disjoint) (auto simp: Basis_prod_def sum.reindex) |
54781 | 332 |
qed |
333 |
||
44129 | 334 |
instance proof |
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
335 |
show "(Basis :: ('a \<times> 'b) set) \<noteq> {}" |
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
336 |
unfolding Basis_prod_def by simp |
44129 | 337 |
next |
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
338 |
show "finite (Basis :: ('a \<times> 'b) set)" |
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
339 |
unfolding Basis_prod_def by simp |
44129 | 340 |
next |
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
341 |
fix u v :: "'a \<times> 'b" |
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
342 |
assume "u \<in> Basis" and "v \<in> Basis" |
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
343 |
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
|
344 |
unfolding Basis_prod_def inner_prod_def |
62390 | 345 |
by (auto simp add: inner_Basis split: if_split_asm) |
44129 | 346 |
next |
347 |
fix x :: "'a \<times> 'b" |
|
44166
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
348 |
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
|
349 |
unfolding Basis_prod_def ball_Un ball_simps |
d12d89a66742
modify euclidean_space class to include basis set
huffman
parents:
44133
diff
changeset
|
350 |
by (simp add: inner_prod_def prod_eq_iff euclidean_all_zero_iff) |
38656 | 351 |
qed |
44129 | 352 |
|
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
|
353 |
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
|
354 |
unfolding Basis_prod_def |
67399 | 355 |
by (subst card_Un_disjoint) (auto intro!: card_image arg_cong2[where f="(+)"] inj_onI) |
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
|
356 |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
36778
diff
changeset
|
357 |
end |
38656 | 358 |
|
68072
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
359 |
|
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
360 |
subsection \<open>Locale instances\<close> |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
361 |
|
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
362 |
lemma finite_dimensional_vector_space_euclidean: |
69064
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68620
diff
changeset
|
363 |
"finite_dimensional_vector_space (*\<^sub>R) Basis" |
68072
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
364 |
proof unfold_locales |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
365 |
show "finite (Basis::'a set)" by (metis finite_Basis) |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
366 |
show "real_vector.independent (Basis::'a set)" |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
367 |
unfolding dependent_def dependent_raw_def[symmetric] |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
368 |
apply (subst span_finite) |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
369 |
apply simp |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
370 |
apply clarify |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
371 |
apply (drule_tac f="inner a" in arg_cong) |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
372 |
apply (simp add: inner_Basis inner_sum_right eq_commute) |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
373 |
done |
69064
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68620
diff
changeset
|
374 |
show "module.span (*\<^sub>R) Basis = UNIV" |
68072
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
375 |
unfolding span_finite [OF finite_Basis] span_raw_def[symmetric] |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
376 |
by (auto intro!: euclidean_representation[symmetric]) |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
377 |
qed |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
378 |
|
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
379 |
interpretation eucl?: finite_dimensional_vector_space "scaleR :: real => 'a => 'a::euclidean_space" "Basis" |
69064
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68620
diff
changeset
|
380 |
rewrites "module.dependent (*\<^sub>R) = dependent" |
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68620
diff
changeset
|
381 |
and "module.representation (*\<^sub>R) = representation" |
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68620
diff
changeset
|
382 |
and "module.subspace (*\<^sub>R) = subspace" |
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68620
diff
changeset
|
383 |
and "module.span (*\<^sub>R) = span" |
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68620
diff
changeset
|
384 |
and "vector_space.extend_basis (*\<^sub>R) = extend_basis" |
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68620
diff
changeset
|
385 |
and "vector_space.dim (*\<^sub>R) = dim" |
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68620
diff
changeset
|
386 |
and "Vector_Spaces.linear (*\<^sub>R) (*\<^sub>R) = linear" |
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68620
diff
changeset
|
387 |
and "Vector_Spaces.linear (*) (*\<^sub>R) = linear" |
68072
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
388 |
and "finite_dimensional_vector_space.dimension Basis = DIM('a)" |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
389 |
and "dimension = DIM('a)" |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
390 |
by (auto simp add: dependent_raw_def representation_raw_def |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
391 |
subspace_raw_def span_raw_def extend_basis_raw_def dim_raw_def linear_def |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
392 |
real_scaleR_def[abs_def] |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
393 |
finite_dimensional_vector_space.dimension_def |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
394 |
intro!: finite_dimensional_vector_space.dimension_def |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
395 |
finite_dimensional_vector_space_euclidean) |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
396 |
|
68620
707437105595
relaxed assumptions for dim_image_eq and dim_image_le
immler
parents:
68617
diff
changeset
|
397 |
interpretation eucl?: finite_dimensional_vector_space_pair_1 |
707437105595
relaxed assumptions for dim_image_eq and dim_image_le
immler
parents:
68617
diff
changeset
|
398 |
"scaleR::real\<Rightarrow>'a::euclidean_space\<Rightarrow>'a" Basis |
707437105595
relaxed assumptions for dim_image_eq and dim_image_le
immler
parents:
68617
diff
changeset
|
399 |
"scaleR::real\<Rightarrow>'b::real_vector \<Rightarrow> 'b" |
707437105595
relaxed assumptions for dim_image_eq and dim_image_le
immler
parents:
68617
diff
changeset
|
400 |
by unfold_locales |
707437105595
relaxed assumptions for dim_image_eq and dim_image_le
immler
parents:
68617
diff
changeset
|
401 |
|
68072
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
402 |
interpretation eucl?: finite_dimensional_vector_space_prod scaleR scaleR Basis Basis |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
403 |
rewrites "Basis_pair = Basis" |
69064
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68620
diff
changeset
|
404 |
and "module_prod.scale (*\<^sub>R) (*\<^sub>R) = (scaleR::_\<Rightarrow>_\<Rightarrow>('a \<times> 'b))" |
68072
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
405 |
proof - |
69064
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68620
diff
changeset
|
406 |
show "finite_dimensional_vector_space_prod (*\<^sub>R) (*\<^sub>R) Basis Basis" |
68072
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
407 |
by unfold_locales |
69064
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68620
diff
changeset
|
408 |
interpret finite_dimensional_vector_space_prod "(*\<^sub>R)" "(*\<^sub>R)" "Basis::'a set" "Basis::'b set" |
68072
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
409 |
by fact |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
410 |
show "Basis_pair = Basis" |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
411 |
unfolding Basis_pair_def Basis_prod_def by auto |
69064
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents:
68620
diff
changeset
|
412 |
show "module_prod.scale (*\<^sub>R) (*\<^sub>R) = scaleR" |
68072
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
413 |
by (fact module_prod_scale_eq_scaleR) |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
414 |
qed |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67962
diff
changeset
|
415 |
|
38656 | 416 |
end |