src/HOL/Real/HahnBanach/HahnBanachSupLemmas.thy
author bauerg
Mon, 17 Jul 2000 13:58:18 +0200
changeset 9374 153853af318b
parent 9262 8baf94ddb345
child 9379 21cfeae6659d
permissions -rw-r--r--
- xsymbols for \<noteq> \<notin> \<in> \<exists> \<forall> \<and> \<inter> \<union> \<Union> - vector space type of {plus, minus, zero}, overload 0 in vector space - syntax |.| and ||.||
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
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    10
text{* This section contains some lemmas that will be used in the
7978
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
    11
proof of the Hahn-Banach Theorem.
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    12
In this section the following context is presumed. 
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
    13
Let $E$ be a real vector space with a seminorm $p$ on $E$. 
7978
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
    14
$F$ is a subspace of $E$ and $f$ a linear form on $F$. We 
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
    15
consider a chain $c$ of norm-preserving extensions of $f$, such that
7927
b50446a33c16 update by Gertrud Bauer;
wenzelm
parents: 7917
diff changeset
    16
$\Union c = \idt{graph}\ap H\ap h$. 
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    17
We will show some properties about the limit function $h$, 
7978
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
    18
i.e.\ the supremum of the chain $c$.
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    19
*} 
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    20
7978
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
    21
text{* Let $c$ be a chain of norm-preserving extensions of the 
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    22
function $f$ and let $\idt{graph}\ap H\ap h$ be the supremum of $c$. 
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    23
Every element in $H$ is member of
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    24
one of the elements of the chain. *}
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:
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    27
  "[| M = norm_pres_extensions E p F f; c \<in> chain M; 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    28
  graph H h = \<Union> c; x \<in> H |]
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    29
   ==> \<exists>H' h'. graph H' h' \<in> c \<and> (x, h x) \<in> graph H' h' 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    30
       \<and> is_linearform H' h' \<and> is_subspace H' E 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    31
       \<and> is_subspace F H' \<and> graph F f \<subseteq> graph H' h' 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    32
       \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    33
proof -
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    34
  assume m: "M = norm_pres_extensions E p F f" and "c \<in> chain M"
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    35
     and u: "graph H h = \<Union> c" "x \<in> H"
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    36
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    37
  have h: "(x, h x) \<in> graph H h" ..
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    38
  with u have "(x, h x) \<in> \<Union> c" by simp
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    39
  hence ex1: "\<exists>g \<in> c. (x, h x) \<in> g" 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    40
    by (simp only: Union_iff)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    41
  thus ?thesis
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    42
  proof (elim bexE)
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    43
    fix g assume g: "g \<in> c" "(x, h x) \<in> g"
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    44
    have "c \<subseteq> M" by (rule chainD2)
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    45
    hence "g \<in> M" ..
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    46
    hence "g \<in> norm_pres_extensions E p F f" by (simp only: m)
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    47
    hence "\<exists>H' h'. graph H' h' = g 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    48
                  \<and> is_linearform H' h'
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    49
                  \<and> is_subspace H' E
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    50
                  \<and> is_subspace F H'
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    51
                  \<and> graph F f \<subseteq> graph H' h'
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    52
                  \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    53
      by (rule norm_pres_extension_D)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    54
    thus ?thesis
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    55
    proof (elim exE conjE) 
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    56
      fix H' h' 
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    57
      assume "graph H' h' = g" "is_linearform H' h'" 
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    58
        "is_subspace H' E" "is_subspace F H'" 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    59
        "graph F f \<subseteq> graph H' h'" "\<forall>x \<in> H'. h' x \<le> p x"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    60
      show ?thesis 
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    61
      proof (intro exI conjI)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    62
        show "graph H' h' \<in> c" by (simp!)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    63
        show "(x, h x) \<in> graph H' h'" by (simp!)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    64
      qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    65
    qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    66
  qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    67
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    68
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    69
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
    70
text{*  \medskip Let $c$ be a chain of norm-preserving extensions of the
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    71
function $f$ and let $\idt{graph}\ap H\ap h$ be the supremum of $c$. 
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    72
Every element in the domain $H$ of the supremum function is member of
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    73
the domain $H'$ of some function $h'$, such that $h$ extends $h'$.
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    74
*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    75
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    76
lemma some_H'h': 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    77
  "[| M = norm_pres_extensions E p F f; c \<in> chain M; 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    78
  graph H h = \<Union> c; x \<in> H |] 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    79
  ==> \<exists>H' h'. x \<in> H' \<and> graph H' h' \<subseteq> graph H h 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    80
      \<and> is_linearform H' h' \<and> is_subspace H' E \<and> is_subspace F H'
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    81
      \<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
    82
proof -
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    83
  assume "M = norm_pres_extensions E p F f" and cM: "c \<in> chain M"
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    84
     and u: "graph H h = \<Union> c" "x \<in> H"  
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    85
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    86
  have "\<exists>H' h'. graph H' h' \<in> c \<and> (x, h x) \<in> graph H' h' 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    87
       \<and> is_linearform H' h' \<and> is_subspace H' E 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    88
       \<and> is_subspace F H' \<and> graph F f \<subseteq> graph H' h' 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    89
       \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    90
    by (rule some_H'h't)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    91
  thus ?thesis 
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    92
  proof (elim exE conjE)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    93
    fix H' h' assume "(x, h x) \<in> graph H' h'" "graph H' h' \<in> c"
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    94
      "is_linearform H' h'" "is_subspace H' E" "is_subspace F H'" 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    95
      "graph F f \<subseteq> graph H' h'" "\<forall>x \<in> H'. h' x \<le> p x"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    96
    show ?thesis
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
    97
    proof (intro exI conjI)
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    98
      show "x \<in> H'" by (rule graphD1)
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
    99
      from cM u show "graph H' h' \<subseteq> graph H h" 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   100
        by (simp! only: chain_ball_Union_upper)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   101
    qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   102
  qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   103
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   104
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   105
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
   106
text{* \medskip Any two elements $x$ and $y$ in the domain $H$ of the 
7978
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
   107
supremum function $h$ are both in the domain $H'$ of some function 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   108
$h'$, such that $h$ extends $h'$. *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   109
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   110
lemma some_H'h'2: 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   111
  "[| M = norm_pres_extensions E p F f; c \<in> chain M; 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   112
  graph H h = \<Union> c;  x \<in> H; y \<in> H |] 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   113
  ==> \<exists>H' h'. x \<in> H' \<and> y \<in> H' \<and> graph H' h' \<subseteq> graph H h 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   114
      \<and> is_linearform H' h' \<and> is_subspace H' E \<and> is_subspace F H'
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   115
      \<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
   116
proof -
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   117
  assume "M = norm_pres_extensions E p F f" "c \<in> chain M" 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   118
         "graph H h = \<Union> c" "x \<in> H" "y \<in> H"
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   119
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   120
  txt {* $x$ is in the domain $H'$ of some function $h'$, 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   121
  such that $h$ extends $h'$. *} 
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   122
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   123
  have e1: "\<exists>H' h'. graph H' h' \<in> c \<and> (x, h x) \<in> graph H' h'
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   124
       \<and> is_linearform H' h' \<and> is_subspace H' E 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   125
       \<and> is_subspace F H' \<and> graph F f \<subseteq> graph H' h' 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   126
       \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   127
    by (rule some_H'h't)
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   128
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   129
  txt {* $y$ is in the domain $H''$ of some function $h''$,
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   130
  such that $h$ extends $h''$. *} 
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   131
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   132
  have e2: "\<exists>H'' h''. graph H'' h'' \<in> c \<and> (y, h y) \<in> graph H'' h''
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   133
       \<and> is_linearform H'' h'' \<and> is_subspace H'' E 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   134
       \<and> is_subspace F H'' \<and> graph F f \<subseteq> graph H'' h'' 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   135
       \<and> (\<forall>x \<in> H''. h'' x \<le> p x)"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   136
    by (rule some_H'h't)
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   137
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   138
  from e1 e2 show ?thesis 
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   139
  proof (elim exE conjE)
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   140
    fix H' h' assume "(y, h y) \<in> graph H' h'" "graph H' h' \<in> c"
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   141
      "is_linearform H' h'" "is_subspace H' E" "is_subspace F H'" 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   142
      "graph F f \<subseteq> graph H' h'" "\<forall>x \<in> H'. h' x \<le> p x"
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   143
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   144
    fix H'' h'' assume "(x, h x) \<in> graph H'' h''" "graph H'' h'' \<in> c"
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   145
      "is_linearform H'' h''" "is_subspace H'' E" "is_subspace F H''"
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   146
      "graph F f \<subseteq> graph H'' h''" "\<forall>x \<in> H''. h'' x \<le> p x"
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   147
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   148
   txt {* Since both $h'$ and $h''$ are elements of the chain,  
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   149
   $h''$ is an extension of $h'$ or vice versa. Thus both 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   150
   $x$ and $y$ are contained in the greater one. \label{cases1}*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   151
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   152
    have "graph H'' h'' \<subseteq> graph H' h' | graph H' h' \<subseteq> graph H'' h''"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   153
      (is "?case1 | ?case2")
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   154
      by (rule chainD)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   155
    thus ?thesis
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   156
    proof 
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   157
      assume ?case1
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   158
      show ?thesis
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   159
      proof (intro exI conjI)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   160
        have "(x, h x) \<in> graph H'' h''" .
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   161
        also have "... \<subseteq> graph H' h'" .
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   162
        finally have xh:"(x, h x) \<in> graph H' h'" .
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   163
        thus x: "x \<in> H'" ..
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   164
        show y: "y \<in> H'" ..
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   165
        show "graph H' h' \<subseteq> graph H h"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   166
          by (simp! only: chain_ball_Union_upper)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   167
      qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   168
    next
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   169
      assume ?case2
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   170
      show ?thesis
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   171
      proof (intro exI conjI)
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   172
        show x: "x \<in> H''" ..
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   173
        have "(y, h y) \<in> graph H' h'" by (simp!)
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   174
        also have "... \<subseteq> graph H'' h''" .
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   175
        finally have yh: "(y, h y) \<in> graph H'' h''" .
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   176
        thus y: "y \<in> H''" ..
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   177
        show "graph H'' h'' \<subseteq> graph H h"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   178
          by (simp! only: chain_ball_Union_upper)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   179
      qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   180
    qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   181
  qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   182
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   183
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   184
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   185
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
   186
text{* \medskip The relation induced by the graph of the supremum
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   187
of a chain $c$ is definite, i.~e.~t is the graph of a function. *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   188
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   189
lemma sup_definite: 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   190
  "[| M == norm_pres_extensions E p F f; c \<in> chain M; 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   191
  (x, y) \<in> \<Union> c; (x, z) \<in> \<Union> c |] ==> z = y"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   192
proof - 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   193
  assume "c \<in> chain M" "M == norm_pres_extensions E p F f"
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   194
    "(x, y) \<in> \<Union> c" "(x, z) \<in> \<Union> c"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   195
  thus ?thesis
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   196
  proof (elim UnionE chainE2)
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   197
7927
b50446a33c16 update by Gertrud Bauer;
wenzelm
parents: 7917
diff changeset
   198
    txt{* Since both $(x, y) \in \Union c$ and $(x, z) \in \Union c$
b50446a33c16 update by Gertrud Bauer;
wenzelm
parents: 7917
diff changeset
   199
    they are members of some graphs $G_1$ and $G_2$, resp., such that
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   200
    both $G_1$ and $G_2$ are members of $c$.*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   201
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   202
    fix G1 G2 assume
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   203
      "(x, y) \<in> G1" "G1 \<in> c" "(x, z) \<in> G2" "G2 \<in> c" "c \<subseteq> M"
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   204
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   205
    have "G1 \<in> M" ..
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   206
    hence e1: "\<exists> H1 h1. graph H1 h1 = G1"  
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   207
      by (force! dest: norm_pres_extension_D)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   208
    have "G2 \<in> M" ..
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   209
    hence e2: "\<exists> H2 h2. graph H2 h2 = G2"  
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   210
      by (force! dest: norm_pres_extension_D)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   211
    from e1 e2 show ?thesis 
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   212
    proof (elim exE)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   213
      fix H1 h1 H2 h2 
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   214
      assume "graph H1 h1 = G1" "graph H2 h2 = G2"
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   215
7978
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
   216
      txt{* $G_1$ is contained in $G_2$ or vice versa, 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   217
      since both $G_1$ and $G_2$ are members of $c$. \label{cases2}*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   218
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   219
      have "G1 \<subseteq> G2 | G2 \<subseteq> G1" (is "?case1 | ?case2") ..
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   220
      thus ?thesis
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   221
      proof
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   222
        assume ?case1
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   223
        have "(x, y) \<in> graph H2 h2" by (force!)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   224
        hence "y = h2 x" ..
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   225
        also have "(x, z) \<in> graph H2 h2" by (simp!)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   226
        hence "z = h2 x" ..
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   227
        finally show ?thesis .
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   228
      next
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   229
        assume ?case2
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   230
        have "(x, y) \<in> graph H1 h1" by (simp!)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   231
        hence "y = h1 x" ..
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   232
        also have "(x, z) \<in> graph H1 h1" by (force!)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   233
        hence "z = h1 x" ..
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   234
        finally show ?thesis .
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   235
      qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   236
    qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   237
  qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   238
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   239
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
   240
text{* \medskip The limit function $h$ is linear. Every element $x$ in the
7927
b50446a33c16 update by Gertrud Bauer;
wenzelm
parents: 7917
diff changeset
   241
domain of $h$ is in the domain of a function $h'$ in the chain of norm
b50446a33c16 update by Gertrud Bauer;
wenzelm
parents: 7917
diff changeset
   242
preserving extensions.  Furthermore, $h$ is an extension of $h'$ so
7978
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
   243
the function values of $x$ are identical for $h'$ and $h$.  Finally, the
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   244
function $h'$ is linear by construction of $M$.  *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   245
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   246
lemma sup_lf: 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   247
  "[| M = norm_pres_extensions E p F f; c \<in> chain M; 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   248
  graph H h = \<Union> c |] ==> is_linearform H h"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   249
proof - 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   250
  assume "M = norm_pres_extensions E p F f" "c \<in> chain M"
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   251
         "graph H h = \<Union> c"
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   252
 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   253
  show "is_linearform H h"
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   254
  proof
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   255
    fix x y assume "x \<in> H" "y \<in> H" 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   256
    have "\<exists>H' h'. x \<in> H' \<and> y \<in> H' \<and> graph H' h' \<subseteq> graph H h 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   257
            \<and> is_linearform H' h' \<and> is_subspace H' E 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   258
            \<and> is_subspace F H' \<and> graph F f \<subseteq> graph H' h'
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   259
            \<and> (\<forall>x \<in> H'. h' x \<le> p x)"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   260
      by (rule some_H'h'2)
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   261
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   262
    txt {* We have to show that $h$ is additive. *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   263
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   264
    thus "h (x + y) = h x + h y" 
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   265
    proof (elim exE conjE)
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   266
      fix H' h' assume "x \<in> H'" "y \<in> H'" 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   267
        and b: "graph H' h' \<subseteq> graph H h" 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   268
        and "is_linearform H' h'" "is_subspace H' E"
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   269
      have "h' (x + y) = h' x + h' y" 
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   270
        by (rule linearform_add)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   271
      also have "h' x = h x" ..
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   272
      also have "h' y = h y" ..
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   273
      also have "x + y \<in> H'" ..
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   274
      with b have "h' (x + y) = h (x + y)" ..
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   275
      finally show ?thesis .
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   276
    qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   277
  next  
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   278
    fix a x assume "x \<in> H"
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   279
    have "\<exists> H' h'. x \<in> H' \<and> graph H' h' \<subseteq> graph H h 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   280
            \<and> is_linearform H' h' \<and> is_subspace H' E
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   281
            \<and> is_subspace F H' \<and> graph F f \<subseteq> graph H' h' 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   282
            \<and> (\<forall> x \<in> H'. h' x \<le> p x)"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   283
      by (rule some_H'h')
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   284
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   285
    txt{* We have to show that $h$ is multiplicative. *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   286
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   287
    thus "h (a \<cdot> x) = a * h x"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   288
    proof (elim exE conjE)
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   289
      fix H' h' assume "x \<in> H'"
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   290
        and b: "graph H' h' \<subseteq> graph H h" 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   291
        and "is_linearform H' h'" "is_subspace H' E"
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   292
      have "h' (a \<cdot> x) = a * h' x" 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   293
        by (rule linearform_mult)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   294
      also have "h' x = h x" ..
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   295
      also have "a \<cdot> x \<in> H'" ..
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   296
      with b have "h' (a \<cdot> x) = h (a \<cdot> x)" ..
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   297
      finally show ?thesis .
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   298
    qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   299
  qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   300
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   301
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
   302
text{* \medskip The limit of a non-empty chain of norm
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   303
preserving extensions of $f$ is an extension of $f$,
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   304
since every element of the chain is an extension
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   305
of $f$ and the supremum is an extension
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   306
for every element of the chain.*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   307
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   308
lemma sup_ext:
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   309
  "[| graph H h = \<Union> c; M = norm_pres_extensions E p F f; 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   310
  c \<in> chain M; \<exists>x. x \<in> c |] ==> graph F f \<subseteq> graph H h"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   311
proof -
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   312
  assume "M = norm_pres_extensions E p F f" "c \<in> chain M" 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   313
         "graph H h = \<Union> c"
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   314
  assume "\<exists>x. x \<in> c"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   315
  thus ?thesis 
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   316
  proof
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   317
    fix x assume "x \<in> c" 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   318
    have "c \<subseteq> M" by (rule chainD2)
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   319
    hence "x \<in> M" ..
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   320
    hence "x \<in> norm_pres_extensions E p F f" by (simp!)
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   321
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   322
    hence "\<exists>G g. graph G g = x
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   323
             \<and> is_linearform G g 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   324
             \<and> is_subspace G E
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   325
             \<and> is_subspace F G
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   326
             \<and> graph F f \<subseteq> graph G g 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   327
             \<and> (\<forall>x \<in> G. g x \<le> p x)"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   328
      by (simp! add: norm_pres_extension_D)
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   329
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   330
    thus ?thesis 
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   331
    proof (elim exE conjE) 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   332
      fix G g assume "graph F f \<subseteq> graph G g"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   333
      also assume "graph G g = x"
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   334
      also have "... \<in> c" .
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   335
      hence "x \<subseteq> \<Union> c" by fast
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   336
      also have [RS sym]: "graph H h = \<Union> c" .
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   337
      finally show ?thesis .
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   338
    qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   339
  qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   340
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   341
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
   342
text{* \medskip The domain $H$ of the limit function is a superspace of $F$,
7927
b50446a33c16 update by Gertrud Bauer;
wenzelm
parents: 7917
diff changeset
   343
since $F$ is a subset of $H$. The existence of the $\zero$ element in
b50446a33c16 update by Gertrud Bauer;
wenzelm
parents: 7917
diff changeset
   344
$F$ and the closure properties follow from the fact that $F$ is a
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   345
vector space. *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   346
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   347
lemma sup_supF: 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   348
  "[|  graph H h = \<Union> c; M = norm_pres_extensions E p F f; 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   349
  c \<in> chain M; \<exists>x. x \<in> c; is_subspace F E; is_vectorspace E |] 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   350
  ==> is_subspace F H"
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   351
proof - 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   352
  assume "M = norm_pres_extensions E p F f" "c \<in> chain M" "\<exists>x. x \<in> c"
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   353
    "graph H h = \<Union> c" "is_subspace F E" "is_vectorspace E"
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   354
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   355
  show ?thesis 
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   356
  proof
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   357
    show "0 \<in> F" ..
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   358
    show "F \<subseteq> H" 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   359
    proof (rule graph_extD2)
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   360
      show "graph F f \<subseteq> graph H h"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   361
        by (rule sup_ext)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   362
    qed
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   363
    show "\<forall>x \<in> F. \<forall>y \<in> F. x + y \<in> F" 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   364
    proof (intro ballI) 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   365
      fix x y assume "x \<in> F" "y \<in> F"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   366
      show "x + y \<in> F" by (simp!)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   367
    qed
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   368
    show "\<forall>x \<in> F. \<forall>a. a \<cdot> x \<in> F"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   369
    proof (intro ballI allI)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   370
      fix x a assume "x\<in>F"
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   371
      show "a \<cdot> x \<in> F" by (simp!)
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   372
    qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   373
  qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   374
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   375
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
   376
text{* \medskip The domain $H$ of the limit function is a subspace
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   377
of $E$. *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   378
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   379
lemma sup_subE: 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   380
  "[| graph H h = \<Union> c; M = norm_pres_extensions E p F f; 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   381
  c \<in> chain M; \<exists>x. x \<in> c; is_subspace F E; is_vectorspace E |] 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   382
  ==> is_subspace H E"
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   383
proof - 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   384
  assume "M = norm_pres_extensions E p F f" "c \<in> chain M" "\<exists>x. x \<in> c"
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   385
    "graph H h = \<Union> c" "is_subspace F E" "is_vectorspace E"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   386
  show ?thesis 
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   387
  proof
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   388
 
7978
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
   389
    txt {* The $\zero$ element is in $H$, as $F$ is a subset 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   390
    of $H$: *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   391
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   392
    have "0 \<in> F" ..
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   393
    also have "is_subspace F H" by (rule sup_supF) 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   394
    hence "F \<subseteq> H" ..
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   395
    finally show "0 \<in> H" .
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   396
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   397
    txt{* $H$ is a subset of $E$: *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   398
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   399
    show "H \<subseteq> E" 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   400
    proof
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   401
      fix x assume "x \<in> H"
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   402
      have "\<exists>H' h'. x \<in> H' \<and> graph H' h' \<subseteq> graph H h
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   403
              \<and> is_linearform H' h' \<and> is_subspace H' E 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   404
              \<and> is_subspace F H' \<and> graph F f \<subseteq> graph H' h' 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   405
              \<and> (\<forall>x \<in> H'. h' x \<le> p x)" 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   406
	by (rule some_H'h')
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   407
      thus "x \<in> E" 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   408
      proof (elim exE conjE)
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   409
	fix H' h' assume "x \<in> H'" "is_subspace H' E"
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   410
        have "H' \<subseteq> E" ..
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   411
	thus "x \<in> E" ..
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   412
      qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   413
    qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   414
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   415
    txt{* $H$ is closed under addition: *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   416
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   417
    show "\<forall>x \<in> H. \<forall>y \<in> H. x + y \<in> H" 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   418
    proof (intro ballI) 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   419
      fix x y assume "x \<in> H" "y \<in> H"
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   420
      have "\<exists>H' h'. x \<in> H' \<and> y \<in> H' \<and> graph H' h' \<subseteq> graph H h 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   421
              \<and> is_linearform H' h' \<and> is_subspace H' E 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   422
              \<and> is_subspace F H' \<and> graph F f \<subseteq> graph H' h' 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   423
              \<and> (\<forall>x \<in> H'. h' x \<le> p x)" 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   424
	by (rule some_H'h'2) 
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   425
      thus "x + y \<in> H" 
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   426
      proof (elim exE conjE) 
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   427
	fix H' h' 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   428
        assume "x \<in> H'" "y \<in> H'" "is_subspace H' E" 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   429
          "graph H' h' \<subseteq> graph H h"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   430
        have "x + y \<in> H'" ..
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   431
	also have "H' \<subseteq> H" ..
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   432
	finally show ?thesis .
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   433
      qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   434
    qed      
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   435
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   436
    txt{* $H$ is closed under scalar multiplication: *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   437
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   438
    show "\<forall>x \<in> H. \<forall>a. a \<cdot> x \<in> H" 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   439
    proof (intro ballI allI) 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   440
      fix x a assume "x \<in> H" 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   441
      have "\<exists>H' h'. x \<in> H' \<and> graph H' h' \<subseteq> graph H h
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   442
              \<and> is_linearform H' h' \<and> is_subspace H' E 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   443
              \<and> is_subspace F H' \<and> graph F f \<subseteq> graph H' h' 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   444
              \<and> (\<forall>x \<in> H'. h' x \<le> p x)" 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   445
	by (rule some_H'h') 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   446
      thus "a \<cdot> x \<in> H" 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   447
      proof (elim exE conjE)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   448
	fix H' h' 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   449
        assume "x \<in> H'" "is_subspace H' E" "graph H' h' \<subseteq> graph H h"
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   450
        have "a \<cdot> x \<in> H'" ..
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   451
        also have "H' \<subseteq> H" ..
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   452
	finally show ?thesis .
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   453
      qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   454
    qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   455
  qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   456
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   457
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
   458
text {* \medskip The limit function is bounded by 
7978
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
   459
the norm $p$ as well, since all elements in the chain are
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
   460
bounded by $p$.
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   461
*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   462
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   463
lemma sup_norm_pres:
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   464
  "[| graph H h = \<Union> c; M = norm_pres_extensions E p F f; c \<in> chain M |] 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   465
  ==> \<forall> x \<in> H. h x \<le> p x"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   466
proof
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   467
  assume "M = norm_pres_extensions E p F f" "c \<in> chain M" 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   468
         "graph H h = \<Union> c"
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   469
  fix x assume "x \<in> H"
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   470
  have "\<exists>H' h'. x \<in> H' \<and> graph H' h' \<subseteq> graph H h 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   471
          \<and> is_linearform H' h' \<and> is_subspace H' E \<and> is_subspace F H'
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   472
          \<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
   473
    by (rule some_H'h')
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   474
  thus "h x \<le> p x" 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   475
  proof (elim exE conjE)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   476
    fix H' h' 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   477
    assume "x \<in> H'" "graph H' h' \<subseteq> graph H h" 
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   478
      and a: "\<forall>x \<in> H'. h' x \<le> p x"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   479
    have [RS sym]: "h' x = h x" ..
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   480
    also from a have "h' x \<le> p x " ..
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   481
    finally show ?thesis .  
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   482
  qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   483
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   484
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   485
8084
c3790c6b4470 small changes;
bauerg
parents: 7978
diff changeset
   486
text{* \medskip The following lemma is a property of linear forms on 
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   487
real vector spaces. It will be used for the lemma 
8838
4eaa99f0d223 replaced rabs by overloaded abs;
wenzelm
parents: 8703
diff changeset
   488
$\idt{abs{\dsh}HahnBanach}$ (see page \pageref{abs-HahnBanach}). \label{abs-ineq-iff}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   489
For real vector spaces the following inequations are equivalent:
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   490
\begin{matharray}{ll} 
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   491
\forall x\in H.\ap |h\ap x|\leq p\ap x& {\rm and}\\ 
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   492
\forall x\in H.\ap h\ap x\leq p\ap x\\ 
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   493
\end{matharray}
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   494
*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   495
8838
4eaa99f0d223 replaced rabs by overloaded abs;
wenzelm
parents: 8703
diff changeset
   496
lemma abs_ineq_iff: 
7978
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
   497
  "[| is_subspace H E; is_vectorspace E; is_seminorm E p; 
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   498
  is_linearform H h |] 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   499
  ==> (\<forall>x \<in> H. |h x| \<le> p x) = (\<forall>x \<in> H. h x \<le> p x)" 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   500
  (concl is "?L = ?R")
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   501
proof -
7978
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
   502
  assume "is_subspace H E" "is_vectorspace E" "is_seminorm E p" 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   503
         "is_linearform H h"
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   504
  have h: "is_vectorspace H" ..
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   505
  show ?thesis
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   506
  proof 
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   507
    assume l: ?L
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   508
    show ?R
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   509
    proof
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   510
      fix x assume x: "x \<in> H"
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   511
      have "h x \<le> |h x|" by (rule abs_ge_self)
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   512
      also from l have "... \<le> p x" ..
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   513
      finally show "h x \<le> p x" .
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   514
    qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   515
  next
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   516
    assume r: ?R
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   517
    show ?L
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   518
    proof 
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   519
      fix x assume "x \<in> H"
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   520
      show "!! a b :: real. [| - a \<le> b; b \<le> a |] ==> |b| \<le> a"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   521
        by arith
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   522
      show "- p x \<le> h x"
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   523
      proof (rule real_minus_le)
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   524
	from h have "- h x = h (- x)" 
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   525
          by (rule linearform_neg [RS sym])
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   526
	also from r have "... \<le> p (- x)" by (simp!)
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   527
	also have "... = p x" 
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   528
          by (rule seminorm_minus [OF _ subspace_subsetD])
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   529
	finally show "- h x \<le> p x" . 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   530
      qed
9374
153853af318b - xsymbols for
bauerg
parents: 9262
diff changeset
   531
      from r show "h x \<le> p x" .. 
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   532
    qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   533
  qed
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   534
qed  
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   535
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   536
9261
879e0f0cd047 removed sorry;
bauerg
parents: 9256
diff changeset
   537
end