src/HOL/Real/HahnBanach/FunctionOrder.thy
author haftmann
Fri, 17 Jun 2005 16:12:49 +0200
changeset 16417 9bc16273c2d4
parent 13515 a6a7025fd7e8
child 19736 d8d0f8f51d69
permissions -rw-r--r--
migrated theory headers to new format
wenzelm@7566
     1
(*  Title:      HOL/Real/HahnBanach/FunctionOrder.thy
wenzelm@7566
     2
    ID:         $Id$
wenzelm@7566
     3
    Author:     Gertrud Bauer, TU Munich
wenzelm@7566
     4
*)
wenzelm@7535
     5
wenzelm@9035
     6
header {* An order on functions *}
wenzelm@7808
     7
haftmann@16417
     8
theory FunctionOrder imports Subspace Linearform begin
wenzelm@7535
     9
wenzelm@9035
    10
subsection {* The graph of a function *}
wenzelm@7808
    11
wenzelm@10687
    12
text {*
wenzelm@10687
    13
  We define the \emph{graph} of a (real) function @{text f} with
wenzelm@10687
    14
  domain @{text F} as the set
wenzelm@10687
    15
  \begin{center}
wenzelm@10687
    16
  @{text "{(x, f x). x \<in> F}"}
wenzelm@10687
    17
  \end{center}
wenzelm@10687
    18
  So we are modeling partial functions by specifying the domain and
wenzelm@10687
    19
  the mapping function. We use the term ``function'' also for its
wenzelm@10687
    20
  graph.
wenzelm@9035
    21
*}
wenzelm@7808
    22
wenzelm@13515
    23
types 'a graph = "('a \<times> real) set"
wenzelm@7535
    24
wenzelm@7535
    25
constdefs
wenzelm@13515
    26
  graph :: "'a set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> 'a graph"
wenzelm@10687
    27
  "graph F f \<equiv> {(x, f x) | x. x \<in> F}"
wenzelm@7535
    28
wenzelm@13515
    29
lemma graphI [intro]: "x \<in> F \<Longrightarrow> (x, f x) \<in> graph F f"
wenzelm@11472
    30
  by (unfold graph_def) blast
wenzelm@7535
    31
wenzelm@13515
    32
lemma graphI2 [intro?]: "x \<in> F \<Longrightarrow> \<exists>t \<in> graph F f. t = (x, f x)"
wenzelm@10687
    33
  by (unfold graph_def) blast
wenzelm@7535
    34
wenzelm@13515
    35
lemma graphE [elim?]:
wenzelm@13515
    36
    "(x, y) \<in> graph F f \<Longrightarrow> (x \<in> F \<Longrightarrow> y = f x \<Longrightarrow> C) \<Longrightarrow> C"
wenzelm@10687
    37
  by (unfold graph_def) blast
wenzelm@10687
    38
wenzelm@7566
    39
wenzelm@9035
    40
subsection {* Functions ordered by domain extension *}
wenzelm@7917
    41
wenzelm@13515
    42
text {*
wenzelm@13515
    43
  A function @{text h'} is an extension of @{text h}, iff the graph of
wenzelm@13515
    44
  @{text h} is a subset of the graph of @{text h'}.
wenzelm@13515
    45
*}
wenzelm@7917
    46
wenzelm@10687
    47
lemma graph_extI:
wenzelm@10687
    48
  "(\<And>x. x \<in> H \<Longrightarrow> h x = h' x) \<Longrightarrow> H \<subseteq> H'
wenzelm@13515
    49
    \<Longrightarrow> graph H h \<subseteq> graph H' h'"
wenzelm@10687
    50
  by (unfold graph_def) blast
wenzelm@7917
    51
wenzelm@13515
    52
lemma graph_extD1 [dest?]:
wenzelm@10687
    53
  "graph H h \<subseteq> graph H' h' \<Longrightarrow> x \<in> H \<Longrightarrow> h x = h' x"
wenzelm@10687
    54
  by (unfold graph_def) blast
wenzelm@7566
    55
wenzelm@13515
    56
lemma graph_extD2 [dest?]:
wenzelm@10687
    57
  "graph H h \<subseteq> graph H' h' \<Longrightarrow> H \<subseteq> H'"
wenzelm@10687
    58
  by (unfold graph_def) blast
wenzelm@7566
    59
wenzelm@13515
    60
wenzelm@9035
    61
subsection {* Domain and function of a graph *}
wenzelm@7917
    62
wenzelm@10687
    63
text {*
wenzelm@13515
    64
  The inverse functions to @{text graph} are @{text domain} and @{text
wenzelm@13515
    65
  funct}.
wenzelm@10687
    66
*}
wenzelm@7917
    67
wenzelm@7917
    68
constdefs
wenzelm@13515
    69
  "domain" :: "'a graph \<Rightarrow> 'a set"
wenzelm@10687
    70
  "domain g \<equiv> {x. \<exists>y. (x, y) \<in> g}"
wenzelm@7917
    71
wenzelm@10687
    72
  funct :: "'a graph \<Rightarrow> ('a \<Rightarrow> real)"
wenzelm@10687
    73
  "funct g \<equiv> \<lambda>x. (SOME y. (x, y) \<in> g)"
wenzelm@7917
    74
wenzelm@10687
    75
text {*
wenzelm@10687
    76
  The following lemma states that @{text g} is the graph of a function
wenzelm@10687
    77
  if the relation induced by @{text g} is unique.
wenzelm@10687
    78
*}
wenzelm@7566
    79
wenzelm@10687
    80
lemma graph_domain_funct:
wenzelm@13515
    81
  assumes uniq: "\<And>x y z. (x, y) \<in> g \<Longrightarrow> (x, z) \<in> g \<Longrightarrow> z = y"
wenzelm@13515
    82
  shows "graph (domain g) (funct g) = g"
wenzelm@13515
    83
proof (unfold domain_def funct_def graph_def, auto)  (* FIXME !? *)
wenzelm@9503
    84
  fix a b assume "(a, b) \<in> g"
paulson@9969
    85
  show "(a, SOME y. (a, y) \<in> g) \<in> g" by (rule someI2)
wenzelm@9503
    86
  show "\<exists>y. (a, y) \<in> g" ..
wenzelm@9503
    87
  show "b = (SOME y. (a, y) \<in> g)"
paulson@9969
    88
  proof (rule some_equality [symmetric])
wenzelm@13515
    89
    fix y assume "(a, y) \<in> g"
wenzelm@13515
    90
    show "y = b" by (rule uniq)
wenzelm@9035
    91
  qed
wenzelm@9035
    92
qed
wenzelm@7535
    93
wenzelm@7808
    94
wenzelm@9035
    95
subsection {* Norm-preserving extensions of a function *}
wenzelm@7917
    96
wenzelm@10687
    97
text {*
wenzelm@10687
    98
  Given a linear form @{text f} on the space @{text F} and a seminorm
wenzelm@10687
    99
  @{text p} on @{text E}. The set of all linear extensions of @{text
wenzelm@10687
   100
  f}, to superspaces @{text H} of @{text F}, which are bounded by
wenzelm@10687
   101
  @{text p}, is defined as follows.
wenzelm@10687
   102
*}
wenzelm@7808
   103
wenzelm@7535
   104
constdefs
wenzelm@10687
   105
  norm_pres_extensions ::
wenzelm@10687
   106
    "'a::{plus, minus, zero} set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> real)
wenzelm@13515
   107
      \<Rightarrow> 'a graph set"
wenzelm@10687
   108
    "norm_pres_extensions E p F f
wenzelm@13515
   109
      \<equiv> {g. \<exists>H h. g = graph H h
wenzelm@13515
   110
          \<and> linearform H h
wenzelm@13515
   111
          \<and> H \<unlhd> E
wenzelm@13515
   112
          \<and> F \<unlhd> H
wenzelm@13515
   113
          \<and> graph F f \<subseteq> graph H h
wenzelm@13515
   114
          \<and> (\<forall>x \<in> H. h x \<le> p x)}"
wenzelm@10687
   115
wenzelm@13515
   116
lemma norm_pres_extensionE [elim]:
wenzelm@9503
   117
  "g \<in> norm_pres_extensions E p F f
wenzelm@13515
   118
  \<Longrightarrow> (\<And>H h. g = graph H h \<Longrightarrow> linearform H h
wenzelm@13515
   119
        \<Longrightarrow> H \<unlhd> E \<Longrightarrow> F \<unlhd> H \<Longrightarrow> graph F f \<subseteq> graph H h
wenzelm@13515
   120
        \<Longrightarrow> \<forall>x \<in> H. h x \<le> p x \<Longrightarrow> C) \<Longrightarrow> C"
wenzelm@10687
   121
  by (unfold norm_pres_extensions_def) blast
wenzelm@7535
   122
wenzelm@10687
   123
lemma norm_pres_extensionI2 [intro]:
wenzelm@13515
   124
  "linearform H h \<Longrightarrow> H \<unlhd> E \<Longrightarrow> F \<unlhd> H
wenzelm@13515
   125
    \<Longrightarrow> graph F f \<subseteq> graph H h \<Longrightarrow> \<forall>x \<in> H. h x \<le> p x
wenzelm@13515
   126
    \<Longrightarrow> graph H h \<in> norm_pres_extensions E p F f"
wenzelm@13515
   127
  by (unfold norm_pres_extensions_def) blast
wenzelm@7535
   128
wenzelm@13515
   129
lemma norm_pres_extensionI:  (* FIXME ? *)
wenzelm@13515
   130
  "\<exists>H h. g = graph H h
wenzelm@13515
   131
    \<and> linearform H h
wenzelm@13515
   132
    \<and> H \<unlhd> E
wenzelm@13515
   133
    \<and> F \<unlhd> H
wenzelm@13515
   134
    \<and> graph F f \<subseteq> graph H h
wenzelm@13515
   135
    \<and> (\<forall>x \<in> H. h x \<le> p x) \<Longrightarrow> g \<in> norm_pres_extensions E p F f"
wenzelm@10687
   136
  by (unfold norm_pres_extensions_def) blast
wenzelm@7535
   137
wenzelm@10687
   138
end