src/HOL/Hahn_Banach/Function_Order.thy
author haftmann
Fri, 23 Oct 2009 17:12:47 +0200
changeset 33085 c1b6cc29496b
parent 31795 be3e1cc5005c
child 41818 6d4c3ee8219d
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31795
be3e1cc5005c standard naming conventions for session and theories;
wenzelm
parents: 29197
diff changeset
     1
(*  Title:      HOL/Hahn_Banach/Function_Order.thy
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     2
    Author:     Gertrud Bauer, TU Munich
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     3
*)
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
     4
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
     5
header {* An order on functions *}
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
     6
31795
be3e1cc5005c standard naming conventions for session and theories;
wenzelm
parents: 29197
diff changeset
     7
theory Function_Order
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 25762
diff changeset
     8
imports Subspace Linearform
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 25762
diff changeset
     9
begin
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    10
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    11
subsection {* The graph of a function *}
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
    12
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
    13
text {*
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
    14
  We define the \emph{graph} of a (real) function @{text f} with
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
    15
  domain @{text F} as the set
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
    16
  \begin{center}
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
    17
  @{text "{(x, f x). x \<in> F}"}
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
    18
  \end{center}
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
    19
  So we are modeling partial functions by specifying the domain and
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
    20
  the mapping function. We use the term ``function'' also for its
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
    21
  graph.
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    22
*}
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    23
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
    24
types 'a graph = "('a \<times> real) set"
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    25
19736
wenzelm
parents: 16417
diff changeset
    26
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    27
  graph :: "'a set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> 'a graph" where
19736
wenzelm
parents: 16417
diff changeset
    28
  "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
    29
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
    30
lemma graphI [intro]: "x \<in> F \<Longrightarrow> (x, f x) \<in> graph F f"
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 25762
diff changeset
    31
  unfolding graph_def by blast
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    32
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
    33
lemma graphI2 [intro?]: "x \<in> F \<Longrightarrow> \<exists>t \<in> graph F f. t = (x, f x)"
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 25762
diff changeset
    34
  unfolding graph_def by blast
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
    35
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
    36
lemma graphE [elim?]:
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
    37
    "(x, y) \<in> graph F f \<Longrightarrow> (x \<in> F \<Longrightarrow> y = f x \<Longrightarrow> C) \<Longrightarrow> C"
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 25762
diff changeset
    38
  unfolding graph_def by blast
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
    39
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    40
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    41
subsection {* Functions ordered by domain extension *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    42
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
    43
text {*
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
    44
  A function @{text h'} is an extension of @{text h}, iff the graph of
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
    45
  @{text h} is a subset of the graph of @{text h'}.
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
    46
*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    47
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
    48
lemma graph_extI:
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
    49
  "(\<And>x. x \<in> H \<Longrightarrow> h x = h' x) \<Longrightarrow> H \<subseteq> H'
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
    50
    \<Longrightarrow> graph H h \<subseteq> graph H' h'"
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 25762
diff changeset
    51
  unfolding graph_def by blast
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    52
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
    53
lemma graph_extD1 [dest?]:
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
    54
  "graph H h \<subseteq> graph H' h' \<Longrightarrow> x \<in> H \<Longrightarrow> h x = h' x"
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 25762
diff changeset
    55
  unfolding graph_def by blast
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
    56
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
    57
lemma graph_extD2 [dest?]:
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
    58
  "graph H h \<subseteq> graph H' h' \<Longrightarrow> H \<subseteq> H'"
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 25762
diff changeset
    59
  unfolding graph_def by blast
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
    60
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
    61
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    62
subsection {* Domain and function of a graph *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    63
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
    64
text {*
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
    65
  The inverse functions to @{text graph} are @{text domain} and @{text
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
    66
  funct}.
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
    67
*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    68
19736
wenzelm
parents: 16417
diff changeset
    69
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    70
  "domain" :: "'a graph \<Rightarrow> 'a set" where
19736
wenzelm
parents: 16417
diff changeset
    71
  "domain g = {x. \<exists>y. (x, y) \<in> g}"
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    72
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    73
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    74
  funct :: "'a graph \<Rightarrow> ('a \<Rightarrow> real)" where
19736
wenzelm
parents: 16417
diff changeset
    75
  "funct g = (\<lambda>x. (SOME y. (x, y) \<in> g))"
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    76
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
    77
text {*
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
    78
  The following lemma states that @{text g} is the graph of a function
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
    79
  if the relation induced by @{text g} is unique.
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
    80
*}
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
    81
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
    82
lemma graph_domain_funct:
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
    83
  assumes uniq: "\<And>x y z. (x, y) \<in> g \<Longrightarrow> (x, z) \<in> g \<Longrightarrow> z = y"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
    84
  shows "graph (domain g) (funct g) = g"
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 25762
diff changeset
    85
  unfolding domain_def funct_def graph_def
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 25762
diff changeset
    86
proof auto  (* FIXME !? *)
23378
1d138d6bb461 tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
    87
  fix a b assume g: "(a, b) \<in> g"
1d138d6bb461 tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
    88
  from g show "(a, SOME y. (a, y) \<in> g) \<in> g" by (rule someI2)
1d138d6bb461 tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
    89
  from g show "\<exists>y. (a, y) \<in> g" ..
1d138d6bb461 tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
    90
  from g show "b = (SOME y. (a, y) \<in> g)"
9969
4753185f1dd2 renamed (most of...) the select rules
paulson
parents: 9623
diff changeset
    91
  proof (rule some_equality [symmetric])
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
    92
    fix y assume "(a, y) \<in> g"
23378
1d138d6bb461 tuned proofs: avoid implicit prems;
wenzelm
parents: 21404
diff changeset
    93
    with g show "y = b" by (rule uniq)
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    94
  qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    95
qed
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    96
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
    97
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8203
diff changeset
    98
subsection {* Norm-preserving extensions of a function *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    99
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
   100
text {*
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
   101
  Given a linear form @{text f} on the space @{text F} and a seminorm
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
   102
  @{text p} on @{text E}. The set of all linear extensions of @{text
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
   103
  f}, to superspaces @{text H} of @{text F}, which are bounded by
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
   104
  @{text p}, is defined as follows.
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
   105
*}
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   106
19736
wenzelm
parents: 16417
diff changeset
   107
definition
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
   108
  norm_pres_extensions ::
25762
c03e9d04b3e4 splitted class uminus from class minus
haftmann
parents: 23378
diff changeset
   109
    "'a::{plus, minus, uminus, zero} set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> real)
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
   110
      \<Rightarrow> 'a graph set" where
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
   111
    "norm_pres_extensions E p F f
19736
wenzelm
parents: 16417
diff changeset
   112
      = {g. \<exists>H h. g = graph H h
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
   113
          \<and> linearform H h
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
   114
          \<and> H \<unlhd> E
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
   115
          \<and> F \<unlhd> H
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
   116
          \<and> graph F f \<subseteq> graph H h
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
   117
          \<and> (\<forall>x \<in> H. h x \<le> p x)}"
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
   118
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
   119
lemma norm_pres_extensionE [elim]:
9503
wenzelm
parents: 9408
diff changeset
   120
  "g \<in> norm_pres_extensions E p F f
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
   121
  \<Longrightarrow> (\<And>H h. g = graph H h \<Longrightarrow> linearform H h
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
   122
        \<Longrightarrow> H \<unlhd> E \<Longrightarrow> F \<unlhd> H \<Longrightarrow> graph F f \<subseteq> graph H h
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
   123
        \<Longrightarrow> \<forall>x \<in> H. h x \<le> p x \<Longrightarrow> C) \<Longrightarrow> C"
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 25762
diff changeset
   124
  unfolding norm_pres_extensions_def by blast
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   125
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
   126
lemma norm_pres_extensionI2 [intro]:
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
   127
  "linearform H h \<Longrightarrow> H \<unlhd> E \<Longrightarrow> F \<unlhd> H
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
   128
    \<Longrightarrow> graph F f \<subseteq> graph H h \<Longrightarrow> \<forall>x \<in> H. h x \<le> p x
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
   129
    \<Longrightarrow> graph H h \<in> norm_pres_extensions E p F f"
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 25762
diff changeset
   130
  unfolding norm_pres_extensions_def by blast
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   131
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
   132
lemma norm_pres_extensionI:  (* FIXME ? *)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
   133
  "\<exists>H h. g = graph H h
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
   134
    \<and> linearform H h
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
   135
    \<and> H \<unlhd> E
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
   136
    \<and> F \<unlhd> H
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
   137
    \<and> graph F f \<subseteq> graph H h
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 11472
diff changeset
   138
    \<and> (\<forall>x \<in> H. h x \<le> p x) \<Longrightarrow> g \<in> norm_pres_extensions E p F f"
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 25762
diff changeset
   139
  unfolding norm_pres_extensions_def by blast
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   140
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9969
diff changeset
   141
end