src/HOL/Real/HahnBanach/Subspace.thy
author paulson
Fri, 15 Sep 2000 12:39:57 +0200
changeset 9969 4753185f1dd2
parent 9941 fe05af7ec816
child 10309 a7f961fb62c6
permissions -rw-r--r--
renamed (most of...) the select rules
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
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
     6
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
     7
header {* Subspaces *}
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
     8
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
     9
theory Subspace = VectorSpace:
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    10
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    11
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    12
subsection {* Definition *}
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    13
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    14
text {* A non-empty subset $U$ of a vector space $V$ is a 
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    15
\emph{subspace} of $V$, iff $U$ is closed under addition and 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    16
scalar multiplication. *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    17
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    18
constdefs 
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    19
  is_subspace ::  "['a::{plus, minus, zero} set, 'a set] => bool"
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    20
  "is_subspace U V == U \<noteq> {} \<and> U <= V 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    21
     \<and> (\<forall>x \<in> U. \<forall>y \<in> U. \<forall>a. x + y \<in> U \<and> a \<cdot> x\<in> U)"
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    22
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
    23
lemma subspaceI [intro]: 
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    24
  "[| 0 \<in> U; U <= V; \<forall>x \<in> U. \<forall>y \<in> U. (x + y \<in> U); 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    25
  \<forall>x \<in> U. \<forall>a. a \<cdot> x \<in> U |]
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    26
  ==> is_subspace U V"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    27
proof (unfold is_subspace_def, intro conjI) 
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    28
  assume "0 \<in> U" thus "U \<noteq> {}" by fast
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    29
qed (simp+)
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    30
9408
d3d56e1d2ec1 classical atts now intro! / intro / intro?;
wenzelm
parents: 9374
diff changeset
    31
lemma subspace_not_empty [intro?]: "is_subspace U V ==> U \<noteq> {}"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    32
  by (unfold is_subspace_def) simp 
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
    33
9408
d3d56e1d2ec1 classical atts now intro! / intro / intro?;
wenzelm
parents: 9374
diff changeset
    34
lemma subspace_subset [intro?]: "is_subspace U V ==> U <= V"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    35
  by (unfold is_subspace_def) simp
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
    36
9408
d3d56e1d2ec1 classical atts now intro! / intro / intro?;
wenzelm
parents: 9374
diff changeset
    37
lemma subspace_subsetD [simp, intro?]: 
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    38
  "[| is_subspace U V; x \<in> U |] ==> x \<in> V"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    39
  by (unfold is_subspace_def) force
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    40
9408
d3d56e1d2ec1 classical atts now intro! / intro / intro?;
wenzelm
parents: 9374
diff changeset
    41
lemma subspace_add_closed [simp, intro?]: 
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    42
  "[| is_subspace U V; x \<in> U; y \<in> U |] ==> x + y \<in> U"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    43
  by (unfold is_subspace_def) simp
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    44
9408
d3d56e1d2ec1 classical atts now intro! / intro / intro?;
wenzelm
parents: 9374
diff changeset
    45
lemma subspace_mult_closed [simp, intro?]: 
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    46
  "[| is_subspace U V; x \<in> U |] ==> a \<cdot> x \<in> U"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    47
  by (unfold is_subspace_def) simp
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
    48
9408
d3d56e1d2ec1 classical atts now intro! / intro / intro?;
wenzelm
parents: 9374
diff changeset
    49
lemma subspace_diff_closed [simp, intro?]: 
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    50
  "[| is_subspace U V; is_vectorspace V; x \<in> U; y \<in> U |] 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    51
  ==> x - y \<in> U"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    52
  by (simp! add: diff_eq1 negate_eq1)
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    53
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    54
text {* Similar as for linear spaces, the existence of the 
7978
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
    55
zero element in every subspace follows from the non-emptiness 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    56
of the carrier set and by vector space laws.*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    57
9408
d3d56e1d2ec1 classical atts now intro! / intro / intro?;
wenzelm
parents: 9374
diff changeset
    58
lemma zero_in_subspace [intro?]:
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    59
  "[| is_subspace U V; is_vectorspace V |] ==> 0 \<in> U"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    60
proof - 
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    61
  assume "is_subspace U V" and v: "is_vectorspace V"
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    62
  have "U \<noteq> {}" ..
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    63
  hence "\<exists>x. x \<in> U" by force
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    64
  thus ?thesis 
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    65
  proof 
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    66
    fix x assume u: "x \<in> U" 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    67
    hence "x \<in> V" by (simp!)
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    68
    with v have "0 = x - x" by (simp!)
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    69
    also have "... \<in> U" by (rule subspace_diff_closed)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    70
    finally show ?thesis .
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    71
  qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    72
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    73
9408
d3d56e1d2ec1 classical atts now intro! / intro / intro?;
wenzelm
parents: 9374
diff changeset
    74
lemma subspace_neg_closed [simp, intro?]: 
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    75
  "[| is_subspace U V; is_vectorspace V; x \<in> U |] ==> - x \<in> U"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    76
  by (simp add: negate_eq1)
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    77
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    78
text_raw {* \medskip *}
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    79
text {* Further derived laws: every subspace is a vector space. *}
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    80
9408
d3d56e1d2ec1 classical atts now intro! / intro / intro?;
wenzelm
parents: 9374
diff changeset
    81
lemma subspace_vs [intro?]:
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    82
  "[| is_subspace U V; is_vectorspace V |] ==> is_vectorspace U"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    83
proof -
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    84
  assume "is_subspace U V" "is_vectorspace V"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    85
  show ?thesis
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    86
  proof 
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    87
    show "0 \<in> U" ..
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    88
    show "\<forall>x \<in> U. \<forall>a. a \<cdot> x \<in> U" by (simp!)
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    89
    show "\<forall>x \<in> U. \<forall>y \<in> U. x + y \<in> U" by (simp!)
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    90
    show "\<forall>x \<in> U. - x = -#1 \<cdot> x" by (simp! add: negate_eq1)
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
    91
    show "\<forall>x \<in> U. \<forall>y \<in> U. x - y =  x + - y" 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    92
      by (simp! add: diff_eq1)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    93
  qed (simp! add: vs_add_mult_distrib1 vs_add_mult_distrib2)+
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    94
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
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
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    98
lemma subspace_refl [intro]: "is_vectorspace V ==> is_subspace V V"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
    99
proof 
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   100
  assume "is_vectorspace V"
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   101
  show "0 \<in> V" ..
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   102
  show "V <= V" ..
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   103
  show "\<forall>x \<in> V. \<forall>y \<in> V. x + y \<in> V" by (simp!)
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   104
  show "\<forall>x \<in> V. \<forall>a. a \<cdot> x \<in> V" by (simp!)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   105
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   106
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   107
text {* The subspace relation is transitive. *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   108
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   109
lemma subspace_trans: 
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   110
  "[| is_subspace U V; is_vectorspace V; is_subspace V W |] 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   111
  ==> is_subspace U W"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   112
proof 
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   113
  assume "is_subspace U V" "is_subspace V W" "is_vectorspace V"
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   114
  show "0 \<in> U" ..
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7567
diff changeset
   115
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   116
  have "U <= V" ..
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   117
  also have "V <= W" ..
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   118
  finally show "U <= W" .
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7567
diff changeset
   119
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   120
  show "\<forall>x \<in> U. \<forall>y \<in> U. x + y \<in> U" 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   121
  proof (intro ballI)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   122
    fix x y assume "x \<in> U" "y \<in> U"
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   123
    show "x + y \<in> U" by (simp!)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   124
  qed
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7567
diff changeset
   125
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   126
  show "\<forall>x \<in> U. \<forall>a. a \<cdot> x \<in> U"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   127
  proof (intro ballI allI)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   128
    fix x a assume "x \<in> U"
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   129
    show "a \<cdot> x \<in> U" by (simp!)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   130
  qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   131
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   132
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   133
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   134
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   135
subsection {* Linear closure *}
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   136
7978
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
   137
text {* The \emph{linear closure} of a vector $x$ is the set of all
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   138
scalar multiples of $x$. *}
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   139
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   140
constdefs
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   141
  lin :: "('a::{minus,plus,zero}) => 'a set"
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   142
  "lin x == {a \<cdot> x | a. True}" 
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   143
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   144
lemma linD: "x \<in> lin v = (\<exists>a::real. x = a \<cdot> v)"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   145
  by (unfold lin_def) fast
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   146
9408
d3d56e1d2ec1 classical atts now intro! / intro / intro?;
wenzelm
parents: 9374
diff changeset
   147
lemma linI [intro?]: "a \<cdot> x0 \<in> lin x0"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   148
  by (unfold lin_def) fast
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7567
diff changeset
   149
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   150
text {* Every vector is contained in its linear closure. *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   151
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   152
lemma x_lin_x: "[| is_vectorspace V; x \<in> V |] ==> x \<in> lin x"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   153
proof (unfold lin_def, intro CollectI exI conjI)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   154
  assume "is_vectorspace V" "x \<in> V"
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   155
  show "x = #1 \<cdot> x" by (simp!)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   156
qed simp
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   157
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   158
text {* Any linear closure is a subspace. *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   159
9408
d3d56e1d2ec1 classical atts now intro! / intro / intro?;
wenzelm
parents: 9374
diff changeset
   160
lemma lin_subspace [intro?]: 
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   161
  "[| is_vectorspace V; x \<in> V |] ==> is_subspace (lin x) V"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   162
proof
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   163
  assume "is_vectorspace V" "x \<in> V"
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   164
  show "0 \<in> lin x" 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   165
  proof (unfold lin_def, intro CollectI exI conjI)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   166
    show "0 = (#0::real) \<cdot> x" by (simp!)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   167
  qed simp
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
   168
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   169
  show "lin x <= V"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   170
  proof (unfold lin_def, intro subsetI, elim CollectE exE conjE) 
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   171
    fix xa a assume "xa = a \<cdot> x" 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   172
    show "xa \<in> V" by (simp!)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   173
  qed
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
   174
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   175
  show "\<forall>x1 \<in> lin x. \<forall>x2 \<in> lin x. x1 + x2 \<in> lin x" 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   176
  proof (intro ballI)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   177
    fix x1 x2 assume "x1 \<in> lin x" "x2 \<in> lin x" 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   178
    thus "x1 + x2 \<in> lin x"
7978
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
   179
    proof (unfold lin_def, elim CollectE exE conjE, 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   180
      intro CollectI exI conjI)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   181
      fix a1 a2 assume "x1 = a1 \<cdot> x" "x2 = a2 \<cdot> x"
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   182
      show "x1 + x2 = (a1 + a2) \<cdot> x" 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   183
        by (simp! add: vs_add_mult_distrib2)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   184
    qed simp
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   185
  qed
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
   186
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   187
  show "\<forall>xa \<in> lin x. \<forall>a. a \<cdot> xa \<in> lin x" 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   188
  proof (intro ballI allI)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   189
    fix x1 a assume "x1 \<in> lin x" 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   190
    thus "a \<cdot> x1 \<in> lin x"
7978
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
   191
    proof (unfold lin_def, elim CollectE exE conjE,
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   192
      intro CollectI exI conjI)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   193
      fix a1 assume "x1 = a1 \<cdot> x"
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   194
      show "a \<cdot> x1 = (a * a1) \<cdot> x" by (simp!)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   195
    qed simp
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   196
  qed 
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   197
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
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
9408
d3d56e1d2ec1 classical atts now intro! / intro / intro?;
wenzelm
parents: 9374
diff changeset
   201
lemma lin_vs [intro?]: 
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   202
  "[| is_vectorspace V; x \<in> V |] ==> is_vectorspace (lin x)"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   203
proof (rule subspace_vs)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   204
  assume "is_vectorspace V" "x \<in> V"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   205
  show "is_subspace (lin x) V" ..
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   206
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   207
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   208
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   209
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   210
subsection {* Sum of two vectorspaces *}
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   211
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   212
text {* The \emph{sum} of two vectorspaces $U$ and $V$ is the set of
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   213
all sums of elements from $U$ and $V$. *}
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   214
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   215
instance set :: (plus) plus by intro_classes
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   216
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   217
defs vs_sum_def:
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   218
  "U + V == {u + v | u v. u \<in> U \<and> v \<in> V}" (***
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   219
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   220
constdefs 
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   221
  vs_sum :: 
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   222
  "['a::{plus, minus, zero} set, 'a set] => 'a set"         (infixl "+" 65)
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   223
  "vs_sum U V == {x. \<exists>u \<in> U. \<exists>v \<in> V. x = u + v}";
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   224
***)
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   225
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   226
lemma vs_sumD: 
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   227
  "x \<in> U + V = (\<exists>u \<in> U. \<exists>v \<in> V. x = u + v)"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   228
    by (unfold vs_sum_def) fast
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   229
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   230
lemmas vs_sumE = vs_sumD [THEN iffD1, elim_format, standard]
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
   231
9408
d3d56e1d2ec1 classical atts now intro! / intro / intro?;
wenzelm
parents: 9374
diff changeset
   232
lemma vs_sumI [intro?]: 
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   233
  "[| x \<in> U; y \<in> V; t= x + y |] ==> t \<in> U + V"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   234
  by (unfold vs_sum_def) fast
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   235
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   236
text{* $U$ is a subspace of $U + V$. *}
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   237
9408
d3d56e1d2ec1 classical atts now intro! / intro / intro?;
wenzelm
parents: 9374
diff changeset
   238
lemma subspace_vs_sum1 [intro?]: 
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   239
  "[| is_vectorspace U; is_vectorspace V |]
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   240
  ==> is_subspace U (U + V)"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   241
proof 
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   242
  assume "is_vectorspace U" "is_vectorspace V"
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   243
  show "0 \<in> U" ..
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   244
  show "U <= U + V"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   245
  proof (intro subsetI vs_sumI)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   246
  fix x assume "x \<in> U"
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   247
    show "x = x + 0" by (simp!)
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   248
    show "0 \<in> V" by (simp!)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   249
  qed
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   250
  show "\<forall>x \<in> U. \<forall>y \<in> U. x + y \<in> U" 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   251
  proof (intro ballI)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   252
    fix x y assume "x \<in> U" "y \<in> U" show "x + y \<in> U" by (simp!)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   253
  qed
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   254
  show "\<forall>x \<in> U. \<forall>a. a \<cdot> x \<in> U" 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   255
  proof (intro ballI allI)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   256
    fix x a assume "x \<in> U" show "a \<cdot> x \<in> U" by (simp!)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   257
  qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   258
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   259
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   260
text{* The sum of two subspaces is again a subspace.*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   261
9408
d3d56e1d2ec1 classical atts now intro! / intro / intro?;
wenzelm
parents: 9374
diff changeset
   262
lemma vs_sum_subspace [intro?]: 
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
   263
  "[| is_subspace U E; is_subspace V E; is_vectorspace E |] 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   264
  ==> is_subspace (U + V) E"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   265
proof 
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   266
  assume "is_subspace U E" "is_subspace V E" "is_vectorspace E"
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   267
  show "0 \<in> U + V"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   268
  proof (intro vs_sumI)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   269
    show "0 \<in> U" ..
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   270
    show "0 \<in> V" ..
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   271
    show "(0::'a) = 0 + 0" by (simp!)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   272
  qed
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
   273
  
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   274
  show "U + V <= E"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   275
  proof (intro subsetI, elim vs_sumE bexE)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   276
    fix x u v assume "u \<in> U" "v \<in> V" "x = u + v"
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   277
    show "x \<in> E" by (simp!)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   278
  qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   279
  
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   280
  show "\<forall>x \<in> U + V. \<forall>y \<in> U + V. x + y \<in> U + V"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   281
  proof (intro ballI)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   282
    fix x y assume "x \<in> U + V" "y \<in> U + V"
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   283
    thus "x + y \<in> U + V"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   284
    proof (elim vs_sumE bexE, intro vs_sumI)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   285
      fix ux vx uy vy 
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   286
      assume "ux \<in> U" "vx \<in> V" "x = ux + vx" 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   287
	and "uy \<in> U" "vy \<in> V" "y = uy + vy"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   288
      show "x + y = (ux + uy) + (vx + vy)" by (simp!)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   289
    qed (simp!)+
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   290
  qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   291
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   292
  show "\<forall>x \<in> U + V. \<forall>a. a \<cdot> x \<in> U + V"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   293
  proof (intro ballI allI)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   294
    fix x a assume "x \<in> U + V"
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   295
    thus "a \<cdot> x \<in> U + V"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   296
    proof (elim vs_sumE bexE, intro vs_sumI)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   297
      fix a x u v assume "u \<in> U" "v \<in> V" "x = u + v"
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   298
      show "a \<cdot> x = (a \<cdot> u) + (a \<cdot> v)" 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   299
        by (simp! add: vs_add_mult_distrib1)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   300
    qed (simp!)+
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   301
  qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   302
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   303
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   304
text{* The sum of two subspaces is a vectorspace. *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   305
9408
d3d56e1d2ec1 classical atts now intro! / intro / intro?;
wenzelm
parents: 9374
diff changeset
   306
lemma vs_sum_vs [intro?]: 
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
   307
  "[| is_subspace U E; is_subspace V E; is_vectorspace E |] 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   308
  ==> is_vectorspace (U + V)"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   309
proof (rule subspace_vs)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   310
  assume "is_subspace U E" "is_subspace V E" "is_vectorspace E"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   311
  show "is_subspace (U + V) E" ..
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   312
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   313
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   314
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   315
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   316
subsection {* Direct sums *}
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   317
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   318
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   319
text {* The sum of $U$ and $V$ is called \emph{direct}, iff the zero 
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   320
element is the only common element of $U$ and $V$. For every element
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   321
$x$ of the direct sum of $U$ and $V$ the decomposition in
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   322
$x = u + v$ with $u \in U$ and $v \in V$ is unique.*} 
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   323
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   324
lemma decomp: 
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   325
  "[| is_vectorspace E; is_subspace U E; is_subspace V E; 
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   326
  U \<inter> V = {0}; u1 \<in> U; u2 \<in> U; v1 \<in> V; v2 \<in> V; u1 + v1 = u2 + v2 |] 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   327
  ==> u1 = u2 \<and> v1 = v2" 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   328
proof 
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   329
  assume "is_vectorspace E" "is_subspace U E" "is_subspace V E"  
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   330
    "U \<inter> V = {0}" "u1 \<in> U" "u2 \<in> U" "v1 \<in> V" "v2 \<in> V" 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   331
    "u1 + v1 = u2 + v2" 
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   332
  have eq: "u1 - u2 = v2 - v1" by (simp! add: vs_add_diff_swap)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   333
  have u: "u1 - u2 \<in> U" by (simp!) 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   334
  with eq have v': "v2 - v1 \<in> U" by simp 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   335
  have v: "v2 - v1 \<in> V" by (simp!) 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   336
  with eq have u': "u1 - u2 \<in> V" by simp
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7567
diff changeset
   337
  
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   338
  show "u1 = u2"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   339
  proof (rule vs_add_minus_eq)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   340
    show "u1 - u2 = 0" by (rule Int_singletonD [OF _ u u']) 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   341
    show "u1 \<in> E" ..
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   342
    show "u2 \<in> E" ..
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   343
  qed
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7567
diff changeset
   344
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   345
  show "v1 = v2"
9623
3ade112482af renamed 'RS' to 'THEN';
wenzelm
parents: 9408
diff changeset
   346
  proof (rule vs_add_minus_eq [symmetric])
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   347
    show "v2 - v1 = 0" by (rule Int_singletonD [OF _ v' v])
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   348
    show "v1 \<in> E" ..
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   349
    show "v2 \<in> E" ..
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   350
  qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   351
qed
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7567
diff changeset
   352
7978
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
   353
text {* An application of the previous lemma will be used in the proof
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   354
of the Hahn-Banach Theorem (see page \pageref{decomp-H-use}): for any
7978
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
   355
element $y + a \mult x_0$ of the direct sum of a vectorspace $H$ and
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
   356
the linear closure of $x_0$ the components $y \in H$ and $a$ are
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   357
uniquely determined. *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   358
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   359
lemma decomp_H': 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   360
  "[| is_vectorspace E; is_subspace H E; y1 \<in> H; y2 \<in> H; 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   361
  x' \<notin> H; x' \<in> E; x' \<noteq> 0; y1 + a1 \<cdot> x' = y2 + a2 \<cdot> x' |]
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   362
  ==> y1 = y2 \<and> a1 = a2"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   363
proof
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7567
diff changeset
   364
  assume "is_vectorspace E" and h: "is_subspace H E"
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   365
     and "y1 \<in> H" "y2 \<in> H" "x' \<notin> H" "x' \<in> E" "x' \<noteq> 0" 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   366
         "y1 + a1 \<cdot> x' = y2 + a2 \<cdot> x'"
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   367
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   368
  have c: "y1 = y2 \<and> a1 \<cdot> x' = a2 \<cdot> x'"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   369
  proof (rule decomp) 
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   370
    show "a1 \<cdot> x' \<in> lin x'" .. 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   371
    show "a2 \<cdot> x' \<in> lin x'" ..
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   372
    show "H \<inter> (lin x') = {0}" 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   373
    proof
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   374
      show "H \<inter> lin x' <= {0}" 
9623
3ade112482af renamed 'RS' to 'THEN';
wenzelm
parents: 9408
diff changeset
   375
      proof (intro subsetI, elim IntE, rule singleton_iff [THEN iffD2])
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   376
        fix x assume "x \<in> H" "x \<in> lin x'" 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   377
        thus "x = 0"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   378
        proof (unfold lin_def, elim CollectE exE conjE)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   379
          fix a assume "x = a \<cdot> x'"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   380
          show ?thesis
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   381
          proof cases
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   382
            assume "a = (#0::real)" show ?thesis by (simp!)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   383
          next
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   384
            assume "a \<noteq> (#0::real)" 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   385
            from h have "rinv a \<cdot> a \<cdot> x' \<in> H" 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   386
              by (rule subspace_mult_closed) (simp!)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   387
            also have "rinv a \<cdot> a \<cdot> x' = x'" by (simp!)
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   388
            finally have "x' \<in> H" .
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   389
            thus ?thesis by contradiction
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   390
          qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   391
       qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   392
      qed
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   393
      show "{0} <= H \<inter> lin x'"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   394
      proof -
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   395
	have "0 \<in> H \<inter> lin x'"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   396
	proof (rule IntI)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   397
	  show "0 \<in> H" ..
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   398
	  from lin_vs show "0 \<in> lin x'" ..
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   399
	qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   400
	thus ?thesis by simp
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   401
      qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   402
    qed
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   403
    show "is_subspace (lin x') E" ..
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   404
  qed
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7567
diff changeset
   405
  
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   406
  from c show "y1 = y2" by simp
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7567
diff changeset
   407
  
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   408
  show  "a1 = a2" 
9623
3ade112482af renamed 'RS' to 'THEN';
wenzelm
parents: 9408
diff changeset
   409
  proof (rule vs_mult_right_cancel [THEN iffD1])
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   410
    from c show "a1 \<cdot> x' = a2 \<cdot> x'" by simp
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   411
  qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   412
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   413
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   414
text {* Since for any element $y + a \mult x'$ of the direct sum 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   415
of a vectorspace $H$ and the linear closure of $x'$ the components
7978
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
   416
$y\in H$ and $a$ are unique, it follows from $y\in H$ that 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   417
$a = 0$.*} 
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   418
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   419
lemma decomp_H'_H: 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   420
  "[| is_vectorspace E; is_subspace H E; t \<in> H; x' \<notin> H; x' \<in> E;
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   421
  x' \<noteq> 0 |] 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   422
  ==> (SOME (y, a). t = y + a \<cdot> x' \<and> y \<in> H) = (t, (#0::real))"
9370
cccba6147dae use split_tupled_all;
wenzelm
parents: 9035
diff changeset
   423
proof (rule, unfold split_tupled_all)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   424
  assume "is_vectorspace E" "is_subspace H E" "t \<in> H" "x' \<notin> H" "x' \<in> E"
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   425
    "x' \<noteq> 0"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   426
  have h: "is_vectorspace H" ..
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   427
  fix y a presume t1: "t = y + a \<cdot> x'" and "y \<in> H"
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   428
  have "y = t \<and> a = (#0::real)" 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   429
    by (rule decomp_H') (assumption | (simp!))+
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   430
  thus "(y, a) = (t, (#0::real))" by (simp!)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   431
qed (simp!)+
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   432
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   433
text {* The components $y\in H$ and $a$ in $y \plus a \mult x'$ 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   434
are unique, so the function $h'$ defined by 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   435
$h' (y \plus a \mult x') = h y + a \cdot \xi$ is definite. *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   436
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   437
lemma h'_definite:
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   438
  "[| h' == (\<lambda>x. let (y, a) = SOME (y, a). (x = y + a \<cdot> x' \<and> y \<in> H)
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
   439
                in (h y) + a * xi);
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   440
  x = y + a \<cdot> x'; is_vectorspace E; is_subspace H E;
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   441
  y \<in> H; x' \<notin> H; x' \<in> E; x' \<noteq> 0 |]
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   442
  ==> h' x = h y + a * xi"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   443
proof -  
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   444
  assume 
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   445
    "h' == (\<lambda>x. let (y, a) = SOME (y, a). (x = y + a \<cdot> x' \<and> y \<in> H)
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   446
               in (h y) + a * xi)"
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   447
    "x = y + a \<cdot> x'" "is_vectorspace E" "is_subspace H E"
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   448
    "y \<in> H" "x' \<notin> H" "x' \<in> E" "x' \<noteq> 0"
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   449
  have "x \<in> H + (lin x')" 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   450
    by (simp! add: vs_sum_def lin_def) force+
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   451
  have "\<exists>! xa. ((\<lambda>(y, a). x = y + a \<cdot> x' \<and> y \<in> H) xa)" 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   452
  proof
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   453
    show "\<exists>xa. ((\<lambda>(y, a). x = y + a \<cdot> x' \<and> y \<in> H) xa)"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   454
      by (force!)
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   455
  next
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   456
    fix xa ya
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   457
    assume "(\<lambda>(y,a). x = y + a \<cdot> x' \<and> y \<in> H) xa"
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   458
           "(\<lambda>(y,a). x = y + a \<cdot> x' \<and> y \<in> H) ya"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   459
    show "xa = ya" 
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   460
    proof -
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   461
      show "fst xa = fst ya \<and> snd xa = snd ya ==> xa = ya" 
9370
cccba6147dae use split_tupled_all;
wenzelm
parents: 9035
diff changeset
   462
        by (simp add: Pair_fst_snd_eq)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   463
      have x: "x = fst xa + snd xa \<cdot> x' \<and> fst xa \<in> H" 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   464
        by (force!)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   465
      have y: "x = fst ya + snd ya \<cdot> x' \<and> fst ya \<in> H" 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   466
        by (force!)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   467
      from x y show "fst xa = fst ya \<and> snd xa = snd ya" 
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   468
        by (elim conjE) (rule decomp_H', (simp!)+)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   469
    qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   470
  qed
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   471
  hence eq: "(SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H) = (y, a)" 
9969
4753185f1dd2 renamed (most of...) the select rules
paulson
parents: 9941
diff changeset
   472
    by (rule some1_equality) (force!)
9374
153853af318b - xsymbols for
bauerg
parents: 9370
diff changeset
   473
  thus "h' x = h y + a * xi" by (simp! add: Let_def)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   474
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   475
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 9013
diff changeset
   476
end