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