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