src/HOL/Hahn_Banach/Subspace.thy
author nipkow
Sat, 24 Feb 2024 11:29:30 +0100
changeset 79714 80cb54976c1c
parent 66453 cc19f7ca2ed6
child 80914 d97fdabd9e2b
permissions -rw-r--r--
timing function generation bug fix by Jonas Stahl
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31795
be3e1cc5005c standard naming conventions for session and theories;
wenzelm
parents: 30729
diff changeset
     1
(*  Title:      HOL/Hahn_Banach/Subspace.thy
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     2
    Author:     Gertrud Bauer, TU Munich
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     3
*)
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
     4
58889
5b7a9633cfa8 modernized header uniformly as section;
wenzelm
parents: 58745
diff changeset
     5
section \<open>Subspaces\<close>
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
     6
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
     7
theory Subspace
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 63910
diff changeset
     8
imports Vector_Space "HOL-Library.Set_Algebras"
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
     9
begin
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    10
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
    11
subsection \<open>Definition\<close>
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    12
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
    13
text \<open>
61540
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    14
  A non-empty subset \<open>U\<close> of a vector space \<open>V\<close> is a \<^emph>\<open>subspace\<close> of \<open>V\<close>, iff
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    15
  \<open>U\<close> is closed under addition and scalar multiplication.
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
    16
\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    17
29234
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
    18
locale subspace =
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 60412
diff changeset
    19
  fixes U :: "'a::{minus, plus, zero, uminus} set" and V
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    20
  assumes non_empty [iff, intro]: "U \<noteq> {}"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    21
    and subset [iff]: "U \<subseteq> V"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    22
    and add_closed [iff]: "x \<in> U \<Longrightarrow> y \<in> U \<Longrightarrow> x + y \<in> U"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    23
    and mult_closed [iff]: "x \<in> U \<Longrightarrow> a \<cdot> x \<in> U"
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    24
21210
c17fd2df4e9e renamed 'const_syntax' to 'notation';
wenzelm
parents: 19736
diff changeset
    25
notation (symbols)
19736
wenzelm
parents: 18689
diff changeset
    26
  subspace  (infix "\<unlhd>" 50)
14254
342634f38451 Isar/Locales: <loc>.intro and <loc>.axioms no longer intro? and elim? by
ballarin
parents: 13547
diff changeset
    27
19736
wenzelm
parents: 18689
diff changeset
    28
declare vectorspace.intro [intro?] subspace.intro [intro?]
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    29
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    30
lemma subspace_subset [elim]: "U \<unlhd> V \<Longrightarrow> U \<subseteq> V"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    31
  by (rule subspace.subset)
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
    32
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    33
lemma (in subspace) subsetD [iff]: "x \<in> U \<Longrightarrow> x \<in> V"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    34
  using subset by blast
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
    35
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    36
lemma subspaceD [elim]: "U \<unlhd> V \<Longrightarrow> x \<in> U \<Longrightarrow> x \<in> V"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    37
  by (rule subspace.subsetD)
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    38
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    39
lemma rev_subspaceD [elim?]: "x \<in> U \<Longrightarrow> U \<unlhd> V \<Longrightarrow> x \<in> V"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    40
  by (rule subspace.subsetD)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    41
13547
wenzelm
parents: 13515
diff changeset
    42
lemma (in subspace) diff_closed [iff]:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
    43
  assumes "vectorspace V"
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
    44
  assumes x: "x \<in> U" and y: "y \<in> U"
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
    45
  shows "x - y \<in> U"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
    46
proof -
29234
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
    47
  interpret vectorspace V by fact
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
    48
  from x y show ?thesis by (simp add: diff_eq1 negate_eq1)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
    49
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    50
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
    51
text \<open>
61486
3590367b0ce9 tuned document;
wenzelm
parents: 61076
diff changeset
    52
  \<^medskip>
61540
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    53
  Similar as for linear spaces, the existence of the zero element in every
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    54
  subspace follows from the non-emptiness of the carrier set and by vector
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    55
  space laws.
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
    56
\<close>
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    57
13547
wenzelm
parents: 13515
diff changeset
    58
lemma (in subspace) zero [intro]:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
    59
  assumes "vectorspace V"
13547
wenzelm
parents: 13515
diff changeset
    60
  shows "0 \<in> U"
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
    61
proof -
30729
461ee3e49ad3 interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents: 29252
diff changeset
    62
  interpret V: vectorspace V by fact
29234
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
    63
  have "U \<noteq> {}" by (rule non_empty)
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    64
  then obtain x where x: "x \<in> U" by blast
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
    65
  then have "x \<in> V" .. then have "0 = x - x" by simp
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
    66
  also from \<open>vectorspace V\<close> x x have "\<dots> \<in> U" by (rule diff_closed)
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    67
  finally show ?thesis .
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    68
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    69
13547
wenzelm
parents: 13515
diff changeset
    70
lemma (in subspace) neg_closed [iff]:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
    71
  assumes "vectorspace V"
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
    72
  assumes x: "x \<in> U"
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
    73
  shows "- x \<in> U"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
    74
proof -
29234
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
    75
  interpret vectorspace V by fact
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
    76
  from x show ?thesis by (simp add: negate_eq1)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
    77
qed
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    78
61486
3590367b0ce9 tuned document;
wenzelm
parents: 61076
diff changeset
    79
text \<open>\<^medskip> Further derived laws: every subspace is a vector space.\<close>
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    80
13547
wenzelm
parents: 13515
diff changeset
    81
lemma (in subspace) vectorspace [iff]:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
    82
  assumes "vectorspace V"
13547
wenzelm
parents: 13515
diff changeset
    83
  shows "vectorspace U"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
    84
proof -
29234
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
    85
  interpret vectorspace V by fact
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
    86
  show ?thesis
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
    87
  proof
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
    88
    show "U \<noteq> {}" ..
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
    89
    fix x y z assume x: "x \<in> U" and y: "y \<in> U" and z: "z \<in> U"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
    90
    fix a b :: real
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
    91
    from x y show "x + y \<in> U" by simp
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
    92
    from x show "a \<cdot> x \<in> U" by simp
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
    93
    from x y z show "(x + y) + z = x + (y + z)" by (simp add: add_ac)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
    94
    from x y show "x + y = y + x" by (simp add: add_ac)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
    95
    from x show "x - x = 0" by simp
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
    96
    from x show "0 + x = x" by simp
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
    97
    from x y show "a \<cdot> (x + y) = a \<cdot> x + a \<cdot> y" by (simp add: distrib)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
    98
    from x show "(a + b) \<cdot> x = a \<cdot> x + b \<cdot> x" by (simp add: distrib)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
    99
    from x show "(a * b) \<cdot> x = a \<cdot> b \<cdot> x" by (simp add: mult_assoc)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   100
    from x show "1 \<cdot> x = x" by simp
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   101
    from x show "- x = - 1 \<cdot> x" by (simp add: negate_eq1)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   102
    from x y show "x - y = x + - y" by (simp add: diff_eq1)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   103
  qed
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   104
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   105
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   106
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   107
text \<open>The subspace relation is reflexive.\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   108
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   109
lemma (in vectorspace) subspace_refl [intro]: "V \<unlhd> V"
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
   110
proof
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   111
  show "V \<noteq> {}" ..
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
   112
  show "V \<subseteq> V" ..
44887
7ca82df6e951 misc tuning and clarification;
wenzelm
parents: 44190
diff changeset
   113
next
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   114
  fix x y assume x: "x \<in> V" and y: "y \<in> V"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   115
  fix a :: real
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   116
  from x y show "x + y \<in> V" by simp
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   117
  from x show "a \<cdot> x \<in> V" by simp
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   118
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   119
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   120
text \<open>The subspace relation is transitive.\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   121
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   122
lemma (in vectorspace) subspace_trans [trans]:
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   123
  "U \<unlhd> V \<Longrightarrow> V \<unlhd> W \<Longrightarrow> U \<unlhd> W"
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
   124
proof
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   125
  assume uv: "U \<unlhd> V" and vw: "V \<unlhd> W"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   126
  from uv show "U \<noteq> {}" by (rule subspace.non_empty)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   127
  show "U \<subseteq> W"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   128
  proof -
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   129
    from uv have "U \<subseteq> V" by (rule subspace.subset)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   130
    also from vw have "V \<subseteq> W" by (rule subspace.subset)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   131
    finally show ?thesis .
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   132
  qed
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   133
  fix x y assume x: "x \<in> U" and y: "y \<in> U"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   134
  from uv and x y show "x + y \<in> U" by (rule subspace.add_closed)
60412
285c7ff27728 tuned proofs;
wenzelm
parents: 58889
diff changeset
   135
  from uv and x show "a \<cdot> x \<in> U" for a by (rule subspace.mult_closed)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   136
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   137
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   138
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   139
subsection \<open>Linear closure\<close>
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   140
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   141
text \<open>
61879
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61540
diff changeset
   142
  The \<^emph>\<open>linear closure\<close> of a vector \<open>x\<close> is the set of all scalar multiples of
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61540
diff changeset
   143
  \<open>x\<close>.
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   144
\<close>
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   145
58745
wenzelm
parents: 58744
diff changeset
   146
definition lin :: "('a::{minus,plus,zero}) \<Rightarrow> 'a set"
44887
7ca82df6e951 misc tuning and clarification;
wenzelm
parents: 44190
diff changeset
   147
  where "lin x = {a \<cdot> x | a. True}"
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   148
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   149
lemma linI [intro]: "y = a \<cdot> x \<Longrightarrow> y \<in> lin x"
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   150
  unfolding lin_def by blast
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   151
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   152
lemma linI' [iff]: "a \<cdot> x \<in> lin x"
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   153
  unfolding lin_def by blast
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   154
60412
285c7ff27728 tuned proofs;
wenzelm
parents: 58889
diff changeset
   155
lemma linE [elim]:
285c7ff27728 tuned proofs;
wenzelm
parents: 58889
diff changeset
   156
  assumes "x \<in> lin v"
285c7ff27728 tuned proofs;
wenzelm
parents: 58889
diff changeset
   157
  obtains a :: real where "x = a \<cdot> v"
285c7ff27728 tuned proofs;
wenzelm
parents: 58889
diff changeset
   158
  using assms unfolding lin_def by blast
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   159
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7567
diff changeset
   160
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   161
text \<open>Every vector is contained in its linear closure.\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   162
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   163
lemma (in vectorspace) x_lin_x [iff]: "x \<in> V \<Longrightarrow> x \<in> lin x"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   164
proof -
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   165
  assume "x \<in> V"
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   166
  then have "x = 1 \<cdot> x" by simp
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   167
  also have "\<dots> \<in> lin x" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   168
  finally show ?thesis .
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   169
qed
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   170
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   171
lemma (in vectorspace) "0_lin_x" [iff]: "x \<in> V \<Longrightarrow> 0 \<in> lin x"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   172
proof
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   173
  assume "x \<in> V"
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   174
  then show "0 = 0 \<cdot> x" by simp
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   175
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   176
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   177
text \<open>Any linear closure is a subspace.\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   178
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   179
lemma (in vectorspace) lin_subspace [intro]:
44887
7ca82df6e951 misc tuning and clarification;
wenzelm
parents: 44190
diff changeset
   180
  assumes x: "x \<in> V"
7ca82df6e951 misc tuning and clarification;
wenzelm
parents: 44190
diff changeset
   181
  shows "lin x \<unlhd> V"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   182
proof
44887
7ca82df6e951 misc tuning and clarification;
wenzelm
parents: 44190
diff changeset
   183
  from x show "lin x \<noteq> {}" by auto
7ca82df6e951 misc tuning and clarification;
wenzelm
parents: 44190
diff changeset
   184
next
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
   185
  show "lin x \<subseteq> V"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   186
  proof
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   187
    fix x' assume "x' \<in> lin x"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   188
    then obtain a where "x' = a \<cdot> x" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   189
    with x show "x' \<in> V" by simp
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   190
  qed
44887
7ca82df6e951 misc tuning and clarification;
wenzelm
parents: 44190
diff changeset
   191
next
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   192
  fix x' x'' assume x': "x' \<in> lin x" and x'': "x'' \<in> lin x"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   193
  show "x' + x'' \<in> lin x"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   194
  proof -
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   195
    from x' obtain a' where "x' = a' \<cdot> x" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   196
    moreover from x'' obtain a'' where "x'' = a'' \<cdot> x" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   197
    ultimately have "x' + x'' = (a' + a'') \<cdot> x"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   198
      using x by (simp add: distrib)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   199
    also have "\<dots> \<in> lin x" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   200
    finally show ?thesis .
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   201
  qed
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   202
  fix a :: real
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   203
  show "a \<cdot> x' \<in> lin x"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   204
  proof -
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   205
    from x' obtain a' where "x' = a' \<cdot> x" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   206
    with x have "a \<cdot> x' = (a * a') \<cdot> x" by (simp add: mult_assoc)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   207
    also have "\<dots> \<in> lin x" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   208
    finally show ?thesis .
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
   209
  qed
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   210
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   211
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   212
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   213
text \<open>Any linear closure is a vector space.\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   214
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   215
lemma (in vectorspace) lin_vectorspace [intro]:
23378
1d138d6bb461 tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   216
  assumes "x \<in> V"
1d138d6bb461 tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   217
  shows "vectorspace (lin x)"
1d138d6bb461 tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   218
proof -
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   219
  from \<open>x \<in> V\<close> have "subspace (lin x) V"
23378
1d138d6bb461 tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   220
    by (rule lin_subspace)
26199
04817a8802f2 explicit referencing of background facts;
wenzelm
parents: 25762
diff changeset
   221
  from this and vectorspace_axioms show ?thesis
23378
1d138d6bb461 tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   222
    by (rule subspace.vectorspace)
1d138d6bb461 tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
   223
qed
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   224
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   225
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   226
subsection \<open>Sum of two vectorspaces\<close>
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   227
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   228
text \<open>
61540
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
   229
  The \<^emph>\<open>sum\<close> of two vectorspaces \<open>U\<close> and \<open>V\<close> is the set of all sums of
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
   230
  elements from \<open>U\<close> and \<open>V\<close>.
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   231
\<close>
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   232
47445
69e96e5500df Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents: 44887
diff changeset
   233
lemma sum_def: "U + V = {u + v | u v. u \<in> U \<and> v \<in> V}"
44190
fe5504984937 HOL-Hahn_Banach: use Set_Algebras library
huffman
parents: 32960
diff changeset
   234
  unfolding set_plus_def by auto
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   235
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   236
lemma sumE [elim]:
47445
69e96e5500df Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents: 44887
diff changeset
   237
    "x \<in> U + V \<Longrightarrow> (\<And>u v. x = u + v \<Longrightarrow> u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> C) \<Longrightarrow> C"
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   238
  unfolding sum_def by blast
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   239
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   240
lemma sumI [intro]:
47445
69e96e5500df Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents: 44887
diff changeset
   241
    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> x = u + v \<Longrightarrow> x \<in> U + V"
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   242
  unfolding sum_def by blast
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
   243
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   244
lemma sumI' [intro]:
47445
69e96e5500df Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents: 44887
diff changeset
   245
    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> u + v \<in> U + V"
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   246
  unfolding sum_def by blast
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   247
61539
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
   248
text \<open>\<open>U\<close> is a subspace of \<open>U + V\<close>.\<close>
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   249
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   250
lemma subspace_sum1 [iff]:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   251
  assumes "vectorspace U" "vectorspace V"
47445
69e96e5500df Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents: 44887
diff changeset
   252
  shows "U \<unlhd> U + V"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   253
proof -
29234
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
   254
  interpret vectorspace U by fact
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
   255
  interpret vectorspace V by fact
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   256
  show ?thesis
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   257
  proof
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   258
    show "U \<noteq> {}" ..
47445
69e96e5500df Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents: 44887
diff changeset
   259
    show "U \<subseteq> U + V"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   260
    proof
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   261
      fix x assume x: "x \<in> U"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   262
      moreover have "0 \<in> V" ..
47445
69e96e5500df Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents: 44887
diff changeset
   263
      ultimately have "x + 0 \<in> U + V" ..
69e96e5500df Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents: 44887
diff changeset
   264
      with x show "x \<in> U + V" by simp
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   265
    qed
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   266
    fix x y assume x: "x \<in> U" and "y \<in> U"
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   267
    then show "x + y \<in> U" by simp
60412
285c7ff27728 tuned proofs;
wenzelm
parents: 58889
diff changeset
   268
    from x show "a \<cdot> x \<in> U" for a by simp
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   269
  qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   270
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   271
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   272
text \<open>The sum of two subspaces is again a subspace.\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   273
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   274
lemma sum_subspace [intro?]:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   275
  assumes "subspace U E" "vectorspace E" "subspace V E"
47445
69e96e5500df Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents: 44887
diff changeset
   276
  shows "U + V \<unlhd> E"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   277
proof -
29234
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
   278
  interpret subspace U E by fact
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
   279
  interpret vectorspace E by fact
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
   280
  interpret subspace V E by fact
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   281
  show ?thesis
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   282
  proof
47445
69e96e5500df Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents: 44887
diff changeset
   283
    have "0 \<in> U + V"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   284
    proof
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   285
      show "0 \<in> U" using \<open>vectorspace E\<close> ..
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   286
      show "0 \<in> V" using \<open>vectorspace E\<close> ..
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   287
      show "(0::'a) = 0 + 0" by simp
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   288
    qed
47445
69e96e5500df Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents: 44887
diff changeset
   289
    then show "U + V \<noteq> {}" by blast
69e96e5500df Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents: 44887
diff changeset
   290
    show "U + V \<subseteq> E"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   291
    proof
47445
69e96e5500df Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents: 44887
diff changeset
   292
      fix x assume "x \<in> U + V"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   293
      then obtain u v where "x = u + v" and
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   294
        "u \<in> U" and "v \<in> V" ..
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   295
      then show "x \<in> E" by simp
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   296
    qed
44887
7ca82df6e951 misc tuning and clarification;
wenzelm
parents: 44190
diff changeset
   297
  next
47445
69e96e5500df Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents: 44887
diff changeset
   298
    fix x y assume x: "x \<in> U + V" and y: "y \<in> U + V"
69e96e5500df Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents: 44887
diff changeset
   299
    show "x + y \<in> U + V"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   300
    proof -
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   301
      from x obtain ux vx where "x = ux + vx" and "ux \<in> U" and "vx \<in> V" ..
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   302
      moreover
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   303
      from y obtain uy vy where "y = uy + vy" and "uy \<in> U" and "vy \<in> V" ..
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   304
      ultimately
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   305
      have "ux + uy \<in> U"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   306
        and "vx + vy \<in> V"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   307
        and "x + y = (ux + uy) + (vx + vy)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   308
        using x y by (simp_all add: add_ac)
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   309
      then show ?thesis ..
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   310
    qed
47445
69e96e5500df Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents: 44887
diff changeset
   311
    fix a show "a \<cdot> x \<in> U + V"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   312
    proof -
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   313
      from x obtain u v where "x = u + v" and "u \<in> U" and "v \<in> V" ..
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   314
      then have "a \<cdot> u \<in> U" and "a \<cdot> v \<in> V"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   315
        and "a \<cdot> x = (a \<cdot> u) + (a \<cdot> v)" by (simp_all add: distrib)
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   316
      then show ?thesis ..
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   317
    qed
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   318
  qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   319
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   320
61879
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61540
diff changeset
   321
text \<open>The sum of two subspaces is a vectorspace.\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   322
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   323
lemma sum_vs [intro?]:
47445
69e96e5500df Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents: 44887
diff changeset
   324
    "U \<unlhd> E \<Longrightarrow> V \<unlhd> E \<Longrightarrow> vectorspace E \<Longrightarrow> vectorspace (U + V)"
13547
wenzelm
parents: 13515
diff changeset
   325
  by (rule subspace.vectorspace) (rule sum_subspace)
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   326
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   327
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   328
subsection \<open>Direct sums\<close>
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   329
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   330
text \<open>
61879
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61540
diff changeset
   331
  The sum of \<open>U\<close> and \<open>V\<close> is called \<^emph>\<open>direct\<close>, iff the zero element is the only
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61540
diff changeset
   332
  common element of \<open>U\<close> and \<open>V\<close>. For every element \<open>x\<close> of the direct sum of
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61540
diff changeset
   333
  \<open>U\<close> and \<open>V\<close> the decomposition in \<open>x = u + v\<close> with \<open>u \<in> U\<close> and \<open>v \<in> V\<close> is
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61540
diff changeset
   334
  unique.
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   335
\<close>
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   336
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
   337
lemma decomp:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   338
  assumes "vectorspace E" "subspace U E" "subspace V E"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   339
  assumes direct: "U \<inter> V = {0}"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   340
    and u1: "u1 \<in> U" and u2: "u2 \<in> U"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   341
    and v1: "v1 \<in> V" and v2: "v2 \<in> V"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   342
    and sum: "u1 + v1 = u2 + v2"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   343
  shows "u1 = u2 \<and> v1 = v2"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   344
proof -
29234
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
   345
  interpret vectorspace E by fact
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
   346
  interpret subspace U E by fact
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
   347
  interpret subspace V E by fact
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   348
  show ?thesis
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   349
  proof
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   350
    have U: "vectorspace U"  (* FIXME: use interpret *)
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   351
      using \<open>subspace U E\<close> \<open>vectorspace E\<close> by (rule subspace.vectorspace)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   352
    have V: "vectorspace V"
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   353
      using \<open>subspace V E\<close> \<open>vectorspace E\<close> by (rule subspace.vectorspace)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   354
    from u1 u2 v1 v2 and sum have eq: "u1 - u2 = v2 - v1"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   355
      by (simp add: add_diff_swap)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   356
    from u1 u2 have u: "u1 - u2 \<in> U"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   357
      by (rule vectorspace.diff_closed [OF U])
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   358
    with eq have v': "v2 - v1 \<in> U" by (simp only:)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   359
    from v2 v1 have v: "v2 - v1 \<in> V"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   360
      by (rule vectorspace.diff_closed [OF V])
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   361
    with eq have u': " u1 - u2 \<in> V" by (simp only:)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   362
    
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   363
    show "u1 = u2"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   364
    proof (rule add_minus_eq)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   365
      from u1 show "u1 \<in> E" ..
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   366
      from u2 show "u2 \<in> E" ..
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   367
      from u u' and direct show "u1 - u2 = 0" by blast
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   368
    qed
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   369
    show "v1 = v2"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   370
    proof (rule add_minus_eq [symmetric])
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   371
      from v1 show "v1 \<in> E" ..
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   372
      from v2 show "v2 \<in> E" ..
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   373
      from v v' and direct show "v2 - v1 = 0" by blast
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   374
    qed
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   375
  qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   376
qed
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7567
diff changeset
   377
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   378
text \<open>
61540
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
   379
  An application of the previous lemma will be used in the proof of the
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
   380
  Hahn-Banach Theorem (see page \pageref{decomp-H-use}): for any element
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
   381
  \<open>y + a \<cdot> x\<^sub>0\<close> of the direct sum of a vectorspace \<open>H\<close> and the linear closure
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
   382
  of \<open>x\<^sub>0\<close> the components \<open>y \<in> H\<close> and \<open>a\<close> are uniquely determined.
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   383
\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   384
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
   385
lemma decomp_H':
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   386
  assumes "vectorspace E" "subspace H E"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   387
  assumes y1: "y1 \<in> H" and y2: "y2 \<in> H"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   388
    and x': "x' \<notin> H"  "x' \<in> E"  "x' \<noteq> 0"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   389
    and eq: "y1 + a1 \<cdot> x' = y2 + a2 \<cdot> x'"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   390
  shows "y1 = y2 \<and> a1 = a2"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   391
proof -
29234
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
   392
  interpret vectorspace E by fact
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
   393
  interpret subspace H E by fact
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   394
  show ?thesis
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   395
  proof
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   396
    have c: "y1 = y2 \<and> a1 \<cdot> x' = a2 \<cdot> x'"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   397
    proof (rule decomp)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   398
      show "a1 \<cdot> x' \<in> lin x'" ..
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   399
      show "a2 \<cdot> x' \<in> lin x'" ..
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   400
      show "H \<inter> lin x' = {0}"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   401
      proof
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   402
        show "H \<inter> lin x' \<subseteq> {0}"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   403
        proof
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   404
          fix x assume x: "x \<in> H \<inter> lin x'"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   405
          then obtain a where xx': "x = a \<cdot> x'"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   406
            by blast
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   407
          have "x = 0"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   408
          proof cases
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   409
            assume "a = 0"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   410
            with xx' and x' show ?thesis by simp
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   411
          next
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   412
            assume a: "a \<noteq> 0"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   413
            from x have "x \<in> H" ..
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   414
            with xx' have "inverse a \<cdot> a \<cdot> x' \<in> H" by simp
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   415
            with a and x' have "x' \<in> H" by (simp add: mult_assoc2)
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   416
            with \<open>x' \<notin> H\<close> show ?thesis by contradiction
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   417
          qed
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   418
          then show "x \<in> {0}" ..
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   419
        qed
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   420
        show "{0} \<subseteq> H \<inter> lin x'"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   421
        proof -
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   422
          have "0 \<in> H" using \<open>vectorspace E\<close> ..
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   423
          moreover have "0 \<in> lin x'" using \<open>x' \<in> E\<close> ..
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   424
          ultimately show ?thesis by blast
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   425
        qed
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   426
      qed
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   427
      show "lin x' \<unlhd> E" using \<open>x' \<in> E\<close> ..
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   428
    qed (rule \<open>vectorspace E\<close>, rule \<open>subspace H E\<close>, rule y1, rule y2, rule eq)
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   429
    then show "y1 = y2" ..
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   430
    from c have "a1 \<cdot> x' = a2 \<cdot> x'" ..
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   431
    with x' show "a1 = a2" by (simp add: mult_right_cancel)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   432
  qed
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   433
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   434
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   435
text \<open>
61540
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
   436
  Since for any element \<open>y + a \<cdot> x'\<close> of the direct sum of a vectorspace \<open>H\<close>
61879
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61540
diff changeset
   437
  and the linear closure of \<open>x'\<close> the components \<open>y \<in> H\<close> and \<open>a\<close> are unique, it
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61540
diff changeset
   438
  follows from \<open>y \<in> H\<close> that \<open>a = 0\<close>.
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   439
\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   440
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
   441
lemma decomp_H'_H:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   442
  assumes "vectorspace E" "subspace H E"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   443
  assumes t: "t \<in> H"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   444
    and x': "x' \<notin> H"  "x' \<in> E"  "x' \<noteq> 0"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   445
  shows "(SOME (y, a). t = y + a \<cdot> x' \<and> y \<in> H) = (t, 0)"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   446
proof -
29234
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
   447
  interpret vectorspace E by fact
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
   448
  interpret subspace H E by fact
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   449
  show ?thesis
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   450
  proof (rule, simp_all only: split_paired_all split_conv)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   451
    from t x' show "t = t + 0 \<cdot> x' \<and> t \<in> H" by simp
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   452
    fix y and a assume ya: "t = y + a \<cdot> x' \<and> y \<in> H"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   453
    have "y = t \<and> a = 0"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   454
    proof (rule decomp_H')
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   455
      from ya x' show "y + a \<cdot> x' = t + 0 \<cdot> x'" by simp
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   456
      from ya show "y \<in> H" ..
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   457
    qed (rule \<open>vectorspace E\<close>, rule \<open>subspace H E\<close>, rule t, (rule x')+)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   458
    with t x' show "(y, a) = (y + a \<cdot> x', 0)" by simp
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   459
  qed
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   460
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   461
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   462
text \<open>
61540
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
   463
  The components \<open>y \<in> H\<close> and \<open>a\<close> in \<open>y + a \<cdot> x'\<close> are unique, so the function
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
   464
  \<open>h'\<close> defined by \<open>h' (y + a \<cdot> x') = h y + a \<cdot> \<xi>\<close> is definite.
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   465
\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   466
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   467
lemma h'_definite:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   468
  fixes H
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   469
  assumes h'_def:
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 61879
diff changeset
   470
    "\<And>x. h' x =
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 61879
diff changeset
   471
      (let (y, a) = SOME (y, a). (x = y + a \<cdot> x' \<and> y \<in> H)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 61879
diff changeset
   472
       in (h y) + a * xi)"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   473
    and x: "x = y + a \<cdot> x'"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26821
diff changeset
   474
  assumes "vectorspace E" "subspace H E"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   475
  assumes y: "y \<in> H"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   476
    and x': "x' \<notin> H"  "x' \<in> E"  "x' \<noteq> 0"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   477
  shows "h' x = h y + a * xi"
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
   478
proof -
29234
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
   479
  interpret vectorspace E by fact
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
   480
  interpret subspace H E by fact
47445
69e96e5500df Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents: 44887
diff changeset
   481
  from x y x' have "x \<in> H + lin x'" by auto
63910
wenzelm
parents: 63040
diff changeset
   482
  have "\<exists>!(y, a). x = y + a \<cdot> x' \<and> y \<in> H" (is "\<exists>!p. ?P p")
18689
a50587cd8414 prefer ex1I over ex_ex1I in single-step reasoning;
wenzelm
parents: 16417
diff changeset
   483
  proof (rule ex_ex1I)
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   484
    from x y show "\<exists>p. ?P p" by blast
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   485
    fix p q assume p: "?P p" and q: "?P q"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   486
    show "p = q"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   487
    proof -
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   488
      from p have xp: "x = fst p + snd p \<cdot> x' \<and> fst p \<in> H"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   489
        by (cases p) simp
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   490
      from q have xq: "x = fst q + snd q \<cdot> x' \<and> fst q \<in> H"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   491
        by (cases q) simp
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   492
      have "fst p = fst q \<and> snd p = snd q"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   493
      proof (rule decomp_H')
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   494
        from xp show "fst p \<in> H" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   495
        from xq show "fst q \<in> H" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   496
        from xp and xq show "fst p + snd p \<cdot> x' = fst q + snd q \<cdot> x'"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   497
          by simp
58744
c434e37f290e update_cartouches;
wenzelm
parents: 47445
diff changeset
   498
      qed (rule \<open>vectorspace E\<close>, rule \<open>subspace H E\<close>, (rule x')+)
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   499
      then show ?thesis by (cases p, cases q) simp
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   500
    qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   501
  qed
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   502
  then have eq: "(SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H) = (y, a)"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   503
    by (rule some1_equality) (simp add: x y)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   504
  with h'_def show "h' x = h y + a * xi" by (simp add: Let_def)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   505
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   506
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
   507
end