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