src/HOL/Real/HahnBanach/HahnBanachSupLemmas.thy
author wenzelm
Thu, 22 Aug 2002 20:49:43 +0200
changeset 13515 a6a7025fd7e8
parent 10687 c186279eecea
child 13547 bf399f3bd7dc
permissions -rw-r--r--
updated to use locales (still some rough edges);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/Real/HahnBanach/HahnBanachSupLemmas.thy
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
     3
    Author:     Gertrud Bauer, TU Munich
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
     4
*)
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
     5
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
     6
header {* The supremum w.r.t.~the function order *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
     7
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
     8
theory HahnBanachSupLemmas = FunctionNorm + ZornLemma:
9256
f9a6670427fb completed TYPES version of HahnBanach;
bauerg
parents: 9035
diff changeset
     9
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
    10
text {*
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
    11
  This section contains some lemmas that will be used in the proof of
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
    12
  the Hahn-Banach Theorem.  In this section the following context is
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
    13
  presumed.  Let @{text E} be a real vector space with a seminorm
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
    14
  @{text p} on @{text E}.  @{text F} is a subspace of @{text E} and
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
    15
  @{text f} a linear form on @{text F}. We consider a chain @{text c}
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    16
  of norm-preserving extensions of @{text f}, such that @{text "\<Union>c =
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    17
  graph H h"}.  We will show some properties about the limit function
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    18
  @{text h}, i.e.\ the supremum of the chain @{text c}.
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    19
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    20
  \medskip Let @{text c} be a chain of norm-preserving extensions of
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    21
  the function @{text f} and let @{text "graph H h"} be the supremum
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    22
  of @{text c}.  Every element in @{text H} is member of one of the
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
    23
  elements of the chain.
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
    24
*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    25
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    26
lemma some_H'h't:
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    27
  assumes M: "M = norm_pres_extensions E p F f"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    28
    and cM: "c \<in> chain M"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    29
    and u: "graph H h = \<Union>c"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    30
    and x: "x \<in> H"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    31
  shows "\<exists>H' h'. graph H' h' \<in> c
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    32
    \<and> (x, h x) \<in> graph H' h'
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    33
    \<and> linearform H' h' \<and> H' \<unlhd> E
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    34
    \<and> F \<unlhd> H' \<and> graph F f \<subseteq> graph H' h'
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    35
    \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    36
proof -
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    37
  from x have "(x, h x) \<in> graph H h" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    38
  also from u have "\<dots> = \<Union>c" .
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    39
  finally obtain g where gc: "g \<in> c" and gh: "(x, h x) \<in> g" by blast
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    40
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    41
  from cM have "c \<subseteq> M" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    42
  with gc have "g \<in> M" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    43
  also from M have "\<dots> = norm_pres_extensions E p F f" .
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    44
  finally obtain H' and h' where g: "g = graph H' h'"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    45
    and * : "linearform H' h'"  "H' \<unlhd> E"  "F \<unlhd> H'"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    46
      "graph F f \<subseteq> graph H' h'"  "\<forall>x \<in> H'. h' x \<le> p x" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    47
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    48
  from gc and g have "graph H' h' \<in> c" by (simp only:)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    49
  moreover from gh and g have "(x, h x) \<in> graph H' h'" by (simp only:)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    50
  ultimately show ?thesis using * by blast
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    51
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    52
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
    53
text {*
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
    54
  \medskip Let @{text c} be a chain of norm-preserving extensions of
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
    55
  the function @{text f} and let @{text "graph H h"} be the supremum
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
    56
  of @{text c}.  Every element in the domain @{text H} of the supremum
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
    57
  function is member of the domain @{text H'} of some function @{text
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
    58
  h'}, such that @{text h} extends @{text h'}.
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    59
*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    60
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
    61
lemma some_H'h':
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    62
  assumes M: "M = norm_pres_extensions E p F f"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    63
    and cM: "c \<in> chain M"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    64
    and u: "graph H h = \<Union>c"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    65
    and x: "x \<in> H"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    66
  shows "\<exists>H' h'. x \<in> H' \<and> graph H' h' \<subseteq> graph H h
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    67
    \<and> linearform H' h' \<and> H' \<unlhd> E \<and> F \<unlhd> H'
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    68
    \<and> graph F f \<subseteq> graph H' h' \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    69
proof -
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    70
  from M cM u x obtain H' h' where
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    71
      x_hx: "(x, h x) \<in> graph H' h'"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    72
    and c: "graph H' h' \<in> c"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    73
    and * : "linearform H' h'"  "H' \<unlhd> E"  "F \<unlhd> H'"
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
    74
      "graph F f \<subseteq> graph H' h'"  "\<forall>x \<in> H'. h' x \<le> p x"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    75
    by (rule some_H'h't [elim_format]) blast
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    76
  from x_hx have "x \<in> H'" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    77
  moreover from cM u c have "graph H' h' \<subseteq> graph H h"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    78
    by (simp only: chain_ball_Union_upper)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    79
  ultimately show ?thesis using * by blast
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    80
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    81
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
    82
text {*
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
    83
  \medskip Any two elements @{text x} and @{text y} in the domain
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
    84
  @{text H} of the supremum function @{text h} are both in the domain
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
    85
  @{text H'} of some function @{text h'}, such that @{text h} extends
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
    86
  @{text h'}.
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
    87
*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    88
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
    89
lemma some_H'h'2:
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    90
  assumes M: "M = norm_pres_extensions E p F f"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    91
    and cM: "c \<in> chain M"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    92
    and u: "graph H h = \<Union>c"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    93
    and x: "x \<in> H"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    94
    and y: "y \<in> H"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    95
  shows "\<exists>H' h'. x \<in> H' \<and> y \<in> H'
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    96
    \<and> graph H' h' \<subseteq> graph H h
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    97
    \<and> linearform H' h' \<and> H' \<unlhd> E \<and> F \<unlhd> H'
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    98
    \<and> graph F f \<subseteq> graph H' h' \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    99
proof -
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   100
  txt {* @{text y} is in the domain @{text H''} of some function @{text h''},
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   101
  such that @{text h} extends @{text h''}. *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   102
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   103
  from M cM u and y obtain H' h' where
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   104
      y_hy: "(y, h y) \<in> graph H' h'"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   105
    and c': "graph H' h' \<in> c"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   106
    and * :
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   107
      "linearform H' h'"  "H' \<unlhd> E"  "F \<unlhd> H'"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   108
      "graph F f \<subseteq> graph H' h'"  "\<forall>x \<in> H'. h' x \<le> p x"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   109
    by (rule some_H'h't [elim_format]) blast
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   110
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   111
  txt {* @{text x} is in the domain @{text H'} of some function @{text h'},
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   112
    such that @{text h} extends @{text h'}. *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   113
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   114
  from M cM u and x obtain H'' h'' where
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   115
      x_hx: "(x, h x) \<in> graph H'' h''"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   116
    and c'': "graph H'' h'' \<in> c"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   117
    and ** :
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   118
      "linearform H'' h''"  "H'' \<unlhd> E"  "F \<unlhd> H''"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   119
      "graph F f \<subseteq> graph H'' h''"  "\<forall>x \<in> H''. h'' x \<le> p x"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   120
    by (rule some_H'h't [elim_format]) blast
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   121
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   122
  txt {* Since both @{text h'} and @{text h''} are elements of the chain,
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   123
    @{text h''} is an extension of @{text h'} or vice versa. Thus both
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   124
    @{text x} and @{text y} are contained in the greater
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   125
    one. \label{cases1}*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   126
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   127
  from cM have "graph H'' h'' \<subseteq> graph H' h' \<or> graph H' h' \<subseteq> graph H'' h''"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   128
    (is "?case1 \<or> ?case2") ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   129
  then show ?thesis
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   130
  proof
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   131
    assume ?case1
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   132
    have "(x, h x) \<in> graph H'' h''" .
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   133
    also have "... \<subseteq> graph H' h'" .
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   134
    finally have xh:"(x, h x) \<in> graph H' h'" .
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   135
    then have "x \<in> H'" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   136
    moreover from y_hy have "y \<in> H'" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   137
    moreover from cM u and c' have "graph H' h' \<subseteq> graph H h"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   138
      by (simp only: chain_ball_Union_upper)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   139
    ultimately show ?thesis using * by blast
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   140
  next
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   141
    assume ?case2
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   142
    from x_hx have "x \<in> H''" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   143
    moreover {
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   144
      from y_hy have "(y, h y) \<in> graph H' h'" .
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   145
      also have "\<dots> \<subseteq> graph H'' h''" .
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   146
      finally have "(y, h y) \<in> graph H'' h''" .
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   147
    } then have "y \<in> H''" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   148
    moreover from cM u and c'' have "graph H'' h'' \<subseteq> graph H h"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   149
      by (simp only: chain_ball_Union_upper)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   150
    ultimately show ?thesis using ** by blast
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   151
  qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   152
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   153
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   154
text {*
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   155
  \medskip The relation induced by the graph of the supremum of a
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   156
  chain @{text c} is definite, i.~e.~t is the graph of a function.
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   157
*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   158
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   159
lemma sup_definite:
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   160
  assumes M_def: "M \<equiv> norm_pres_extensions E p F f"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   161
    and cM: "c \<in> chain M"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   162
    and xy: "(x, y) \<in> \<Union>c"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   163
    and xz: "(x, z) \<in> \<Union>c"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   164
  shows "z = y"
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   165
proof -
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   166
  from cM have c: "c \<subseteq> M" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   167
  from xy obtain G1 where xy': "(x, y) \<in> G1" and G1: "G1 \<in> c" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   168
  from xz obtain G2 where xz': "(x, z) \<in> G2" and G2: "G2 \<in> c" ..
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   169
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   170
  from G1 c have "G1 \<in> M" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   171
  then obtain H1 h1 where G1_rep: "G1 = graph H1 h1"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   172
    by (unfold M_def) blast
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   173
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   174
  from G2 c have "G2 \<in> M" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   175
  then obtain H2 h2 where G2_rep: "G2 = graph H2 h2"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   176
    by (unfold M_def) blast
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   177
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   178
  txt {* @{text "G\<^sub>1"} is contained in @{text "G\<^sub>2"}
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   179
    or vice versa, since both @{text "G\<^sub>1"} and @{text
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   180
    "G\<^sub>2"} are members of @{text c}. \label{cases2}*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   181
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   182
  from cM G1 G2 have "G1 \<subseteq> G2 \<or> G2 \<subseteq> G1" (is "?case1 \<or> ?case2") ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   183
  then show ?thesis
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   184
  proof
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   185
    assume ?case1
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   186
    with xy' G2_rep have "(x, y) \<in> graph H2 h2" by blast
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   187
    hence "y = h2 x" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   188
    also
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   189
    from xz' G2_rep have "(x, z) \<in> graph H2 h2" by (simp only:)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   190
    hence "z = h2 x" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   191
    finally show ?thesis .
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   192
  next
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   193
    assume ?case2
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   194
    with xz' G1_rep have "(x, z) \<in> graph H1 h1" by blast
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   195
    hence "z = h1 x" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   196
    also
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   197
    from xy' G1_rep have "(x, y) \<in> graph H1 h1" by (simp only:)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   198
    hence "y = h1 x" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   199
    finally show ?thesis ..
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   200
  qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   201
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   202
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   203
text {*
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   204
  \medskip The limit function @{text h} is linear. Every element
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   205
  @{text x} in the domain of @{text h} is in the domain of a function
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   206
  @{text h'} in the chain of norm preserving extensions.  Furthermore,
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   207
  @{text h} is an extension of @{text h'} so the function values of
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   208
  @{text x} are identical for @{text h'} and @{text h}.  Finally, the
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   209
  function @{text h'} is linear by construction of @{text M}.
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   210
*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   211
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   212
lemma sup_lf:
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   213
  assumes M: "M = norm_pres_extensions E p F f"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   214
    and cM: "c \<in> chain M"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   215
    and u: "graph H h = \<Union>c"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   216
  shows "linearform H h"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   217
proof
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   218
  fix x y assume x: "x \<in> H" and y: "y \<in> H"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   219
  with M cM u obtain H' h' where
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   220
        x': "x \<in> H'" and y': "y \<in> H'"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   221
      and b: "graph H' h' \<subseteq> graph H h"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   222
      and linearform: "linearform H' h'"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   223
      and subspace: "H' \<unlhd> E"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   224
    by (rule some_H'h'2 [elim_format]) blast
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   225
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   226
  show "h (x + y) = h x + h y"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   227
  proof -
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   228
    from linearform x' y' have "h' (x + y) = h' x + h' y"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   229
      by (rule linearform.add)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   230
    also from b x' have "h' x = h x" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   231
    also from b y' have "h' y = h y" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   232
    also from subspace x' y' have "x + y \<in> H'"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   233
      by (rule subspace.add_closed)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   234
    with b have "h' (x + y) = h (x + y)" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   235
    finally show ?thesis .
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   236
  qed
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   237
next
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   238
  fix x a assume x: "x \<in> H"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   239
  with M cM u obtain H' h' where
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   240
        x': "x \<in> H'"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   241
      and b: "graph H' h' \<subseteq> graph H h"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   242
      and linearform: "linearform H' h'"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   243
      and subspace: "H' \<unlhd> E"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   244
    by (rule some_H'h' [elim_format]) blast
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   245
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   246
  show "h (a \<cdot> x) = a * h x"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   247
  proof -
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   248
    from linearform x' have "h' (a \<cdot> x) = a * h' x"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   249
      by (rule linearform.mult)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   250
    also from b x' have "h' x = h x" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   251
    also from subspace x' have "a \<cdot> x \<in> H'"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   252
      by (rule subspace.mult_closed)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   253
    with b have "h' (a \<cdot> x) = h (a \<cdot> x)" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   254
    finally show ?thesis .
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   255
  qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   256
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   257
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   258
text {*
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   259
  \medskip The limit of a non-empty chain of norm preserving
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   260
  extensions of @{text f} is an extension of @{text f}, since every
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   261
  element of the chain is an extension of @{text f} and the supremum
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   262
  is an extension for every element of the chain.
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   263
*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   264
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   265
lemma sup_ext:
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   266
  assumes graph: "graph H h = \<Union>c"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   267
    and M: "M = norm_pres_extensions E p F f"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   268
    and cM: "c \<in> chain M"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   269
    and ex: "\<exists>x. x \<in> c"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   270
  shows "graph F f \<subseteq> graph H h"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   271
proof -
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   272
  from ex obtain x where xc: "x \<in> c" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   273
  from cM have "c \<subseteq> M" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   274
  with xc have "x \<in> M" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   275
  with M have "x \<in> norm_pres_extensions E p F f"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   276
    by (simp only:)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   277
  then obtain G g where "x = graph G g" and "graph F f \<subseteq> graph G g" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   278
  then have "graph F f \<subseteq> x" by (simp only:)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   279
  also from xc have "\<dots> \<subseteq> \<Union>c" by blast
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   280
  also from graph have "\<dots> = graph H h" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   281
  finally show ?thesis .
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   282
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   283
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   284
text {*
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   285
  \medskip The domain @{text H} of the limit function is a superspace
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   286
  of @{text F}, since @{text F} is a subset of @{text H}. The
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   287
  existence of the @{text 0} element in @{text F} and the closure
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   288
  properties follow from the fact that @{text F} is a vector space.
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   289
*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   290
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   291
lemma sup_supF:
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   292
  assumes graph: "graph H h = \<Union>c"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   293
    and M: "M = norm_pres_extensions E p F f"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   294
    and cM: "c \<in> chain M"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   295
    and ex: "\<exists>x. x \<in> c"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   296
    and FE: "F \<unlhd> E"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   297
  shows "F \<unlhd> H"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   298
proof
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   299
  from FE show "F \<noteq> {}" by (rule subspace.non_empty)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   300
  from graph M cM ex have "graph F f \<subseteq> graph H h" by (rule sup_ext)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   301
  then show "F \<subseteq> H" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   302
  fix x y assume "x \<in> F" and "y \<in> F"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   303
  with FE show "x + y \<in> F" by (rule subspace.add_closed)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   304
next
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   305
  fix x a assume "x \<in> F"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   306
  with FE show "a \<cdot> x \<in> F" by (rule subspace.mult_closed)
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   307
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   308
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   309
text {*
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   310
  \medskip The domain @{text H} of the limit function is a subspace of
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   311
  @{text E}.
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   312
*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   313
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   314
lemma sup_subE:
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   315
  assumes graph: "graph H h = \<Union>c"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   316
    and M: "M = norm_pres_extensions E p F f"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   317
    and cM: "c \<in> chain M"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   318
    and ex: "\<exists>x. x \<in> c"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   319
    and FE: "F \<unlhd> E"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   320
    and E: "vectorspace E"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   321
  shows "H \<unlhd> E"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   322
proof
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   323
  show "H \<noteq> {}"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   324
  proof -
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   325
    from FE E have "0 \<in> F" by (rule subvectorspace.zero)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   326
    also from graph M cM ex FE have "F \<unlhd> H" by (rule sup_supF)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   327
    then have "F \<subseteq> H" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   328
    finally show ?thesis by blast
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   329
  qed
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   330
  show "H \<subseteq> E"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   331
  proof
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   332
    fix x assume "x \<in> H"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   333
    with M cM graph
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   334
    obtain H' h' where x: "x \<in> H'" and H'E: "H' \<unlhd> E"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   335
      by (rule some_H'h' [elim_format]) blast
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   336
    from H'E have "H' \<subseteq> E" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   337
    with x show "x \<in> E" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   338
  qed
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   339
  fix x y assume x: "x \<in> H" and y: "y \<in> H"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   340
  show "x + y \<in> H"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   341
  proof -
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   342
    from M cM graph x y obtain H' h' where
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   343
          x': "x \<in> H'" and y': "y \<in> H'" and H'E: "H' \<unlhd> E"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   344
        and graphs: "graph H' h' \<subseteq> graph H h"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   345
      by (rule some_H'h'2 [elim_format]) blast
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   346
    from H'E x' y' have "x + y \<in> H'"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   347
      by (rule subspace.add_closed)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   348
    also from graphs have "H' \<subseteq> H" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   349
    finally show ?thesis .
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   350
  qed
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   351
next
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   352
  fix x a assume x: "x \<in> H"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   353
  show "a \<cdot> x \<in> H"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   354
  proof -
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   355
    from M cM graph x
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   356
    obtain H' h' where x': "x \<in> H'" and H'E: "H' \<unlhd> E"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   357
        and graphs: "graph H' h' \<subseteq> graph H h"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   358
      by (rule some_H'h' [elim_format]) blast
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   359
    from H'E x' have "a \<cdot> x \<in> H'" by (rule subspace.mult_closed)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   360
    also from graphs have "H' \<subseteq> H" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   361
    finally show ?thesis .
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   362
  qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   363
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   364
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   365
text {*
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   366
  \medskip The limit function is bounded by the norm @{text p} as
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   367
  well, since all elements in the chain are bounded by @{text p}.
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   368
*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   369
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   370
lemma sup_norm_pres:
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   371
  assumes graph: "graph H h = \<Union>c"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   372
    and M: "M = norm_pres_extensions E p F f"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   373
    and cM: "c \<in> chain M"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   374
  shows "\<forall>x \<in> H. h x \<le> p x"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   375
proof
9503
wenzelm
parents: 9394
diff changeset
   376
  fix x assume "x \<in> H"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   377
  with M cM graph obtain H' h' where x': "x \<in> H'"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   378
      and graphs: "graph H' h' \<subseteq> graph H h"
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   379
      and a: "\<forall>x \<in> H'. h' x \<le> p x"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   380
    by (rule some_H'h' [elim_format]) blast
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   381
  from graphs x' have [symmetric]: "h' x = h x" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   382
  also from a x' have "h' x \<le> p x " ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   383
  finally show "h x \<le> p x" .
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   384
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   385
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   386
text {*
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   387
  \medskip The following lemma is a property of linear forms on real
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   388
  vector spaces. It will be used for the lemma @{text abs_HahnBanach}
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   389
  (see page \pageref{abs-HahnBanach}). \label{abs-ineq-iff} For real
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   390
  vector spaces the following inequations are equivalent:
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   391
  \begin{center}
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   392
  \begin{tabular}{lll}
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   393
  @{text "\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x"} & and &
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   394
  @{text "\<forall>x \<in> H. h x \<le> p x"} \\
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   395
  \end{tabular}
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   396
  \end{center}
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   397
*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   398
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   399
lemma abs_ineq_iff:
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   400
  includes subvectorspace H E + seminorm E p + linearform H h
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   401
  shows "(\<forall>x \<in> H. \<bar>h x\<bar> \<le> p x) = (\<forall>x \<in> H. h x \<le> p x)" (is "?L = ?R")
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   402
proof
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   403
  have h: "vectorspace H" by (rule vectorspace)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   404
  {
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   405
    assume l: ?L
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   406
    show ?R
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   407
    proof
9503
wenzelm
parents: 9394
diff changeset
   408
      fix x assume x: "x \<in> H"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   409
      have "h x \<le> \<bar>h x\<bar>" by arith
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   410
      also from l x have "\<dots> \<le> p x" ..
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   411
      finally show "h x \<le> p x" .
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   412
    qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   413
  next
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   414
    assume r: ?R
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   415
    show ?L
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   416
    proof
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   417
      fix x assume x: "x \<in> H"
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   418
      show "\<And>a b :: real. - a \<le> b \<Longrightarrow> b \<le> a \<Longrightarrow> \<bar>b\<bar> \<le> a"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   419
        by arith
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   420
      show "- p x \<le> h x"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   421
      proof (rule real_minus_le)
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   422
        have "linearform H h" .
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   423
        from h this x have "- h x = h (- x)"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   424
          by (rule vectorspace_linearform.neg [symmetric])
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   425
        also from r x have "\<dots> \<le> p (- x)" by simp
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   426
        also have "\<dots> = p x"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   427
        proof (rule seminorm_vectorspace.minus)
9503
wenzelm
parents: 9394
diff changeset
   428
          show "x \<in> E" ..
9379
21cfeae6659d tuded presentation;
bauerg
parents: 9374
diff changeset
   429
        qed
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   430
        finally show "- h x \<le> p x" .
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   431
      qed
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   432
      from r x show "h x \<le> p x" ..
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   433
    qed
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
   434
  }
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   435
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   436
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9623
diff changeset
   437
end