src/HOL/Real/HahnBanach/FunctionOrder.thy
author bauerg
Mon, 17 Jul 2000 13:58:18 +0200
changeset 9374 153853af318b
parent 9035 371f023d3dbd
child 9379 21cfeae6659d
permissions -rw-r--r--
- xsymbols for \<noteq> \<notin> \<in> \<exists> \<forall> \<and> \<inter> \<union> \<Union> - vector space type of {plus, minus, zero}, overload 0 in vector space - syntax |.| and ||.||
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     1
(*  Title:      HOL/Real/HahnBanach/FunctionOrder.thy
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     2
    ID:         $Id$
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     3
    Author:     Gertrud Bauer, TU Munich
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     4
*)
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
     5
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
     6
header {* An order on functions *}
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
     7
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
     8
theory FunctionOrder = Subspace + Linearform:
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
     9
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    10
subsection {* The graph of a function *}
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
    11
7978
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
    12
text{* We define the \emph{graph} of a (real) function $f$ with
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    13
domain $F$ as the set 
7927
b50446a33c16 update by Gertrud Bauer;
wenzelm
parents: 7917
diff changeset
    14
\[
9374
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
    15
\{(x, f\ap x). \ap x \in F\}
7927
b50446a33c16 update by Gertrud Bauer;
wenzelm
parents: 7917
diff changeset
    16
\]
7978
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
    17
So we are modeling partial functions by specifying the domain and 
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
    18
the mapping function. We use the term ``function'' also for its graph.
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    19
*}
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    20
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    21
types 'a graph = "('a * real) set"
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    22
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    23
constdefs
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    24
  graph :: "['a set, 'a => real] => 'a graph "
9374
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
    25
  "graph F f == {(x, f x) | x. x \<in> F}" 
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    26
9374
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
    27
lemma graphI [intro??]: "x \<in> F ==> (x, f x) \<in> graph F f"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    28
  by (unfold graph_def, intro CollectI exI) force
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    29
9374
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
    30
lemma graphI2 [intro??]: "x \<in> F ==> \<exists>t\<in> (graph F f). t = (x, f x)"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    31
  by (unfold graph_def, force)
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
    32
9374
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
    33
lemma graphD1 [intro??]: "(x, y) \<in> graph F f ==> x \<in> F"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    34
  by (unfold graph_def, elim CollectE exE) force
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
    35
9374
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
    36
lemma graphD2 [intro??]: "(x, y) \<in> graph H h ==> y = h x"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    37
  by (unfold graph_def, elim CollectE exE) force 
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    38
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    39
subsection {* Functions ordered by domain extension *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    40
7978
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
    41
text{* A function $h'$ is an extension of $h$, iff the graph of 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    42
$h$ is a subset of the graph of $h'$.*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    43
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    44
lemma graph_extI: 
9374
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
    45
  "[| !! x. x \<in> H ==> h x = h' x; H <= H'|]
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    46
  ==> graph H h <= graph H' h'"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    47
  by (unfold graph_def, force)
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    48
8203
2fcc6017cb72 intro/elim/dest attributes: changed ! / !! flags to ? / ??;
wenzelm
parents: 7978
diff changeset
    49
lemma graph_extD1 [intro??]: 
9374
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
    50
  "[| graph H h <= graph H' h'; x \<in> H |] ==> h x = h' x"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    51
  by (unfold graph_def, force)
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
    52
8203
2fcc6017cb72 intro/elim/dest attributes: changed ! / !! flags to ? / ??;
wenzelm
parents: 7978
diff changeset
    53
lemma graph_extD2 [intro??]: 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    54
  "[| graph H h <= graph H' h' |] ==> H <= H'"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    55
  by (unfold graph_def, force)
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
    56
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    57
subsection {* Domain and function of a graph *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    58
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    59
text{* The inverse functions to $\idt{graph}$ are $\idt{domain}$ and 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    60
$\idt{funct}$.*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    61
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    62
constdefs
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    63
  domain :: "'a graph => 'a set" 
9374
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
    64
  "domain g == {x. \<exists>y. (x, y) \<in> g}"
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    65
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    66
  funct :: "'a graph => ('a => real)"
9374
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
    67
  "funct g == \<lambda>x. (SOME y. (x, y) \<in> g)"
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    68
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    69
(*text{*  The equations 
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    70
\begin{matharray}
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    71
\idt{domain} graph F f = F {\rm and}\\ 
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    72
\idt{funct} graph F f = f
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    73
\end{matharray}
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    74
hold, but are not proved here.
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    75
*}*)
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    76
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    77
text {* The following lemma states that $g$ is the graph of a function
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    78
if the relation induced by $g$ is unique. *}
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
    79
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
    80
lemma graph_domain_funct: 
9374
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
    81
  "(!!x y z. (x, y) \<in> g ==> (x, z) \<in> g ==> z = y) 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    82
  ==> graph (domain g) (funct g) = g"
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    83
proof (unfold domain_def funct_def graph_def, auto)
9374
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
    84
  fix a b assume "(a, b) \<in> g"
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
    85
  show "(a, SOME y. (a, y) \<in> g) \<in> g" by (rule selectI2)
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
    86
  show "\<exists>y. (a, y) \<in> g" ..
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
    87
  assume uniq: "!!x y z. (x, y) \<in> g ==> (x, z) \<in> g ==> z = y"
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
    88
  show "b = (SOME y. (a, y) \<in> g)"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    89
  proof (rule select_equality [RS sym])
9374
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
    90
    fix y assume "(a, y) \<in> g" show "y = b" by (rule uniq)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    91
  qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    92
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    93
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
    94
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
    95
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    96
subsection {* Norm-preserving extensions of a function *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    97
7978
1b99ee57d131 final update by Gertrud Bauer;
wenzelm
parents: 7927
diff changeset
    98
text {* Given a linear form $f$ on the space $F$ and a seminorm $p$ on 
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    99
$E$. The set of all linear extensions of $f$, to superspaces $H$ of 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
   100
$F$, which are bounded by $p$, is defined as follows. *}
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   101
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   102
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   103
constdefs
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   104
  norm_pres_extensions :: 
9374
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
   105
    "['a::{plus, minus, zero} set, 'a  => real, 'a set, 'a => real] 
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   106
    => 'a graph set"
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   107
    "norm_pres_extensions E p F f 
9374
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
   108
    == {g. \<exists>H h. graph H h = g 
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
   109
                \<and> is_linearform H h 
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
   110
                \<and> is_subspace H E
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
   111
                \<and> is_subspace F H
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
   112
                \<and> graph F f <= graph H h
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
   113
                \<and> (\<forall>x \<in> H. h x <= p x)}"
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   114
                       
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   115
lemma norm_pres_extension_D:  
9374
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
   116
  "g \<in> norm_pres_extensions E p F f
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
   117
  ==> \<exists>H h. graph H h = g 
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
   118
            \<and> is_linearform H h 
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
   119
            \<and> is_subspace H E
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
   120
            \<and> is_subspace F H
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
   121
            \<and> graph F f <= graph H h
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
   122
            \<and> (\<forall>x \<in> H. h x <= p x)"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
   123
  by (unfold norm_pres_extensions_def) force
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   124
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   125
lemma norm_pres_extensionI2 [intro]:  
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   126
  "[| is_linearform H h; is_subspace H E; is_subspace F H;
9374
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
   127
  graph F f <= graph H h; \<forall>x \<in> H. h x <= p x |]
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
   128
  ==> (graph H h \<in> norm_pres_extensions E p F f)"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
   129
 by (unfold norm_pres_extensions_def) force
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   130
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   131
lemma norm_pres_extensionI [intro]:  
9374
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
   132
  "\<exists>H h. graph H h = g 
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
   133
         \<and> is_linearform H h    
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
   134
         \<and> is_subspace H E
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
   135
         \<and> is_subspace F H
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
   136
         \<and> graph F f <= graph H h
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
   137
         \<and> (\<forall>x \<in> H. h x <= p x)
153853af318b - xsymbols for
bauerg
parents: 9035
diff changeset
   138
  ==> g\<in> norm_pres_extensions E p F f"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
   139
  by (unfold norm_pres_extensions_def) force
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   140
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
   141
end