diff -r 71af1fd6a5e4 -r be3e1cc5005c src/HOL/Hahn_Banach/Function_Order.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hahn_Banach/Function_Order.thy Wed Jun 24 21:46:54 2009 +0200 @@ -0,0 +1,141 @@ +(* Title: HOL/Hahn_Banach/Function_Order.thy + Author: Gertrud Bauer, TU Munich +*) + +header {* An order on functions *} + +theory Function_Order +imports Subspace Linearform +begin + +subsection {* The graph of a function *} + +text {* + We define the \emph{graph} of a (real) function @{text f} with + domain @{text F} as the set + \begin{center} + @{text "{(x, f x). x \ F}"} + \end{center} + So we are modeling partial functions by specifying the domain and + the mapping function. We use the term ``function'' also for its + graph. +*} + +types 'a graph = "('a \ real) set" + +definition + graph :: "'a set \ ('a \ real) \ 'a graph" where + "graph F f = {(x, f x) | x. x \ F}" + +lemma graphI [intro]: "x \ F \ (x, f x) \ graph F f" + unfolding graph_def by blast + +lemma graphI2 [intro?]: "x \ F \ \t \ graph F f. t = (x, f x)" + unfolding graph_def by blast + +lemma graphE [elim?]: + "(x, y) \ graph F f \ (x \ F \ y = f x \ C) \ C" + unfolding graph_def by blast + + +subsection {* Functions ordered by domain extension *} + +text {* + A function @{text h'} is an extension of @{text h}, iff the graph of + @{text h} is a subset of the graph of @{text h'}. +*} + +lemma graph_extI: + "(\x. x \ H \ h x = h' x) \ H \ H' + \ graph H h \ graph H' h'" + unfolding graph_def by blast + +lemma graph_extD1 [dest?]: + "graph H h \ graph H' h' \ x \ H \ h x = h' x" + unfolding graph_def by blast + +lemma graph_extD2 [dest?]: + "graph H h \ graph H' h' \ H \ H'" + unfolding graph_def by blast + + +subsection {* Domain and function of a graph *} + +text {* + The inverse functions to @{text graph} are @{text domain} and @{text + funct}. +*} + +definition + "domain" :: "'a graph \ 'a set" where + "domain g = {x. \y. (x, y) \ g}" + +definition + funct :: "'a graph \ ('a \ real)" where + "funct g = (\x. (SOME y. (x, y) \ g))" + +text {* + The following lemma states that @{text g} is the graph of a function + if the relation induced by @{text g} is unique. +*} + +lemma graph_domain_funct: + assumes uniq: "\x y z. (x, y) \ g \ (x, z) \ g \ z = y" + shows "graph (domain g) (funct g) = g" + unfolding domain_def funct_def graph_def +proof auto (* FIXME !? *) + fix a b assume g: "(a, b) \ g" + from g show "(a, SOME y. (a, y) \ g) \ g" by (rule someI2) + from g show "\y. (a, y) \ g" .. + from g show "b = (SOME y. (a, y) \ g)" + proof (rule some_equality [symmetric]) + fix y assume "(a, y) \ g" + with g show "y = b" by (rule uniq) + qed +qed + + +subsection {* Norm-preserving extensions of a function *} + +text {* + Given a linear form @{text f} on the space @{text F} and a seminorm + @{text p} on @{text E}. The set of all linear extensions of @{text + f}, to superspaces @{text H} of @{text F}, which are bounded by + @{text p}, is defined as follows. +*} + +definition + norm_pres_extensions :: + "'a::{plus, minus, uminus, zero} set \ ('a \ real) \ 'a set \ ('a \ real) + \ 'a graph set" where + "norm_pres_extensions E p F f + = {g. \H h. g = graph H h + \ linearform H h + \ H \ E + \ F \ H + \ graph F f \ graph H h + \ (\x \ H. h x \ p x)}" + +lemma norm_pres_extensionE [elim]: + "g \ norm_pres_extensions E p F f + \ (\H h. g = graph H h \ linearform H h + \ H \ E \ F \ H \ graph F f \ graph H h + \ \x \ H. h x \ p x \ C) \ C" + unfolding norm_pres_extensions_def by blast + +lemma norm_pres_extensionI2 [intro]: + "linearform H h \ H \ E \ F \ H + \ graph F f \ graph H h \ \x \ H. h x \ p x + \ graph H h \ norm_pres_extensions E p F f" + unfolding norm_pres_extensions_def by blast + +lemma norm_pres_extensionI: (* FIXME ? *) + "\H h. g = graph H h + \ linearform H h + \ H \ E + \ F \ H + \ graph F f \ graph H h + \ (\x \ H. h x \ p x) \ g \ norm_pres_extensions E p F f" + unfolding norm_pres_extensions_def by blast + +end