| author | wenzelm | 
| Sun, 09 Oct 2016 16:27:01 +0200 | |
| changeset 64121 | f2c8f6b11dcf | 
| parent 61879 | e4f9d8f094fe | 
| permissions | -rw-r--r-- | 
| 
31795
 
be3e1cc5005c
standard naming conventions for session and theories;
 
wenzelm 
parents: 
29197 
diff
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  | 
|
| 58889 | 5  | 
section \<open>An order on functions\<close>  | 
| 7808 | 6  | 
|
| 
31795
 
be3e1cc5005c
standard naming conventions for session and theories;
 
wenzelm 
parents: 
29197 
diff
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  | 
|
| 58744 | 11  | 
subsection \<open>The graph of a function\<close>  | 
| 7808 | 12  | 
|
| 58744 | 13  | 
text \<open>  | 
| 61879 | 14  | 
We define the \<^emph>\<open>graph\<close> of a (real) function \<open>f\<close> with domain \<open>F\<close> as the set  | 
| 10687 | 15  | 
  \begin{center}
 | 
| 61539 | 16  | 
  \<open>{(x, f x). x \<in> F}\<close>
 | 
| 10687 | 17  | 
  \end{center}
 | 
| 61879 | 18  | 
So we are modeling partial functions by specifying the domain and the  | 
19  | 
mapping function. We use the term ``function'' also for its graph.  | 
|
| 58744 | 20  | 
\<close>  | 
| 
7535
 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 
wenzelm 
parents:  
diff
changeset
 | 
21  | 
|
| 41818 | 22  | 
type_synonym 'a graph = "('a \<times> real) set"
 | 
| 
7535
 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 
wenzelm 
parents:  
diff
changeset
 | 
23  | 
|
| 44887 | 24  | 
definition graph :: "'a set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> 'a graph"
 | 
25  | 
  where "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  | 
|
| 13515 | 27  | 
lemma graphI [intro]: "x \<in> F \<Longrightarrow> (x, f x) \<in> graph F f"  | 
| 27612 | 28  | 
unfolding graph_def by blast  | 
| 
7535
 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 
wenzelm 
parents:  
diff
changeset
 | 
29  | 
|
| 13515 | 30  | 
lemma graphI2 [intro?]: "x \<in> F \<Longrightarrow> \<exists>t \<in> graph F f. t = (x, f x)"  | 
| 27612 | 31  | 
unfolding graph_def by blast  | 
| 7566 | 32  | 
|
| 13515 | 33  | 
lemma graphE [elim?]:  | 
| 44887 | 34  | 
assumes "(x, y) \<in> graph F f"  | 
35  | 
obtains "x \<in> F" and "y = f x"  | 
|
36  | 
using assms unfolding graph_def by blast  | 
|
| 10687 | 37  | 
|
| 
7535
 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 
wenzelm 
parents:  
diff
changeset
 | 
38  | 
|
| 58744 | 39  | 
subsection \<open>Functions ordered by domain extension\<close>  | 
| 7917 | 40  | 
|
| 58744 | 41  | 
text \<open>  | 
| 61879 | 42  | 
A function \<open>h'\<close> is an extension of \<open>h\<close>, iff the graph of \<open>h\<close> is a subset of  | 
43  | 
the graph of \<open>h'\<close>.  | 
|
| 58744 | 44  | 
\<close>  | 
| 7917 | 45  | 
|
| 10687 | 46  | 
lemma graph_extI:  | 
47  | 
"(\<And>x. x \<in> H \<Longrightarrow> h x = h' x) \<Longrightarrow> H \<subseteq> H'  | 
|
| 13515 | 48  | 
\<Longrightarrow> graph H h \<subseteq> graph H' h'"  | 
| 27612 | 49  | 
unfolding graph_def by blast  | 
| 7917 | 50  | 
|
| 44887 | 51  | 
lemma graph_extD1 [dest?]: "graph H h \<subseteq> graph H' h' \<Longrightarrow> x \<in> H \<Longrightarrow> h x = h' x"  | 
| 27612 | 52  | 
unfolding graph_def by blast  | 
| 7566 | 53  | 
|
| 44887 | 54  | 
lemma graph_extD2 [dest?]: "graph H h \<subseteq> graph H' h' \<Longrightarrow> H \<subseteq> H'"  | 
| 27612 | 55  | 
unfolding graph_def by blast  | 
| 7566 | 56  | 
|
| 13515 | 57  | 
|
| 58744 | 58  | 
subsection \<open>Domain and function of a graph\<close>  | 
| 7917 | 59  | 
|
| 58744 | 60  | 
text \<open>  | 
| 61539 | 61  | 
The inverse functions to \<open>graph\<close> are \<open>domain\<close> and \<open>funct\<close>.  | 
| 58744 | 62  | 
\<close>  | 
| 7917 | 63  | 
|
| 44887 | 64  | 
definition domain :: "'a graph \<Rightarrow> 'a set"  | 
65  | 
  where "domain g = {x. \<exists>y. (x, y) \<in> g}"
 | 
|
| 7917 | 66  | 
|
| 44887 | 67  | 
definition funct :: "'a graph \<Rightarrow> ('a \<Rightarrow> real)"
 | 
68  | 
where "funct g = (\<lambda>x. (SOME y. (x, y) \<in> g))"  | 
|
| 7917 | 69  | 
|
| 58744 | 70  | 
text \<open>  | 
| 61540 | 71  | 
The following lemma states that \<open>g\<close> is the graph of a function if the  | 
72  | 
relation induced by \<open>g\<close> is unique.  | 
|
| 58744 | 73  | 
\<close>  | 
| 7566 | 74  | 
|
| 10687 | 75  | 
lemma graph_domain_funct:  | 
| 13515 | 76  | 
assumes uniq: "\<And>x y z. (x, y) \<in> g \<Longrightarrow> (x, z) \<in> g \<Longrightarrow> z = y"  | 
77  | 
shows "graph (domain g) (funct g) = g"  | 
|
| 27612 | 78  | 
unfolding domain_def funct_def graph_def  | 
79  | 
proof auto (* FIXME !? *)  | 
|
| 23378 | 80  | 
fix a b assume g: "(a, b) \<in> g"  | 
81  | 
from g show "(a, SOME y. (a, y) \<in> g) \<in> g" by (rule someI2)  | 
|
82  | 
from g show "\<exists>y. (a, y) \<in> g" ..  | 
|
83  | 
from g show "b = (SOME y. (a, y) \<in> g)"  | 
|
| 9969 | 84  | 
proof (rule some_equality [symmetric])  | 
| 13515 | 85  | 
fix y assume "(a, y) \<in> g"  | 
| 23378 | 86  | 
with g show "y = b" by (rule uniq)  | 
| 9035 | 87  | 
qed  | 
88  | 
qed  | 
|
| 
7535
 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 
wenzelm 
parents:  
diff
changeset
 | 
89  | 
|
| 7808 | 90  | 
|
| 58744 | 91  | 
subsection \<open>Norm-preserving extensions of a function\<close>  | 
| 7917 | 92  | 
|
| 58744 | 93  | 
text \<open>  | 
| 61879 | 94  | 
Given a linear form \<open>f\<close> on the space \<open>F\<close> and a seminorm \<open>p\<close> on \<open>E\<close>. The set  | 
95  | 
of all linear extensions of \<open>f\<close>, to superspaces \<open>H\<close> of \<open>F\<close>, which are  | 
|
| 61540 | 96  | 
bounded by \<open>p\<close>, is defined as follows.  | 
| 58744 | 97  | 
\<close>  | 
| 7808 | 98  | 
|
| 19736 | 99  | 
definition  | 
| 10687 | 100  | 
norm_pres_extensions ::  | 
| 58745 | 101  | 
    "'a::{plus,minus,uminus,zero} set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> real)
 | 
| 44887 | 102  | 
\<Rightarrow> 'a graph set"  | 
103  | 
where  | 
|
104  | 
"norm_pres_extensions E p F f  | 
|
105  | 
    = {g. \<exists>H h. g = graph H h
 | 
|
106  | 
\<and> linearform H h  | 
|
107  | 
\<and> H \<unlhd> E  | 
|
108  | 
\<and> F \<unlhd> H  | 
|
109  | 
\<and> graph F f \<subseteq> graph H h  | 
|
110  | 
\<and> (\<forall>x \<in> H. h x \<le> p x)}"  | 
|
| 10687 | 111  | 
|
| 13515 | 112  | 
lemma norm_pres_extensionE [elim]:  | 
| 44887 | 113  | 
assumes "g \<in> norm_pres_extensions E p F f"  | 
114  | 
obtains H h  | 
|
115  | 
where "g = graph H h"  | 
|
116  | 
and "linearform H h"  | 
|
117  | 
and "H \<unlhd> E"  | 
|
118  | 
and "F \<unlhd> H"  | 
|
119  | 
and "graph F f \<subseteq> graph H h"  | 
|
120  | 
and "\<forall>x \<in> H. h x \<le> p x"  | 
|
121  | 
using assms unfolding norm_pres_extensions_def by blast  | 
|
| 
7535
 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 
wenzelm 
parents:  
diff
changeset
 | 
122  | 
|
| 10687 | 123  | 
lemma norm_pres_extensionI2 [intro]:  | 
| 13515 | 124  | 
"linearform H h \<Longrightarrow> H \<unlhd> E \<Longrightarrow> F \<unlhd> H  | 
125  | 
\<Longrightarrow> graph F f \<subseteq> graph H h \<Longrightarrow> \<forall>x \<in> H. h x \<le> p x  | 
|
126  | 
\<Longrightarrow> graph H h \<in> norm_pres_extensions E p F f"  | 
|
| 27612 | 127  | 
unfolding norm_pres_extensions_def by blast  | 
| 
7535
 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 
wenzelm 
parents:  
diff
changeset
 | 
128  | 
|
| 13515 | 129  | 
lemma norm_pres_extensionI: (* FIXME ? *)  | 
130  | 
"\<exists>H h. g = graph H h  | 
|
131  | 
\<and> linearform H h  | 
|
132  | 
\<and> H \<unlhd> E  | 
|
133  | 
\<and> F \<unlhd> H  | 
|
134  | 
\<and> graph F f \<subseteq> graph H h  | 
|
135  | 
\<and> (\<forall>x \<in> H. h x \<le> p x) \<Longrightarrow> g \<in> norm_pres_extensions E p F f"  | 
|
| 27612 | 136  | 
unfolding norm_pres_extensions_def by blast  | 
| 
7535
 
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
 
wenzelm 
parents:  
diff
changeset
 | 
137  | 
|
| 10687 | 138  | 
end  |