src/HOL/Analysis/Euclidean_Space.thy
author wenzelm
Sun, 25 Aug 2024 22:54:56 +0200
changeset 80771 fd01ef524169
parent 80768 c7723cc15de8
child 80914 d97fdabd9e2b
permissions -rw-r--r--
use nicer notation, following 783406dd051e;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63627
6ddb43c6b711 rename HOL-Multivariate_Analysis to HOL-Analysis.
hoelzl
parents: 63589
diff changeset
     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
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
     4
*)
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
     5
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 58877
diff changeset
     6
section \<open>Finite-Dimensional Inner Product Spaces\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
     7
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
     8
theory Euclidean_Space
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
     9
imports
69511
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
    10
  L2_Norm
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
    11
  Inner_Product
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
    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
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69597
diff changeset
    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
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 58877
diff changeset
    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
75129a73aca3 more economic tagging
nipkow
parents: 68310
diff changeset
    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'(_')))")
80768
c7723cc15de8 more markup for syntax consts;
wenzelm
parents: 77490
diff changeset
    37
syntax_consts "_type_dimension" \<rightleftharpoons> card
63141
7e5084ad95aa recovered printing of DIM('a) (cf. 899c9c4e4a4c);
wenzelm
parents: 63114
diff changeset
    38
translations "DIM('a)" \<rightharpoonup> "CONST card (CONST Basis :: 'a set)"
7e5084ad95aa recovered printing of DIM('a) (cf. 899c9c4e4a4c);
wenzelm
parents: 63114
diff changeset
    39
typed_print_translation \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69516
diff changeset
    40
  [(\<^const_syntax>\<open>card\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69516
diff changeset
    41
    fn ctxt => fn _ => fn [Const (\<^const_syntax>\<open>Basis\<close>, Type (\<^type_name>\<open>set\<close>, [T]))] =>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69516
diff changeset
    42
      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
    43
\<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
    44
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
    45
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
    46
  unfolding norm_eq_sqrt_inner by (simp add: inner_Basis)
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
    47
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
    48
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
    49
  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
    50
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
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
    52
  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
    53
44166
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
    54
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
    55
  unfolding sgn_div_norm by (simp add: scaleR_one)
44166
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
    56
77490
2c86ea8961b5 Some new lemmas. Some tidying up
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
    57
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
    58
  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
    59
44166
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
    60
lemma (in euclidean_space) Basis_zero [simp]: "0 \<notin> Basis"
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
    61
proof
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
    62
  assume "0 \<in> Basis" thus "False"
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
    63
    using inner_Basis [of 0 0] by simp
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
    64
qed
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
    65
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
    66
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
    67
  by clarsimp
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
    68
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
    69
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
    70
  by (metis ex_in_conv nonempty_Basis someI_ex)
44166
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
    71
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
    72
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
    73
  by (simp add: SOME_Basis)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
    74
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63971
diff changeset
    75
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
    76
    "b \<in> Basis \<Longrightarrow> inner (\<Sum>i\<in>Basis. f i *\<^sub>R i) b = f b"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63971
diff changeset
    77
  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
    78
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
    79
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
    80
  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
    81
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
    82
  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
    83
    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
    84
  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
    85
    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
    86
qed
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36778
diff changeset
    87
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
    88
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
    89
  "x = y \<longleftrightarrow> (\<forall>b\<in>Basis. inner x b = inner y b)"
44129
286bd57858b9 simplified definition of class euclidean_space;
huffman
parents: 44128
diff changeset
    90
  by (auto intro: euclidean_eqI)
286bd57858b9 simplified definition of class euclidean_space;
huffman
parents: 44128
diff changeset
    91
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63971
diff changeset
    92
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
    93
  "(\<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
    94
  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
    95
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63971
diff changeset
    96
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
    97
  "b = (\<Sum>i\<in>Basis. f i *\<^sub>R i) \<longleftrightarrow> (\<forall>i\<in>Basis. f i = inner b i)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63971
diff changeset
    98
  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
    99
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
   100
lemma (in euclidean_space) euclidean_representation: "(\<Sum>b\<in>Basis. inner x b *\<^sub>R b) = x"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63971
diff changeset
   101
  unfolding euclidean_representation_sum by simp
44129
286bd57858b9 simplified definition of class euclidean_space;
huffman
parents: 44128
diff changeset
   102
67685
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67399
diff changeset
   103
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
   104
  by (subst (1 2) euclidean_representation [symmetric])
bdff8bf0a75b moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents: 67399
diff changeset
   105
    (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
   106
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
   107
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
   108
  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
   109
  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
   110
  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
   111
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
   112
  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
   113
  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
   114
    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
   115
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
   116
63952
354808e9f44b new material connected with HOL Light measure theory, plus more rationalisation
paulson <lp15@cam.ac.uk>
parents: 63938
diff changeset
   117
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
   118
  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
   119
  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
   120
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
   121
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63971
diff changeset
   122
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
   123
    "(\<lambda>x. \<Sum>b\<in>Basis. inner (f x) b *\<^sub>R b) = f"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63971
diff changeset
   124
  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
   125
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60420
diff changeset
   126
lemma euclidean_isCont:
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60420
diff changeset
   127
  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
   128
    shows "isCont f x"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63971
diff changeset
   129
  apply (subst euclidean_representation_sum_fun [symmetric])
b9a1486e79be setsum -> sum
nipkow
parents: 63971
diff changeset
   130
  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
   131
  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
   132
  done
6a6f15d8fbc4 New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents: 60420
diff changeset
   133
63938
f6ce08859d4c More mainly topological results
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
   134
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
   135
  by (simp add: card_gt_0_iff)
44628
bd17b7543af1 move lemmas from Topology_Euclidean_Space to Euclidean_Space
huffman
parents: 44571
diff changeset
   136
63938
f6ce08859d4c More mainly topological results
paulson <lp15@cam.ac.uk>
parents: 63627
diff changeset
   137
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
   138
  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
   139
63114
27afe7af7379 Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
   140
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63971
diff changeset
   141
lemma sum_inner_Basis_scaleR [simp]:
63114
27afe7af7379 Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
   142
  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
   143
  assumes "b \<in> Basis" shows "(\<Sum>i\<in>Basis. (inner i b) *\<^sub>R f i) = f b"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63971
diff changeset
   144
  by (simp add: comm_monoid_add_class.sum.remove [OF finite_Basis assms]
b9a1486e79be setsum -> sum
nipkow
parents: 63971
diff changeset
   145
         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
   146
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63971
diff changeset
   147
lemma sum_inner_Basis_eq [simp]:
63114
27afe7af7379 Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
   148
  assumes "b \<in> Basis" shows "(\<Sum>i\<in>Basis. (inner i b) * f i) = f b"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63971
diff changeset
   149
  by (simp add: comm_monoid_add_class.sum.remove [OF finite_Basis assms]
b9a1486e79be setsum -> sum
nipkow
parents: 63971
diff changeset
   150
         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
   151
66154
bc5e6461f759 Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
   152
lemma sum_if_inner [simp]:
bc5e6461f759 Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
   153
  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
   154
    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
   155
proof (cases "i=j")
bc5e6461f759 Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
   156
  case True
bc5e6461f759 Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
   157
  with assms show ?thesis
bc5e6461f759 Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
   158
    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
   159
next
bc5e6461f759 Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
   160
  case False
bc5e6461f759 Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
   161
  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
   162
        (\<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
   163
    apply (rule sum.cong)
bc5e6461f759 Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
   164
    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
   165
  also have "... = g j"
bc5e6461f759 Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
   166
    using assms by auto
bc5e6461f759 Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
   167
  finally show ?thesis
bc5e6461f759 Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
   168
    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
   169
qed
bc5e6461f759 Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
   170
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   171
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
   172
   "(\<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
   173
  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
   174
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   175
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
   176
  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
   177
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   178
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
   179
  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
   180
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   181
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
   182
  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
   183
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   184
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
   185
  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
   186
  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
   187
  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
   188
  done
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   189
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   190
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
   191
  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
   192
  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
   193
    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
   194
  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
   195
proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   196
  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
   197
    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
   198
  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
   199
    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
   200
  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
   201
  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
   202
    fix i :: 'n
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   203
    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
   204
    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
   205
      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
   206
      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
   207
        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
   208
    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
   209
      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
   210
      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
   211
    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
   212
  qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   213
  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
   214
  finally show ?thesis .
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   215
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   216
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   217
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69597
diff changeset
   218
subsection\<^marker>\<open>tag unimportant\<close> \<open>Subclass relationships\<close>
44571
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44531
diff changeset
   219
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44531
diff changeset
   220
instance euclidean_space \<subseteq> perfect_space
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44531
diff changeset
   221
proof
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44531
diff changeset
   222
  fix x :: 'a show "\<not> open {x}"
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44531
diff changeset
   223
  proof
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44531
diff changeset
   224
    assume "open {x}"
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44531
diff changeset
   225
    then obtain e where "0 < e" and e: "\<forall>y. dist y x < e \<longrightarrow> y = x"
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44531
diff changeset
   226
      unfolding open_dist by fast
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 63007
diff changeset
   227
    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
   228
    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
   229
      by (rule someI_ex) (auto simp: ex_in_conv)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 58877
diff changeset
   230
    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
   231
      unfolding y_def by (auto intro!: nonzero_Basis)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 58877
diff changeset
   232
    from \<open>0 < e\<close> have "dist y x < e"
53939
eb25bddf6a22 tuned proofs
huffman
parents: 50526
diff changeset
   233
      unfolding y_def by (simp add: dist_norm)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 58877
diff changeset
   234
    from \<open>y \<noteq> x\<close> and \<open>dist y x < e\<close> show "False"
44571
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44531
diff changeset
   235
      using e by simp
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44531
diff changeset
   236
  qed
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44531
diff changeset
   237
qed
bd91b77c4cd6 move class perfect_space into RealVector.thy;
huffman
parents: 44531
diff changeset
   238
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 58877
diff changeset
   239
subsection \<open>Class instances\<close>
33175
2083bde13ce1 distinguished session for multivariate analysis
himmelma
parents:
diff changeset
   240
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69597
diff changeset
   241
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
   242
68617
75129a73aca3 more economic tagging
nipkow
parents: 68310
diff changeset
   243
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
   244
begin
44129
286bd57858b9 simplified definition of class euclidean_space;
huffman
parents: 44128
diff changeset
   245
63627
6ddb43c6b711 rename HOL-Multivariate_Analysis to HOL-Analysis.
hoelzl
parents: 63589
diff changeset
   246
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
   247
  [simp]: "Basis = {1::real}"
44129
286bd57858b9 simplified definition of class euclidean_space;
huffman
parents: 44128
diff changeset
   248
286bd57858b9 simplified definition of class euclidean_space;
huffman
parents: 44128
diff changeset
   249
instance
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 60974
diff changeset
   250
  by standard auto
44129
286bd57858b9 simplified definition of class euclidean_space;
huffman
parents: 44128
diff changeset
   251
286bd57858b9 simplified definition of class euclidean_space;
huffman
parents: 44128
diff changeset
   252
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
   253
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
   254
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
   255
  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
   256
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69597
diff changeset
   257
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
   258
68617
75129a73aca3 more economic tagging
nipkow
parents: 68310
diff changeset
   259
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
   260
begin
44129
286bd57858b9 simplified definition of class euclidean_space;
huffman
parents: 44128
diff changeset
   261
63589
58aab4745e85 more symbols;
wenzelm
parents: 63141
diff changeset
   262
definition Basis_complex_def: "Basis = {1, \<i>}"
44166
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
   263
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
   264
instance
62390
842917225d56 more canonical names
nipkow
parents: 61169
diff changeset
   265
  by standard (auto simp add: Basis_complex_def intro: complex_eqI split: if_split_asm)
44129
286bd57858b9 simplified definition of class euclidean_space;
huffman
parents: 44128
diff changeset
   266
286bd57858b9 simplified definition of class euclidean_space;
huffman
parents: 44128
diff changeset
   267
end
286bd57858b9 simplified definition of class euclidean_space;
huffman
parents: 44128
diff changeset
   268
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36778
diff changeset
   269
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
   270
  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
   271
68310
d0a7ddf5450e more general tidying
paulson <lp15@cam.ac.uk>
parents: 68072
diff changeset
   272
lemma complex_Basis_1 [iff]: "(1::complex) \<in> Basis"
d0a7ddf5450e more general tidying
paulson <lp15@cam.ac.uk>
parents: 68072
diff changeset
   273
  by (simp add: Basis_complex_def)
d0a7ddf5450e more general tidying
paulson <lp15@cam.ac.uk>
parents: 68072
diff changeset
   274
d0a7ddf5450e more general tidying
paulson <lp15@cam.ac.uk>
parents: 68072
diff changeset
   275
lemma complex_Basis_i [iff]: "\<i> \<in> Basis"
d0a7ddf5450e more general tidying
paulson <lp15@cam.ac.uk>
parents: 68072
diff changeset
   276
  by (simp add: Basis_complex_def)
d0a7ddf5450e more general tidying
paulson <lp15@cam.ac.uk>
parents: 68072
diff changeset
   277
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69597
diff changeset
   278
subsubsection\<^marker>\<open>tag unimportant\<close> \<open>Type \<^typ>\<open>'a \<times> 'b\<close>\<close>
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   279
69511
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   280
instantiation prod :: (real_inner, real_inner) real_inner
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   281
begin
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   282
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   283
definition inner_prod_def:
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   284
  "inner x y = inner (fst x) (fst y) + inner (snd x) (snd y)"
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   285
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   286
lemma inner_Pair [simp]: "inner (a, b) (c, d) = inner a c + inner b d"
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   287
  unfolding inner_prod_def by simp
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   288
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   289
instance
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   290
proof
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   291
  fix r :: real
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   292
  fix x y z :: "'a::real_inner \<times> 'b::real_inner"
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   293
  show "inner x y = inner y x"
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   294
    unfolding inner_prod_def
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   295
    by (simp add: inner_commute)
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   296
  show "inner (x + y) z = inner x z + inner y z"
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   297
    unfolding inner_prod_def
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   298
    by (simp add: inner_add_left)
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   299
  show "inner (scaleR r x) y = r * inner x y"
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   300
    unfolding inner_prod_def
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   301
    by (simp add: distrib_left)
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   302
  show "0 \<le> inner x x"
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   303
    unfolding inner_prod_def
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   304
    by (intro add_nonneg_nonneg inner_ge_zero)
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   305
  show "inner x x = 0 \<longleftrightarrow> x = 0"
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   306
    unfolding inner_prod_def prod_eq_iff
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   307
    by (simp add: add_nonneg_eq_0_iff)
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   308
  show "norm x = sqrt (inner x x)"
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   309
    unfolding norm_prod_def inner_prod_def
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   310
    by (simp add: power2_norm_eq_inner)
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   311
qed
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   312
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   313
end
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   314
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   315
lemma inner_Pair_0: "inner x (0, b) = inner (snd x) b" "inner x (a, 0) = inner (fst x) a"
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   316
    by (cases x, simp)+
4cc5e4a528f8 moved dependency
immler
parents: 69064
diff changeset
   317
68617
75129a73aca3 more economic tagging
nipkow
parents: 68310
diff changeset
   318
instantiation prod :: (euclidean_space, euclidean_space) euclidean_space
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   319
begin
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   320
44129
286bd57858b9 simplified definition of class euclidean_space;
huffman
parents: 44128
diff changeset
   321
definition
44166
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
   322
  "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
   323
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63971
diff changeset
   324
lemma sum_Basis_prod_eq:
54781
fe462aa28c3d additional lemmas
immler
parents: 54776
diff changeset
   325
  fixes f::"('a*'b)\<Rightarrow>('a*'b)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63971
diff changeset
   326
  shows "sum f Basis = sum (\<lambda>i. f (i, 0)) Basis + sum (\<lambda>i. f (0, i)) Basis"
54781
fe462aa28c3d additional lemmas
immler
parents: 54776
diff changeset
   327
proof -
fe462aa28c3d additional lemmas
immler
parents: 54776
diff changeset
   328
  have "inj_on (\<lambda>u. (u::'a, 0::'b)) Basis" "inj_on (\<lambda>u. (0::'a, u::'b)) Basis"
fe462aa28c3d additional lemmas
immler
parents: 54776
diff changeset
   329
    by (auto intro!: inj_onI Pair_inject)
fe462aa28c3d additional lemmas
immler
parents: 54776
diff changeset
   330
  thus ?thesis
fe462aa28c3d additional lemmas
immler
parents: 54776
diff changeset
   331
    unfolding Basis_prod_def
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63971
diff changeset
   332
    by (subst sum.union_disjoint) (auto simp: Basis_prod_def sum.reindex)
54781
fe462aa28c3d additional lemmas
immler
parents: 54776
diff changeset
   333
qed
fe462aa28c3d additional lemmas
immler
parents: 54776
diff changeset
   334
44129
286bd57858b9 simplified definition of class euclidean_space;
huffman
parents: 44128
diff changeset
   335
instance proof
44166
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
   336
  show "(Basis :: ('a \<times> 'b) set) \<noteq> {}"
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
   337
    unfolding Basis_prod_def by simp
44129
286bd57858b9 simplified definition of class euclidean_space;
huffman
parents: 44128
diff changeset
   338
next
44166
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
   339
  show "finite (Basis :: ('a \<times> 'b) set)"
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
   340
    unfolding Basis_prod_def by simp
44129
286bd57858b9 simplified definition of class euclidean_space;
huffman
parents: 44128
diff changeset
   341
next
44166
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
   342
  fix u v :: "'a \<times> 'b"
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
   343
  assume "u \<in> Basis" and "v \<in> Basis"
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
   344
  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
   345
    unfolding Basis_prod_def inner_prod_def
62390
842917225d56 more canonical names
nipkow
parents: 61169
diff changeset
   346
    by (auto simp add: inner_Basis split: if_split_asm)
44129
286bd57858b9 simplified definition of class euclidean_space;
huffman
parents: 44128
diff changeset
   347
next
286bd57858b9 simplified definition of class euclidean_space;
huffman
parents: 44128
diff changeset
   348
  fix x :: "'a \<times> 'b"
44166
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
   349
  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
   350
    unfolding Basis_prod_def ball_Un ball_simps
d12d89a66742 modify euclidean_space class to include basis set
huffman
parents: 44133
diff changeset
   351
    by (simp add: inner_prod_def prod_eq_iff euclidean_all_zero_iff)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   352
qed
44129
286bd57858b9 simplified definition of class euclidean_space;
huffman
parents: 44128
diff changeset
   353
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
   354
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
   355
  unfolding Basis_prod_def
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66154
diff changeset
   356
  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
   357
37489
44e42d392c6e Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents: 36778
diff changeset
   358
end
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   359
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   360
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   361
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
   362
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   363
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
   364
  "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
   365
proof unfold_locales
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   366
  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
   367
  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
   368
    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
   369
    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
   370
    apply simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   371
    apply clarify
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   372
    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
   373
    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
   374
    done
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68620
diff changeset
   375
  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
   376
    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
   377
    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
   378
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   379
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   380
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
   381
  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
   382
    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
   383
    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
   384
    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
   385
    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
   386
    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
   387
    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
   388
    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
   389
    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
   390
    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
   391
  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
   392
      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
   393
      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
   394
      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
      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
   396
      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
   397
68620
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68617
diff changeset
   398
interpretation eucl?: finite_dimensional_vector_space_pair_1
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68617
diff changeset
   399
  "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
   400
  "scaleR::real\<Rightarrow>'b::real_vector \<Rightarrow> 'b"
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68617
diff changeset
   401
  by unfold_locales
707437105595 relaxed assumptions for dim_image_eq and dim_image_le
immler
parents: 68617
diff changeset
   402
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   403
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
   404
  rewrites "Basis_pair = Basis"
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68620
diff changeset
   405
    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
   406
proof -
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68620
diff changeset
   407
  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
   408
    by unfold_locales
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68620
diff changeset
   409
  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
   410
    by fact
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   411
  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
   412
    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
   413
  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
   414
    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
   415
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents: 67962
diff changeset
   416
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 38642
diff changeset
   417
end