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