src/HOL/Hahn_Banach/Hahn_Banach_Ext_Lemmas.thy
author wenzelm
Fri, 17 Jan 2025 15:39:40 +0100
changeset 81856 4af2e864c26c
parent 81464 2575f1bd226b
permissions -rw-r--r--
clarified inst_type: more direct Thm.instantiate_frees;
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_Ext_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>Extending non-maximal functions\<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_Ext_Lemmas
be3e1cc5005c standard naming conventions for session and theories;
wenzelm
parents: 29252
diff changeset
     8
imports Function_Norm
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
     9
begin
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    10
58744
c434e37f290e update_cartouches;
wenzelm
parents: 49962
diff changeset
    11
text \<open>
61879
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61540
diff changeset
    12
  In this section the following context is presumed. Let \<open>E\<close> be a real vector
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61540
diff changeset
    13
  space with a seminorm \<open>q\<close> on \<open>E\<close>. \<open>F\<close> is a subspace of \<open>E\<close> and \<open>f\<close> a linear
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61540
diff changeset
    14
  function on \<open>F\<close>. We consider a subspace \<open>H\<close> of \<open>E\<close> that is a superspace of
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61540
diff changeset
    15
  \<open>F\<close> and a linear form \<open>h\<close> on \<open>H\<close>. \<open>H\<close> is a not equal to \<open>E\<close> and \<open>x\<^sub>0\<close> is an
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61540
diff changeset
    16
  element in \<open>E - H\<close>. \<open>H\<close> is extended to the direct sum \<open>H' = H + lin x\<^sub>0\<close>, so
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61540
diff changeset
    17
  for any \<open>x \<in> H'\<close> the decomposition of \<open>x = y + a \<cdot> x\<close> with \<open>y \<in> H\<close> is
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61540
diff changeset
    18
  unique. \<open>h'\<close> is defined on \<open>H'\<close> by \<open>h' x = h y + a \<cdot> \<xi>\<close> for a certain \<open>\<xi>\<close>.
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    19
61540
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    20
  Subsequently we show some properties of this extension \<open>h'\<close> of \<open>h\<close>.
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    21
61486
3590367b0ce9 tuned document;
wenzelm
parents: 60458
diff changeset
    22
  \<^medskip>
61540
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    23
  This lemma will be used to show the existence of a linear extension of \<open>f\<close>
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    24
  (see page \pageref{ex-xi-use}). It is a consequence of the completeness of
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    25
  \<open>\<real>\<close>. To show
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
    26
  \begin{center}
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
    27
  \begin{tabular}{l}
61539
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
    28
  \<open>\<exists>\<xi>. \<forall>y \<in> F. a y \<le> \<xi> \<and> \<xi> \<le> b y\<close>
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
    29
  \end{tabular}
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
    30
  \end{center}
61540
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
    31
  \<^noindent> it suffices to show that
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
    32
  \begin{center}
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
    33
  \begin{tabular}{l}
61539
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
    34
  \<open>\<forall>u \<in> F. \<forall>v \<in> F. a u \<le> b v\<close>
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
    35
  \end{tabular}
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
    36
  \end{center}
58744
c434e37f290e update_cartouches;
wenzelm
parents: 49962
diff changeset
    37
\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    38
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
    39
lemma ex_xi:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
    40
  assumes "vectorspace F"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    41
  assumes r: "\<And>u v. u \<in> F \<Longrightarrow> v \<in> F \<Longrightarrow> a u \<le> b v"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    42
  shows "\<exists>xi::real. \<forall>y \<in> F. a y \<le> xi \<and> xi \<le> b y"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9906
diff changeset
    43
proof -
29234
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
    44
  interpret vectorspace F by fact
58744
c434e37f290e update_cartouches;
wenzelm
parents: 49962
diff changeset
    45
  txt \<open>From the completeness of the reals follows:
61539
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
    46
    The set \<open>S = {a u. u \<in> F}\<close> has a supremum, if it is
58744
c434e37f290e update_cartouches;
wenzelm
parents: 49962
diff changeset
    47
    non-empty and has an upper bound.\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    48
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    49
  let ?S = "{a u | u. u \<in> F}"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    50
  have "\<exists>xi. lub ?S xi"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    51
  proof (rule real_complete)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    52
    have "a 0 \<in> ?S" by blast
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    53
    then show "\<exists>X. X \<in> ?S" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    54
    have "\<forall>y \<in> ?S. y \<le> b 0"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    55
    proof
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    56
      fix y assume y: "y \<in> ?S"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    57
      then obtain u where u: "u \<in> F" and y: "y = a u" by blast
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    58
      from u and zero have "a u \<le> b 0" by (rule r)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    59
      with y show "y \<le> b 0" by (simp only:)
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9906
diff changeset
    60
    qed
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    61
    then show "\<exists>u. \<forall>y \<in> ?S. y \<le> u" ..
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9906
diff changeset
    62
  qed
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    63
  then obtain xi where xi: "lub ?S xi" ..
81464
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
    64
  have "a y \<le> xi" if "y \<in> F" for y
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
    65
  proof -
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
    66
    from that have "a y \<in> ?S" by blast
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
    67
    with xi show ?thesis by (rule lub.upper)
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
    68
  qed
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
    69
  moreover have "xi \<le> b y" if y: "y \<in> F" for y
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
    70
  proof -
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
    71
    from xi
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
    72
    show ?thesis
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    73
    proof (rule lub.least)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    74
      fix au assume "au \<in> ?S"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    75
      then obtain u where u: "u \<in> F" and au: "au = a u" by blast
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    76
      from u y have "a u \<le> b y" by (rule r)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    77
      with au show "au \<le> b y" by (simp only:)
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9906
diff changeset
    78
    qed
81464
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
    79
  qed
60458
0d10ae17e3ee tuned proofs;
wenzelm
parents: 58999
diff changeset
    80
  ultimately show "\<exists>xi. \<forall>y \<in> F. a y \<le> xi \<and> xi \<le> b y" by blast
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9906
diff changeset
    81
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    82
58744
c434e37f290e update_cartouches;
wenzelm
parents: 49962
diff changeset
    83
text \<open>
61486
3590367b0ce9 tuned document;
wenzelm
parents: 60458
diff changeset
    84
  \<^medskip>
61879
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61540
diff changeset
    85
  The function \<open>h'\<close> is defined as a \<open>h' x = h y + a \<cdot> \<xi>\<close> where \<open>x = y + a \<cdot> \<xi>\<close>
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61540
diff changeset
    86
  is a linear extension of \<open>h\<close> to \<open>H'\<close>.
58744
c434e37f290e update_cartouches;
wenzelm
parents: 49962
diff changeset
    87
\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    88
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 10606
diff changeset
    89
lemma h'_lf:
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 61879
diff changeset
    90
  assumes h'_def: "\<And>x. h' x = (let (y, a) =
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 61879
diff changeset
    91
      SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi)"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 61879
diff changeset
    92
    and H'_def: "H' = H + lin x0"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    93
    and HE: "H \<unlhd> E"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
    94
  assumes "linearform H h"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    95
  assumes x0: "x0 \<notin> H"  "x0 \<in> E"  "x0 \<noteq> 0"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
    96
  assumes E: "vectorspace E"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
    97
  shows "linearform H' h'"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
    98
proof -
29234
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
    99
  interpret linearform H h by fact
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
   100
  interpret vectorspace E by fact
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   101
  show ?thesis
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   102
  proof
58744
c434e37f290e update_cartouches;
wenzelm
parents: 49962
diff changeset
   103
    note E = \<open>vectorspace E\<close>
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   104
    have H': "vectorspace H'"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   105
    proof (unfold H'_def)
58744
c434e37f290e update_cartouches;
wenzelm
parents: 49962
diff changeset
   106
      from \<open>x0 \<in> E\<close>
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   107
      have "lin x0 \<unlhd> E" ..
47445
69e96e5500df Set_Algebras: removed syntax \<oplus> and \<otimes>, in favour of plain + and *
krauss
parents: 44190
diff changeset
   108
      with HE show "vectorspace (H + lin x0)" using E ..
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   109
    qed
81464
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   110
    show "h' (x1 + x2) = h' x1 + h' x2" if x1: "x1 \<in> H'" and x2: "x2 \<in> H'" for x1 x2
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   111
    proof -
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   112
      from H' x1 x2 have "x1 + x2 \<in> H'"
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   113
        by (rule vectorspace.add_closed)
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   114
      with x1 x2 obtain y y1 y2 a a1 a2 where
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   115
        x1x2: "x1 + x2 = y + a \<cdot> x0" and y: "y \<in> H"
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   116
        and x1_rep: "x1 = y1 + a1 \<cdot> x0" and y1: "y1 \<in> H"
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   117
        and x2_rep: "x2 = y2 + a2 \<cdot> x0" and y2: "y2 \<in> H"
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   118
        unfolding H'_def sum_def lin_def by blast
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   119
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   120
      have ya: "y1 + y2 = y \<and> a1 + a2 = a" using E HE _ y x0
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   121
      proof (rule decomp_H') text_raw \<open>\label{decomp-H-use}\<close>
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   122
        from HE y1 y2 show "y1 + y2 \<in> H"
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   123
          by (rule subspace.add_closed)
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   124
        from x0 and HE y y1 y2
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   125
        have "x0 \<in> E"  "y \<in> E"  "y1 \<in> E"  "y2 \<in> E" by auto
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   126
        with x1_rep x2_rep have "(y1 + y2) + (a1 + a2) \<cdot> x0 = x1 + x2"
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   127
          by (simp add: add_ac add_mult_distrib2)
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   128
        also note x1x2
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   129
        finally show "(y1 + y2) + (a1 + a2) \<cdot> x0 = y + a \<cdot> x0" .
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9906
diff changeset
   130
      qed
81464
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   131
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   132
      from h'_def x1x2 E HE y x0
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   133
      have "h' (x1 + x2) = h y + a * xi"
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   134
        by (rule h'_definite)
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   135
      also have "\<dots> = h (y1 + y2) + (a1 + a2) * xi"
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   136
        by (simp only: ya)
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   137
      also from y1 y2 have "h (y1 + y2) = h y1 + h y2"
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   138
        by simp
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   139
      also have "\<dots> + (a1 + a2) * xi = (h y1 + a1 * xi) + (h y2 + a2 * xi)"
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   140
        by (simp add: distrib_right)
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   141
      also from h'_def x1_rep E HE y1 x0
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   142
      have "h y1 + a1 * xi = h' x1"
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   143
        by (rule h'_definite [symmetric])
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   144
      also from h'_def x2_rep E HE y2 x0
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   145
      have "h y2 + a2 * xi = h' x2"
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   146
        by (rule h'_definite [symmetric])
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   147
      finally show ?thesis .
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   148
    qed
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   149
    show "h' (c \<cdot> x1) = c * (h' x1)" if x1: "x1 \<in> H'" for x1 c
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   150
    proof -
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   151
      from H' x1 have ax1: "c \<cdot> x1 \<in> H'"
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   152
        by (rule vectorspace.mult_closed)
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   153
      with x1 obtain y a y1 a1 where
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   154
          cx1_rep: "c \<cdot> x1 = y + a \<cdot> x0" and y: "y \<in> H"
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   155
        and x1_rep: "x1 = y1 + a1 \<cdot> x0" and y1: "y1 \<in> H"
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   156
        unfolding H'_def sum_def lin_def by blast
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   157
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   158
      have ya: "c \<cdot> y1 = y \<and> c * a1 = a" using E HE _ y x0
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   159
      proof (rule decomp_H')
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   160
        from HE y1 show "c \<cdot> y1 \<in> H"
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   161
          by (rule subspace.mult_closed)
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   162
        from x0 and HE y y1
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   163
        have "x0 \<in> E"  "y \<in> E"  "y1 \<in> E" by auto
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   164
        with x1_rep have "c \<cdot> y1 + (c * a1) \<cdot> x0 = c \<cdot> x1"
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   165
          by (simp add: mult_assoc add_mult_distrib1)
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   166
        also note cx1_rep
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   167
        finally show "c \<cdot> y1 + (c * a1) \<cdot> x0 = y + a \<cdot> x0" .
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9906
diff changeset
   168
      qed
81464
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   169
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   170
      from h'_def cx1_rep E HE y x0 have "h' (c \<cdot> x1) = h y + a * xi"
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   171
        by (rule h'_definite)
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   172
      also have "\<dots> = h (c \<cdot> y1) + (c * a1) * xi"
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   173
        by (simp only: ya)
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   174
      also from y1 have "h (c \<cdot> y1) = c * h y1"
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   175
        by simp
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   176
      also have "\<dots> + (c * a1) * xi = c * (h y1 + a1 * xi)"
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   177
        by (simp only: distrib_left)
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   178
      also from h'_def x1_rep E HE y1 x0 have "h y1 + a1 * xi = h' x1"
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   179
        by (rule h'_definite [symmetric])
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   180
      finally show ?thesis .
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   181
    qed
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   182
  qed
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9906
diff changeset
   183
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   184
61486
3590367b0ce9 tuned document;
wenzelm
parents: 60458
diff changeset
   185
text \<open>
3590367b0ce9 tuned document;
wenzelm
parents: 60458
diff changeset
   186
  \<^medskip>
61540
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
   187
  The linear extension \<open>h'\<close> of \<open>h\<close> is bounded by the seminorm \<open>p\<close>.
f92bf6674699 tuned document;
wenzelm
parents: 61539
diff changeset
   188
\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   189
9374
153853af318b - xsymbols for
bauerg
parents: 9256
diff changeset
   190
lemma h'_norm_pres:
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 61879
diff changeset
   191
  assumes h'_def: "\<And>x. h' x = (let (y, a) =
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 61879
diff changeset
   192
      SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi)"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 61879
diff changeset
   193
    and H'_def: "H' = H + lin x0"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   194
    and x0: "x0 \<notin> H"  "x0 \<in> E"  "x0 \<noteq> 0"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   195
  assumes E: "vectorspace E" and HE: "subspace H E"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   196
    and "seminorm E p" and "linearform H h"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   197
  assumes a: "\<forall>y \<in> H. h y \<le> p y"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   198
    and a': "\<forall>y \<in> H. - p (y + x0) - h y \<le> xi \<and> xi \<le> p (y + x0) - h y"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   199
  shows "\<forall>x \<in> H'. h' x \<le> p x"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   200
proof -
29234
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
   201
  interpret vectorspace E by fact
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
   202
  interpret subspace H E by fact
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
   203
  interpret seminorm E p by fact
60f7fb56f8cd More porting to new locales.
ballarin
parents: 27612
diff changeset
   204
  interpret linearform H h by fact
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   205
  show ?thesis
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   206
  proof
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   207
    fix x assume x': "x \<in> H'"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   208
    show "h' x \<le> p x"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   209
    proof -
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   210
      from a' have a1: "\<forall>ya \<in> H. - p (ya + x0) - h ya \<le> xi"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   211
        and a2: "\<forall>ya \<in> H. xi \<le> p (ya + x0) - h ya" by auto
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   212
      from x' obtain y a where
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
   213
          x_rep: "x = y + a \<cdot> x0" and y: "y \<in> H"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   214
        unfolding H'_def sum_def lin_def by blast
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   215
      from y have y': "y \<in> E" ..
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   216
      from y have ay: "inverse a \<cdot> y \<in> H" by simp
81464
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   217
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   218
      from h'_def x_rep E HE y x0 have "h' x = h y + a * xi"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   219
        by (rule h'_definite)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   220
      also have "\<dots> \<le> p (y + a \<cdot> x0)"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   221
      proof (rule linorder_cases)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   222
        assume z: "a = 0"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   223
        then have "h y + a * xi = h y" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   224
        also from a y have "\<dots> \<le> p y" ..
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   225
        also from x0 y' z have "p y = p (y + a \<cdot> x0)" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   226
        finally show ?thesis .
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   227
      next
61539
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
   228
        txt \<open>In the case \<open>a < 0\<close>, we use \<open>a\<^sub>1\<close>
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
   229
          with \<open>ya\<close> taken as \<open>y / a\<close>:\<close>
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   230
        assume lz: "a < 0" then have nz: "a \<noteq> 0" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   231
        from a1 ay
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   232
        have "- p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y) \<le> xi" ..
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   233
        with lz have "a * xi \<le>
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   234
          a * (- p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y))"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   235
          by (simp add: mult_left_mono_neg order_less_imp_le)
81464
2575f1bd226b tuned proofs;
wenzelm
parents: 63040
diff changeset
   236
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   237
        also have "\<dots> =
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   238
          - a * (p (inverse a \<cdot> y + x0)) - a * (h (inverse a \<cdot> y))"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   239
          by (simp add: right_diff_distrib)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   240
        also from lz x0 y' have "- a * (p (inverse a \<cdot> y + x0)) =
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   241
          p (a \<cdot> (inverse a \<cdot> y + x0))"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   242
          by (simp add: abs_homogenous)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   243
        also from nz x0 y' have "\<dots> = p (y + a \<cdot> x0)"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   244
          by (simp add: add_mult_distrib1 mult_assoc [symmetric])
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   245
        also from nz y have "a * (h (inverse a \<cdot> y)) =  h y"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   246
          by simp
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   247
        finally have "a * xi \<le> p (y + a \<cdot> x0) - h y" .
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   248
        then show ?thesis by simp
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   249
      next
61539
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
   250
        txt \<open>In the case \<open>a > 0\<close>, we use \<open>a\<^sub>2\<close>
a29295dac1ca isabelle update_cartouches -t;
wenzelm
parents: 61486
diff changeset
   251
          with \<open>ya\<close> taken as \<open>y / a\<close>:\<close>
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   252
        assume gz: "0 < a" then have nz: "a \<noteq> 0" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   253
        from a2 ay
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   254
        have "xi \<le> p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y)" ..
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   255
        with gz have "a * xi \<le>
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   256
          a * (p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y))"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   257
          by simp
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   258
        also have "\<dots> = a * p (inverse a \<cdot> y + x0) - a * h (inverse a \<cdot> y)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   259
          by (simp add: right_diff_distrib)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   260
        also from gz x0 y'
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   261
        have "a * p (inverse a \<cdot> y + x0) = p (a \<cdot> (inverse a \<cdot> y + x0))"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   262
          by (simp add: abs_homogenous)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   263
        also from nz x0 y' have "\<dots> = p (y + a \<cdot> x0)"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   264
          by (simp add: add_mult_distrib1 mult_assoc [symmetric])
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   265
        also from nz y have "a * h (inverse a \<cdot> y) = h y"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   266
          by simp
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   267
        finally have "a * xi \<le> p (y + a \<cdot> x0) - h y" .
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
   268
        then show ?thesis by simp
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   269
      qed
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   270
      also from x_rep have "\<dots> = p x" by (simp only:)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 23378
diff changeset
   271
      finally show ?thesis .
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9906
diff changeset
   272
    qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9906
diff changeset
   273
  qed
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 12018
diff changeset
   274
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
   275
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9906
diff changeset
   276
end