huffman@15600
|
1 |
(* Title: HOLCF/Cfun.thy
|
huffman@15576
|
2 |
Author: Franz Regensburger
|
huffman@15576
|
3 |
|
huffman@15576
|
4 |
Definition of the type -> of continuous functions.
|
huffman@15576
|
5 |
*)
|
huffman@15576
|
6 |
|
huffman@15576
|
7 |
header {* The type of continuous functions *}
|
huffman@15576
|
8 |
|
huffman@15577
|
9 |
theory Cfun
|
huffman@29533
|
10 |
imports Pcpodef Ffun Product_Cpo
|
huffman@15577
|
11 |
begin
|
huffman@15576
|
12 |
|
huffman@15576
|
13 |
defaultsort cpo
|
huffman@15576
|
14 |
|
huffman@15589
|
15 |
subsection {* Definition of continuous function type *}
|
huffman@15589
|
16 |
|
huffman@16699
|
17 |
lemma Ex_cont: "\<exists>f. cont f"
|
huffman@16699
|
18 |
by (rule exI, rule cont_const)
|
huffman@16699
|
19 |
|
huffman@16699
|
20 |
lemma adm_cont: "adm cont"
|
huffman@16699
|
21 |
by (rule admI, rule cont_lub_fun)
|
huffman@16699
|
22 |
|
huffman@17817
|
23 |
cpodef (CFun) ('a, 'b) "->" (infixr "->" 0) = "{f::'a => 'b. cont f}"
|
wenzelm@29063
|
24 |
by (simp_all add: Ex_cont adm_cont)
|
huffman@15576
|
25 |
|
huffman@17816
|
26 |
syntax (xsymbols)
|
huffman@17816
|
27 |
"->" :: "[type, type] => type" ("(_ \<rightarrow>/ _)" [1,0]0)
|
huffman@17816
|
28 |
|
wenzelm@25131
|
29 |
notation
|
wenzelm@25131
|
30 |
Rep_CFun ("(_$/_)" [999,1000] 999)
|
huffman@15576
|
31 |
|
wenzelm@25131
|
32 |
notation (xsymbols)
|
wenzelm@25131
|
33 |
Rep_CFun ("(_\<cdot>/_)" [999,1000] 999)
|
huffman@15576
|
34 |
|
wenzelm@25131
|
35 |
notation (HTML output)
|
wenzelm@25131
|
36 |
Rep_CFun ("(_\<cdot>/_)" [999,1000] 999)
|
huffman@17816
|
37 |
|
huffman@17832
|
38 |
subsection {* Syntax for continuous lambda abstraction *}
|
huffman@17832
|
39 |
|
huffman@18078
|
40 |
syntax "_cabs" :: "'a"
|
huffman@18078
|
41 |
|
huffman@18078
|
42 |
parse_translation {*
|
huffman@18087
|
43 |
(* rewrites (_cabs x t) => (Abs_CFun (%x. t)) *)
|
wenzelm@25131
|
44 |
[mk_binder_tr ("_cabs", @{const_syntax Abs_CFun})];
|
huffman@18078
|
45 |
*}
|
huffman@17816
|
46 |
|
huffman@17832
|
47 |
text {* To avoid eta-contraction of body: *}
|
huffman@18087
|
48 |
typed_print_translation {*
|
huffman@18078
|
49 |
let
|
huffman@18087
|
50 |
fun cabs_tr' _ _ [Abs abs] = let
|
huffman@18087
|
51 |
val (x,t) = atomic_abs_tr' abs
|
huffman@18087
|
52 |
in Syntax.const "_cabs" $ x $ t end
|
huffman@18087
|
53 |
|
huffman@18087
|
54 |
| cabs_tr' _ T [t] = let
|
huffman@18087
|
55 |
val xT = domain_type (domain_type T);
|
huffman@18087
|
56 |
val abs' = ("x",xT,(incr_boundvars 1 t)$Bound 0);
|
huffman@18087
|
57 |
val (x,t') = atomic_abs_tr' abs';
|
huffman@18087
|
58 |
in Syntax.const "_cabs" $ x $ t' end;
|
huffman@18087
|
59 |
|
wenzelm@25131
|
60 |
in [(@{const_syntax Abs_CFun}, cabs_tr')] end;
|
huffman@17816
|
61 |
*}
|
huffman@17816
|
62 |
|
huffman@18087
|
63 |
text {* Syntax for nested abstractions *}
|
huffman@17832
|
64 |
|
huffman@17832
|
65 |
syntax
|
huffman@18078
|
66 |
"_Lambda" :: "[cargs, 'a] \<Rightarrow> logic" ("(3LAM _./ _)" [1000, 10] 10)
|
huffman@17832
|
67 |
|
huffman@17832
|
68 |
syntax (xsymbols)
|
huffman@25927
|
69 |
"_Lambda" :: "[cargs, 'a] \<Rightarrow> logic" ("(3\<Lambda> _./ _)" [1000, 10] 10)
|
huffman@17832
|
70 |
|
huffman@17816
|
71 |
parse_ast_translation {*
|
huffman@18087
|
72 |
(* rewrites (LAM x y z. t) => (_cabs x (_cabs y (_cabs z t))) *)
|
huffman@18078
|
73 |
(* cf. Syntax.lambda_ast_tr from Syntax/syn_trans.ML *)
|
huffman@18078
|
74 |
let
|
huffman@18078
|
75 |
fun Lambda_ast_tr [pats, body] =
|
huffman@18078
|
76 |
Syntax.fold_ast_p "_cabs" (Syntax.unfold_ast "_cargs" pats, body)
|
huffman@18078
|
77 |
| Lambda_ast_tr asts = raise Syntax.AST ("Lambda_ast_tr", asts);
|
huffman@18078
|
78 |
in [("_Lambda", Lambda_ast_tr)] end;
|
huffman@17816
|
79 |
*}
|
huffman@17816
|
80 |
|
huffman@17816
|
81 |
print_ast_translation {*
|
huffman@18087
|
82 |
(* rewrites (_cabs x (_cabs y (_cabs z t))) => (LAM x y z. t) *)
|
huffman@18078
|
83 |
(* cf. Syntax.abs_ast_tr' from Syntax/syn_trans.ML *)
|
huffman@18078
|
84 |
let
|
huffman@18078
|
85 |
fun cabs_ast_tr' asts =
|
huffman@18078
|
86 |
(case Syntax.unfold_ast_p "_cabs"
|
huffman@18078
|
87 |
(Syntax.Appl (Syntax.Constant "_cabs" :: asts)) of
|
huffman@18078
|
88 |
([], _) => raise Syntax.AST ("cabs_ast_tr'", asts)
|
huffman@18078
|
89 |
| (xs, body) => Syntax.Appl
|
huffman@18078
|
90 |
[Syntax.Constant "_Lambda", Syntax.fold_ast "_cargs" xs, body]);
|
huffman@18078
|
91 |
in [("_cabs", cabs_ast_tr')] end;
|
huffman@17816
|
92 |
*}
|
huffman@15641
|
93 |
|
huffman@18087
|
94 |
text {* Dummy patterns for continuous abstraction *}
|
huffman@18079
|
95 |
translations
|
wenzelm@25131
|
96 |
"\<Lambda> _. t" => "CONST Abs_CFun (\<lambda> _. t)"
|
huffman@18087
|
97 |
|
huffman@18079
|
98 |
|
huffman@17832
|
99 |
subsection {* Continuous function space is pointed *}
|
huffman@15589
|
100 |
|
huffman@16098
|
101 |
lemma UU_CFun: "\<bottom> \<in> CFun"
|
huffman@16098
|
102 |
by (simp add: CFun_def inst_fun_pcpo cont_const)
|
huffman@16098
|
103 |
|
huffman@25827
|
104 |
instance "->" :: (finite_po, finite_po) finite_po
|
huffman@25827
|
105 |
by (rule typedef_finite_po [OF type_definition_CFun])
|
huffman@25827
|
106 |
|
huffman@25827
|
107 |
instance "->" :: (finite_po, chfin) chfin
|
huffman@31076
|
108 |
by (rule typedef_chfin [OF type_definition_CFun below_CFun_def])
|
huffman@25827
|
109 |
|
huffman@26025
|
110 |
instance "->" :: (cpo, discrete_cpo) discrete_cpo
|
huffman@31076
|
111 |
by intro_classes (simp add: below_CFun_def Rep_CFun_inject)
|
huffman@26025
|
112 |
|
huffman@16098
|
113 |
instance "->" :: (cpo, pcpo) pcpo
|
huffman@31076
|
114 |
by (rule typedef_pcpo [OF type_definition_CFun below_CFun_def UU_CFun])
|
huffman@16098
|
115 |
|
huffman@16209
|
116 |
lemmas Rep_CFun_strict =
|
huffman@31076
|
117 |
typedef_Rep_strict [OF type_definition_CFun below_CFun_def UU_CFun]
|
huffman@16209
|
118 |
|
huffman@16209
|
119 |
lemmas Abs_CFun_strict =
|
huffman@31076
|
120 |
typedef_Abs_strict [OF type_definition_CFun below_CFun_def UU_CFun]
|
huffman@16098
|
121 |
|
huffman@17832
|
122 |
text {* function application is strict in its first argument *}
|
huffman@17832
|
123 |
|
huffman@17832
|
124 |
lemma Rep_CFun_strict1 [simp]: "\<bottom>\<cdot>x = \<bottom>"
|
huffman@17832
|
125 |
by (simp add: Rep_CFun_strict)
|
huffman@17832
|
126 |
|
huffman@17832
|
127 |
text {* for compatibility with old HOLCF-Version *}
|
huffman@17832
|
128 |
lemma inst_cfun_pcpo: "\<bottom> = (\<Lambda> x. \<bottom>)"
|
huffman@17832
|
129 |
by (simp add: inst_fun_pcpo [symmetric] Abs_CFun_strict)
|
huffman@17832
|
130 |
|
huffman@17832
|
131 |
subsection {* Basic properties of continuous functions *}
|
huffman@17832
|
132 |
|
huffman@17832
|
133 |
text {* Beta-equality for continuous functions *}
|
huffman@16209
|
134 |
|
huffman@16209
|
135 |
lemma Abs_CFun_inverse2: "cont f \<Longrightarrow> Rep_CFun (Abs_CFun f) = f"
|
huffman@16209
|
136 |
by (simp add: Abs_CFun_inverse CFun_def)
|
huffman@16098
|
137 |
|
huffman@16209
|
138 |
lemma beta_cfun [simp]: "cont f \<Longrightarrow> (\<Lambda> x. f x)\<cdot>u = f u"
|
huffman@16209
|
139 |
by (simp add: Abs_CFun_inverse2)
|
huffman@16209
|
140 |
|
huffman@16209
|
141 |
text {* Eta-equality for continuous functions *}
|
huffman@16209
|
142 |
|
huffman@16209
|
143 |
lemma eta_cfun: "(\<Lambda> x. f\<cdot>x) = f"
|
huffman@16209
|
144 |
by (rule Rep_CFun_inverse)
|
huffman@16209
|
145 |
|
huffman@16209
|
146 |
text {* Extensionality for continuous functions *}
|
huffman@16209
|
147 |
|
huffman@17832
|
148 |
lemma expand_cfun_eq: "(f = g) = (\<forall>x. f\<cdot>x = g\<cdot>x)"
|
huffman@17832
|
149 |
by (simp add: Rep_CFun_inject [symmetric] expand_fun_eq)
|
huffman@17832
|
150 |
|
huffman@16209
|
151 |
lemma ext_cfun: "(\<And>x. f\<cdot>x = g\<cdot>x) \<Longrightarrow> f = g"
|
huffman@17832
|
152 |
by (simp add: expand_cfun_eq)
|
huffman@17832
|
153 |
|
huffman@17832
|
154 |
text {* Extensionality wrt. ordering for continuous functions *}
|
huffman@15576
|
155 |
|
huffman@31076
|
156 |
lemma expand_cfun_below: "f \<sqsubseteq> g = (\<forall>x. f\<cdot>x \<sqsubseteq> g\<cdot>x)"
|
huffman@31076
|
157 |
by (simp add: below_CFun_def expand_fun_below)
|
huffman@17832
|
158 |
|
huffman@31076
|
159 |
lemma below_cfun_ext: "(\<And>x. f\<cdot>x \<sqsubseteq> g\<cdot>x) \<Longrightarrow> f \<sqsubseteq> g"
|
huffman@31076
|
160 |
by (simp add: expand_cfun_below)
|
huffman@17832
|
161 |
|
huffman@17832
|
162 |
text {* Congruence for continuous function application *}
|
huffman@15589
|
163 |
|
huffman@16209
|
164 |
lemma cfun_cong: "\<lbrakk>f = g; x = y\<rbrakk> \<Longrightarrow> f\<cdot>x = g\<cdot>y"
|
huffman@15589
|
165 |
by simp
|
huffman@15589
|
166 |
|
huffman@16209
|
167 |
lemma cfun_fun_cong: "f = g \<Longrightarrow> f\<cdot>x = g\<cdot>x"
|
huffman@15589
|
168 |
by simp
|
huffman@15589
|
169 |
|
huffman@16209
|
170 |
lemma cfun_arg_cong: "x = y \<Longrightarrow> f\<cdot>x = f\<cdot>y"
|
huffman@15589
|
171 |
by simp
|
huffman@15589
|
172 |
|
huffman@16209
|
173 |
subsection {* Continuity of application *}
|
huffman@15576
|
174 |
|
huffman@16209
|
175 |
lemma cont_Rep_CFun1: "cont (\<lambda>f. f\<cdot>x)"
|
huffman@18092
|
176 |
by (rule cont_Rep_CFun [THEN cont2cont_fun])
|
huffman@15576
|
177 |
|
huffman@16209
|
178 |
lemma cont_Rep_CFun2: "cont (\<lambda>x. f\<cdot>x)"
|
huffman@18092
|
179 |
apply (cut_tac x=f in Rep_CFun)
|
huffman@18092
|
180 |
apply (simp add: CFun_def)
|
huffman@15576
|
181 |
done
|
huffman@15576
|
182 |
|
huffman@16209
|
183 |
lemmas monofun_Rep_CFun = cont_Rep_CFun [THEN cont2mono]
|
huffman@16209
|
184 |
lemmas contlub_Rep_CFun = cont_Rep_CFun [THEN cont2contlub]
|
huffman@15589
|
185 |
|
huffman@16209
|
186 |
lemmas monofun_Rep_CFun1 = cont_Rep_CFun1 [THEN cont2mono, standard]
|
huffman@16209
|
187 |
lemmas contlub_Rep_CFun1 = cont_Rep_CFun1 [THEN cont2contlub, standard]
|
huffman@16209
|
188 |
lemmas monofun_Rep_CFun2 = cont_Rep_CFun2 [THEN cont2mono, standard]
|
huffman@16209
|
189 |
lemmas contlub_Rep_CFun2 = cont_Rep_CFun2 [THEN cont2contlub, standard]
|
huffman@16209
|
190 |
|
huffman@16209
|
191 |
text {* contlub, cont properties of @{term Rep_CFun} in each argument *}
|
huffman@16209
|
192 |
|
huffman@27413
|
193 |
lemma contlub_cfun_arg: "chain Y \<Longrightarrow> f\<cdot>(\<Squnion>i. Y i) = (\<Squnion>i. f\<cdot>(Y i))"
|
huffman@16209
|
194 |
by (rule contlub_Rep_CFun2 [THEN contlubE])
|
huffman@15576
|
195 |
|
huffman@27413
|
196 |
lemma cont_cfun_arg: "chain Y \<Longrightarrow> range (\<lambda>i. f\<cdot>(Y i)) <<| f\<cdot>(\<Squnion>i. Y i)"
|
huffman@16209
|
197 |
by (rule cont_Rep_CFun2 [THEN contE])
|
huffman@16209
|
198 |
|
huffman@27413
|
199 |
lemma contlub_cfun_fun: "chain F \<Longrightarrow> (\<Squnion>i. F i)\<cdot>x = (\<Squnion>i. F i\<cdot>x)"
|
huffman@16209
|
200 |
by (rule contlub_Rep_CFun1 [THEN contlubE])
|
huffman@15576
|
201 |
|
huffman@27413
|
202 |
lemma cont_cfun_fun: "chain F \<Longrightarrow> range (\<lambda>i. F i\<cdot>x) <<| (\<Squnion>i. F i)\<cdot>x"
|
huffman@16209
|
203 |
by (rule cont_Rep_CFun1 [THEN contE])
|
huffman@15576
|
204 |
|
huffman@16209
|
205 |
text {* monotonicity of application *}
|
huffman@16209
|
206 |
|
huffman@16209
|
207 |
lemma monofun_cfun_fun: "f \<sqsubseteq> g \<Longrightarrow> f\<cdot>x \<sqsubseteq> g\<cdot>x"
|
huffman@31076
|
208 |
by (simp add: expand_cfun_below)
|
huffman@15576
|
209 |
|
huffman@16209
|
210 |
lemma monofun_cfun_arg: "x \<sqsubseteq> y \<Longrightarrow> f\<cdot>x \<sqsubseteq> f\<cdot>y"
|
huffman@16209
|
211 |
by (rule monofun_Rep_CFun2 [THEN monofunE])
|
huffman@15576
|
212 |
|
huffman@16209
|
213 |
lemma monofun_cfun: "\<lbrakk>f \<sqsubseteq> g; x \<sqsubseteq> y\<rbrakk> \<Longrightarrow> f\<cdot>x \<sqsubseteq> g\<cdot>y"
|
huffman@31076
|
214 |
by (rule below_trans [OF monofun_cfun_fun monofun_cfun_arg])
|
huffman@15576
|
215 |
|
huffman@16209
|
216 |
text {* ch2ch - rules for the type @{typ "'a -> 'b"} *}
|
huffman@15576
|
217 |
|
huffman@16209
|
218 |
lemma chain_monofun: "chain Y \<Longrightarrow> chain (\<lambda>i. f\<cdot>(Y i))"
|
huffman@16209
|
219 |
by (erule monofun_Rep_CFun2 [THEN ch2ch_monofun])
|
huffman@16209
|
220 |
|
huffman@16209
|
221 |
lemma ch2ch_Rep_CFunR: "chain Y \<Longrightarrow> chain (\<lambda>i. f\<cdot>(Y i))"
|
huffman@16209
|
222 |
by (rule monofun_Rep_CFun2 [THEN ch2ch_monofun])
|
huffman@15576
|
223 |
|
huffman@16209
|
224 |
lemma ch2ch_Rep_CFunL: "chain F \<Longrightarrow> chain (\<lambda>i. (F i)\<cdot>x)"
|
huffman@16209
|
225 |
by (rule monofun_Rep_CFun1 [THEN ch2ch_monofun])
|
huffman@15576
|
226 |
|
huffman@18076
|
227 |
lemma ch2ch_Rep_CFun [simp]:
|
huffman@18076
|
228 |
"\<lbrakk>chain F; chain Y\<rbrakk> \<Longrightarrow> chain (\<lambda>i. (F i)\<cdot>(Y i))"
|
huffman@25884
|
229 |
by (simp add: chain_def monofun_cfun)
|
huffman@15576
|
230 |
|
huffman@25884
|
231 |
lemma ch2ch_LAM [simp]:
|
huffman@25884
|
232 |
"\<lbrakk>\<And>x. chain (\<lambda>i. S i x); \<And>i. cont (\<lambda>x. S i x)\<rbrakk> \<Longrightarrow> chain (\<lambda>i. \<Lambda> x. S i x)"
|
huffman@31076
|
233 |
by (simp add: chain_def expand_cfun_below)
|
huffman@18092
|
234 |
|
huffman@16209
|
235 |
text {* contlub, cont properties of @{term Rep_CFun} in both arguments *}
|
huffman@15576
|
236 |
|
huffman@16209
|
237 |
lemma contlub_cfun:
|
huffman@16209
|
238 |
"\<lbrakk>chain F; chain Y\<rbrakk> \<Longrightarrow> (\<Squnion>i. F i)\<cdot>(\<Squnion>i. Y i) = (\<Squnion>i. F i\<cdot>(Y i))"
|
huffman@18076
|
239 |
by (simp add: contlub_cfun_fun contlub_cfun_arg diag_lub)
|
huffman@15576
|
240 |
|
huffman@16209
|
241 |
lemma cont_cfun:
|
huffman@16209
|
242 |
"\<lbrakk>chain F; chain Y\<rbrakk> \<Longrightarrow> range (\<lambda>i. F i\<cdot>(Y i)) <<| (\<Squnion>i. F i)\<cdot>(\<Squnion>i. Y i)"
|
huffman@16209
|
243 |
apply (rule thelubE)
|
huffman@16209
|
244 |
apply (simp only: ch2ch_Rep_CFun)
|
huffman@16209
|
245 |
apply (simp only: contlub_cfun)
|
huffman@16209
|
246 |
done
|
huffman@16209
|
247 |
|
huffman@18092
|
248 |
lemma contlub_LAM:
|
huffman@18092
|
249 |
"\<lbrakk>\<And>x. chain (\<lambda>i. F i x); \<And>i. cont (\<lambda>x. F i x)\<rbrakk>
|
huffman@18092
|
250 |
\<Longrightarrow> (\<Lambda> x. \<Squnion>i. F i x) = (\<Squnion>i. \<Lambda> x. F i x)"
|
huffman@25884
|
251 |
apply (simp add: thelub_CFun)
|
huffman@18092
|
252 |
apply (simp add: Abs_CFun_inverse2)
|
huffman@18092
|
253 |
apply (simp add: thelub_fun ch2ch_lambda)
|
huffman@18092
|
254 |
done
|
huffman@18092
|
255 |
|
huffman@25901
|
256 |
lemmas lub_distribs =
|
huffman@25901
|
257 |
contlub_cfun [symmetric]
|
huffman@25901
|
258 |
contlub_LAM [symmetric]
|
huffman@25901
|
259 |
|
huffman@16209
|
260 |
text {* strictness *}
|
huffman@16209
|
261 |
|
huffman@16209
|
262 |
lemma strictI: "f\<cdot>x = \<bottom> \<Longrightarrow> f\<cdot>\<bottom> = \<bottom>"
|
huffman@16209
|
263 |
apply (rule UU_I)
|
huffman@15576
|
264 |
apply (erule subst)
|
huffman@15576
|
265 |
apply (rule minimal [THEN monofun_cfun_arg])
|
huffman@15576
|
266 |
done
|
huffman@15576
|
267 |
|
huffman@16209
|
268 |
text {* the lub of a chain of continous functions is monotone *}
|
huffman@15576
|
269 |
|
huffman@16209
|
270 |
lemma lub_cfun_mono: "chain F \<Longrightarrow> monofun (\<lambda>x. \<Squnion>i. F i\<cdot>x)"
|
huffman@16209
|
271 |
apply (drule ch2ch_monofun [OF monofun_Rep_CFun])
|
huffman@16209
|
272 |
apply (simp add: thelub_fun [symmetric])
|
huffman@16209
|
273 |
apply (erule monofun_lub_fun)
|
huffman@16209
|
274 |
apply (simp add: monofun_Rep_CFun2)
|
huffman@15576
|
275 |
done
|
huffman@15576
|
276 |
|
huffman@16386
|
277 |
text {* a lemma about the exchange of lubs for type @{typ "'a -> 'b"} *}
|
huffman@15576
|
278 |
|
huffman@16699
|
279 |
lemma ex_lub_cfun:
|
huffman@16699
|
280 |
"\<lbrakk>chain F; chain Y\<rbrakk> \<Longrightarrow> (\<Squnion>j. \<Squnion>i. F j\<cdot>(Y i)) = (\<Squnion>i. \<Squnion>j. F j\<cdot>(Y i))"
|
huffman@18076
|
281 |
by (simp add: diag_lub)
|
huffman@15576
|
282 |
|
huffman@15589
|
283 |
text {* the lub of a chain of cont. functions is continuous *}
|
huffman@15576
|
284 |
|
huffman@16209
|
285 |
lemma cont_lub_cfun: "chain F \<Longrightarrow> cont (\<lambda>x. \<Squnion>i. F i\<cdot>x)"
|
huffman@16209
|
286 |
apply (rule cont2cont_lub)
|
huffman@16209
|
287 |
apply (erule monofun_Rep_CFun [THEN ch2ch_monofun])
|
huffman@16209
|
288 |
apply (rule cont_Rep_CFun2)
|
huffman@15576
|
289 |
done
|
huffman@15576
|
290 |
|
huffman@15589
|
291 |
text {* type @{typ "'a -> 'b"} is chain complete *}
|
huffman@15576
|
292 |
|
huffman@16920
|
293 |
lemma lub_cfun: "chain F \<Longrightarrow> range F <<| (\<Lambda> x. \<Squnion>i. F i\<cdot>x)"
|
huffman@16920
|
294 |
by (simp only: contlub_cfun_fun [symmetric] eta_cfun thelubE)
|
huffman@15576
|
295 |
|
huffman@27413
|
296 |
lemma thelub_cfun: "chain F \<Longrightarrow> (\<Squnion>i. F i) = (\<Lambda> x. \<Squnion>i. F i\<cdot>x)"
|
huffman@16920
|
297 |
by (rule lub_cfun [THEN thelubI])
|
huffman@15576
|
298 |
|
huffman@17832
|
299 |
subsection {* Continuity simplification procedure *}
|
huffman@15589
|
300 |
|
huffman@15589
|
301 |
text {* cont2cont lemma for @{term Rep_CFun} *}
|
huffman@15576
|
302 |
|
huffman@29530
|
303 |
lemma cont2cont_Rep_CFun [cont2cont]:
|
huffman@29049
|
304 |
assumes f: "cont (\<lambda>x. f x)"
|
huffman@29049
|
305 |
assumes t: "cont (\<lambda>x. t x)"
|
huffman@29049
|
306 |
shows "cont (\<lambda>x. (f x)\<cdot>(t x))"
|
huffman@29049
|
307 |
proof -
|
huffman@29049
|
308 |
have "cont (\<lambda>x. Rep_CFun (f x))"
|
huffman@29049
|
309 |
using cont_Rep_CFun f by (rule cont2cont_app3)
|
huffman@29049
|
310 |
thus "cont (\<lambda>x. (f x)\<cdot>(t x))"
|
huffman@29049
|
311 |
using cont_Rep_CFun2 t by (rule cont2cont_app2)
|
huffman@29049
|
312 |
qed
|
huffman@15576
|
313 |
|
huffman@15589
|
314 |
text {* cont2mono Lemma for @{term "%x. LAM y. c1(x)(y)"} *}
|
huffman@15576
|
315 |
|
huffman@15576
|
316 |
lemma cont2mono_LAM:
|
huffman@29049
|
317 |
"\<lbrakk>\<And>x. cont (\<lambda>y. f x y); \<And>y. monofun (\<lambda>x. f x y)\<rbrakk>
|
huffman@29049
|
318 |
\<Longrightarrow> monofun (\<lambda>x. \<Lambda> y. f x y)"
|
huffman@31076
|
319 |
unfolding monofun_def expand_cfun_below by simp
|
huffman@15576
|
320 |
|
huffman@29049
|
321 |
text {* cont2cont Lemma for @{term "%x. LAM y. f x y"} *}
|
huffman@15576
|
322 |
|
huffman@29530
|
323 |
text {*
|
huffman@29530
|
324 |
Not suitable as a cont2cont rule, because on nested lambdas
|
huffman@29530
|
325 |
it causes exponential blow-up in the number of subgoals.
|
huffman@29530
|
326 |
*}
|
huffman@29530
|
327 |
|
huffman@15576
|
328 |
lemma cont2cont_LAM:
|
huffman@29049
|
329 |
assumes f1: "\<And>x. cont (\<lambda>y. f x y)"
|
huffman@29049
|
330 |
assumes f2: "\<And>y. cont (\<lambda>x. f x y)"
|
huffman@29049
|
331 |
shows "cont (\<lambda>x. \<Lambda> y. f x y)"
|
huffman@29049
|
332 |
proof (rule cont_Abs_CFun)
|
huffman@29049
|
333 |
fix x
|
huffman@29049
|
334 |
from f1 show "f x \<in> CFun" by (simp add: CFun_def)
|
huffman@29049
|
335 |
from f2 show "cont f" by (rule cont2cont_lambda)
|
huffman@29049
|
336 |
qed
|
huffman@15576
|
337 |
|
huffman@29530
|
338 |
text {*
|
huffman@29530
|
339 |
This version does work as a cont2cont rule, since it
|
huffman@29530
|
340 |
has only a single subgoal.
|
huffman@29530
|
341 |
*}
|
huffman@29530
|
342 |
|
huffman@29530
|
343 |
lemma cont2cont_LAM' [cont2cont]:
|
huffman@29530
|
344 |
fixes f :: "'a::cpo \<Rightarrow> 'b::cpo \<Rightarrow> 'c::cpo"
|
huffman@29530
|
345 |
assumes f: "cont (\<lambda>p. f (fst p) (snd p))"
|
huffman@29530
|
346 |
shows "cont (\<lambda>x. \<Lambda> y. f x y)"
|
huffman@29530
|
347 |
proof (rule cont2cont_LAM)
|
huffman@31041
|
348 |
fix x :: 'a show "cont (\<lambda>y. f x y)"
|
huffman@31041
|
349 |
using f by (rule cont_fst_snd_D2)
|
huffman@29530
|
350 |
next
|
huffman@31041
|
351 |
fix y :: 'b show "cont (\<lambda>x. f x y)"
|
huffman@31041
|
352 |
using f by (rule cont_fst_snd_D1)
|
huffman@29530
|
353 |
qed
|
huffman@29530
|
354 |
|
huffman@29530
|
355 |
lemma cont2cont_LAM_discrete [cont2cont]:
|
huffman@29530
|
356 |
"(\<And>y::'a::discrete_cpo. cont (\<lambda>x. f x y)) \<Longrightarrow> cont (\<lambda>x. \<Lambda> y. f x y)"
|
huffman@29530
|
357 |
by (simp add: cont2cont_LAM)
|
huffman@15576
|
358 |
|
huffman@16055
|
359 |
lemmas cont_lemmas1 =
|
huffman@16055
|
360 |
cont_const cont_id cont_Rep_CFun2 cont2cont_Rep_CFun cont2cont_LAM
|
huffman@16055
|
361 |
|
huffman@17832
|
362 |
subsection {* Miscellaneous *}
|
huffman@17832
|
363 |
|
huffman@17832
|
364 |
text {* Monotonicity of @{term Abs_CFun} *}
|
huffman@15576
|
365 |
|
huffman@17832
|
366 |
lemma semi_monofun_Abs_CFun:
|
huffman@17832
|
367 |
"\<lbrakk>cont f; cont g; f \<sqsubseteq> g\<rbrakk> \<Longrightarrow> Abs_CFun f \<sqsubseteq> Abs_CFun g"
|
huffman@31076
|
368 |
by (simp add: below_CFun_def Abs_CFun_inverse2)
|
huffman@15576
|
369 |
|
huffman@15589
|
370 |
text {* some lemmata for functions with flat/chfin domain/range types *}
|
huffman@15576
|
371 |
|
huffman@15576
|
372 |
lemma chfin_Rep_CFunR: "chain (Y::nat => 'a::cpo->'b::chfin)
|
huffman@27413
|
373 |
==> !s. ? n. (LUB i. Y i)$s = Y n$s"
|
huffman@15576
|
374 |
apply (rule allI)
|
huffman@15576
|
375 |
apply (subst contlub_cfun_fun)
|
huffman@15576
|
376 |
apply assumption
|
huffman@15576
|
377 |
apply (fast intro!: thelubI chfin lub_finch2 chfin2finch ch2ch_Rep_CFunL)
|
huffman@15576
|
378 |
done
|
huffman@15576
|
379 |
|
huffman@18089
|
380 |
lemma adm_chfindom: "adm (\<lambda>(u::'a::cpo \<rightarrow> 'b::chfin). P(u\<cdot>s))"
|
huffman@18089
|
381 |
by (rule adm_subst, simp, rule adm_chfin)
|
huffman@18089
|
382 |
|
huffman@16085
|
383 |
subsection {* Continuous injection-retraction pairs *}
|
huffman@15589
|
384 |
|
huffman@16085
|
385 |
text {* Continuous retractions are strict. *}
|
huffman@15576
|
386 |
|
huffman@16085
|
387 |
lemma retraction_strict:
|
huffman@16085
|
388 |
"\<forall>x. f\<cdot>(g\<cdot>x) = x \<Longrightarrow> f\<cdot>\<bottom> = \<bottom>"
|
huffman@15576
|
389 |
apply (rule UU_I)
|
huffman@16085
|
390 |
apply (drule_tac x="\<bottom>" in spec)
|
huffman@16085
|
391 |
apply (erule subst)
|
huffman@16085
|
392 |
apply (rule monofun_cfun_arg)
|
huffman@16085
|
393 |
apply (rule minimal)
|
huffman@15576
|
394 |
done
|
huffman@15576
|
395 |
|
huffman@16085
|
396 |
lemma injection_eq:
|
huffman@16085
|
397 |
"\<forall>x. f\<cdot>(g\<cdot>x) = x \<Longrightarrow> (g\<cdot>x = g\<cdot>y) = (x = y)"
|
huffman@16085
|
398 |
apply (rule iffI)
|
huffman@16085
|
399 |
apply (drule_tac f=f in cfun_arg_cong)
|
huffman@16085
|
400 |
apply simp
|
huffman@16085
|
401 |
apply simp
|
huffman@15576
|
402 |
done
|
huffman@15576
|
403 |
|
huffman@31076
|
404 |
lemma injection_below:
|
huffman@16314
|
405 |
"\<forall>x. f\<cdot>(g\<cdot>x) = x \<Longrightarrow> (g\<cdot>x \<sqsubseteq> g\<cdot>y) = (x \<sqsubseteq> y)"
|
huffman@16314
|
406 |
apply (rule iffI)
|
huffman@16314
|
407 |
apply (drule_tac f=f in monofun_cfun_arg)
|
huffman@16314
|
408 |
apply simp
|
huffman@16314
|
409 |
apply (erule monofun_cfun_arg)
|
huffman@16314
|
410 |
done
|
huffman@16314
|
411 |
|
huffman@16085
|
412 |
lemma injection_defined_rev:
|
huffman@16085
|
413 |
"\<lbrakk>\<forall>x. f\<cdot>(g\<cdot>x) = x; g\<cdot>z = \<bottom>\<rbrakk> \<Longrightarrow> z = \<bottom>"
|
huffman@16085
|
414 |
apply (drule_tac f=f in cfun_arg_cong)
|
huffman@16085
|
415 |
apply (simp add: retraction_strict)
|
huffman@15576
|
416 |
done
|
huffman@15576
|
417 |
|
huffman@16085
|
418 |
lemma injection_defined:
|
huffman@16085
|
419 |
"\<lbrakk>\<forall>x. f\<cdot>(g\<cdot>x) = x; z \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> g\<cdot>z \<noteq> \<bottom>"
|
huffman@16085
|
420 |
by (erule contrapos_nn, rule injection_defined_rev)
|
huffman@16085
|
421 |
|
huffman@16085
|
422 |
text {* propagation of flatness and chain-finiteness by retractions *}
|
huffman@16085
|
423 |
|
huffman@16085
|
424 |
lemma chfin2chfin:
|
huffman@16085
|
425 |
"\<forall>y. (f::'a::chfin \<rightarrow> 'b)\<cdot>(g\<cdot>y) = y
|
huffman@16085
|
426 |
\<Longrightarrow> \<forall>Y::nat \<Rightarrow> 'b. chain Y \<longrightarrow> (\<exists>n. max_in_chain n Y)"
|
huffman@16085
|
427 |
apply clarify
|
huffman@16085
|
428 |
apply (drule_tac f=g in chain_monofun)
|
huffman@25921
|
429 |
apply (drule chfin)
|
huffman@16085
|
430 |
apply (unfold max_in_chain_def)
|
huffman@16085
|
431 |
apply (simp add: injection_eq)
|
huffman@16085
|
432 |
done
|
huffman@16085
|
433 |
|
huffman@16085
|
434 |
lemma flat2flat:
|
huffman@16085
|
435 |
"\<forall>y. (f::'a::flat \<rightarrow> 'b::pcpo)\<cdot>(g\<cdot>y) = y
|
huffman@16085
|
436 |
\<Longrightarrow> \<forall>x y::'b. x \<sqsubseteq> y \<longrightarrow> x = \<bottom> \<or> x = y"
|
huffman@16085
|
437 |
apply clarify
|
huffman@16209
|
438 |
apply (drule_tac f=g in monofun_cfun_arg)
|
huffman@25920
|
439 |
apply (drule ax_flat)
|
huffman@16085
|
440 |
apply (erule disjE)
|
huffman@16085
|
441 |
apply (simp add: injection_defined_rev)
|
huffman@16085
|
442 |
apply (simp add: injection_eq)
|
huffman@15576
|
443 |
done
|
huffman@15576
|
444 |
|
huffman@15589
|
445 |
text {* a result about functions with flat codomain *}
|
huffman@15576
|
446 |
|
huffman@16085
|
447 |
lemma flat_eqI: "\<lbrakk>(x::'a::flat) \<sqsubseteq> y; x \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> x = y"
|
huffman@25920
|
448 |
by (drule ax_flat, simp)
|
huffman@16085
|
449 |
|
huffman@16085
|
450 |
lemma flat_codom:
|
huffman@16085
|
451 |
"f\<cdot>x = (c::'b::flat) \<Longrightarrow> f\<cdot>\<bottom> = \<bottom> \<or> (\<forall>z. f\<cdot>z = c)"
|
huffman@16085
|
452 |
apply (case_tac "f\<cdot>x = \<bottom>")
|
huffman@15576
|
453 |
apply (rule disjI1)
|
huffman@15576
|
454 |
apply (rule UU_I)
|
huffman@16085
|
455 |
apply (erule_tac t="\<bottom>" in subst)
|
huffman@15576
|
456 |
apply (rule minimal [THEN monofun_cfun_arg])
|
huffman@16085
|
457 |
apply clarify
|
huffman@16085
|
458 |
apply (rule_tac a = "f\<cdot>\<bottom>" in refl [THEN box_equals])
|
huffman@16085
|
459 |
apply (erule minimal [THEN monofun_cfun_arg, THEN flat_eqI])
|
huffman@16085
|
460 |
apply (erule minimal [THEN monofun_cfun_arg, THEN flat_eqI])
|
huffman@15589
|
461 |
done
|
huffman@15589
|
462 |
|
huffman@15589
|
463 |
|
huffman@15589
|
464 |
subsection {* Identity and composition *}
|
huffman@15589
|
465 |
|
wenzelm@25135
|
466 |
definition
|
wenzelm@25135
|
467 |
ID :: "'a \<rightarrow> 'a" where
|
wenzelm@25135
|
468 |
"ID = (\<Lambda> x. x)"
|
wenzelm@25135
|
469 |
|
wenzelm@25135
|
470 |
definition
|
wenzelm@25135
|
471 |
cfcomp :: "('b \<rightarrow> 'c) \<rightarrow> ('a \<rightarrow> 'b) \<rightarrow> 'a \<rightarrow> 'c" where
|
wenzelm@25135
|
472 |
oo_def: "cfcomp = (\<Lambda> f g x. f\<cdot>(g\<cdot>x))"
|
huffman@15589
|
473 |
|
wenzelm@25131
|
474 |
abbreviation
|
wenzelm@25131
|
475 |
cfcomp_syn :: "['b \<rightarrow> 'c, 'a \<rightarrow> 'b] \<Rightarrow> 'a \<rightarrow> 'c" (infixr "oo" 100) where
|
wenzelm@25131
|
476 |
"f oo g == cfcomp\<cdot>f\<cdot>g"
|
huffman@15589
|
477 |
|
huffman@16085
|
478 |
lemma ID1 [simp]: "ID\<cdot>x = x"
|
huffman@16085
|
479 |
by (simp add: ID_def)
|
huffman@15576
|
480 |
|
huffman@16085
|
481 |
lemma cfcomp1: "(f oo g) = (\<Lambda> x. f\<cdot>(g\<cdot>x))"
|
huffman@15589
|
482 |
by (simp add: oo_def)
|
huffman@15576
|
483 |
|
huffman@16085
|
484 |
lemma cfcomp2 [simp]: "(f oo g)\<cdot>x = f\<cdot>(g\<cdot>x)"
|
huffman@15589
|
485 |
by (simp add: cfcomp1)
|
huffman@15576
|
486 |
|
huffman@27274
|
487 |
lemma cfcomp_LAM: "cont g \<Longrightarrow> f oo (\<Lambda> x. g x) = (\<Lambda> x. f\<cdot>(g x))"
|
huffman@27274
|
488 |
by (simp add: cfcomp1)
|
huffman@27274
|
489 |
|
huffman@19709
|
490 |
lemma cfcomp_strict [simp]: "\<bottom> oo f = \<bottom>"
|
huffman@19709
|
491 |
by (simp add: expand_cfun_eq)
|
huffman@19709
|
492 |
|
huffman@15589
|
493 |
text {*
|
huffman@15589
|
494 |
Show that interpretation of (pcpo,@{text "_->_"}) is a category.
|
huffman@15589
|
495 |
The class of objects is interpretation of syntactical class pcpo.
|
huffman@15589
|
496 |
The class of arrows between objects @{typ 'a} and @{typ 'b} is interpret. of @{typ "'a -> 'b"}.
|
huffman@15589
|
497 |
The identity arrow is interpretation of @{term ID}.
|
huffman@15589
|
498 |
The composition of f and g is interpretation of @{text "oo"}.
|
huffman@15589
|
499 |
*}
|
huffman@15576
|
500 |
|
huffman@16085
|
501 |
lemma ID2 [simp]: "f oo ID = f"
|
huffman@15589
|
502 |
by (rule ext_cfun, simp)
|
huffman@15576
|
503 |
|
huffman@16085
|
504 |
lemma ID3 [simp]: "ID oo f = f"
|
huffman@15589
|
505 |
by (rule ext_cfun, simp)
|
huffman@15576
|
506 |
|
huffman@15576
|
507 |
lemma assoc_oo: "f oo (g oo h) = (f oo g) oo h"
|
huffman@15589
|
508 |
by (rule ext_cfun, simp)
|
huffman@15576
|
509 |
|
huffman@16085
|
510 |
|
huffman@16085
|
511 |
subsection {* Strictified functions *}
|
huffman@16085
|
512 |
|
huffman@16085
|
513 |
defaultsort pcpo
|
huffman@16085
|
514 |
|
wenzelm@25131
|
515 |
definition
|
wenzelm@25131
|
516 |
strictify :: "('a \<rightarrow> 'b) \<rightarrow> 'a \<rightarrow> 'b" where
|
wenzelm@25131
|
517 |
"strictify = (\<Lambda> f x. if x = \<bottom> then \<bottom> else f\<cdot>x)"
|
huffman@16085
|
518 |
|
huffman@16085
|
519 |
text {* results about strictify *}
|
huffman@16085
|
520 |
|
huffman@17815
|
521 |
lemma cont_strictify1: "cont (\<lambda>f. if x = \<bottom> then \<bottom> else f\<cdot>x)"
|
huffman@17815
|
522 |
by (simp add: cont_if)
|
huffman@16085
|
523 |
|
huffman@17815
|
524 |
lemma monofun_strictify2: "monofun (\<lambda>x. if x = \<bottom> then \<bottom> else f\<cdot>x)"
|
huffman@17815
|
525 |
apply (rule monofunI)
|
huffman@25786
|
526 |
apply (auto simp add: monofun_cfun_arg)
|
huffman@16085
|
527 |
done
|
huffman@16085
|
528 |
|
huffman@17815
|
529 |
(*FIXME: long proof*)
|
huffman@25723
|
530 |
lemma contlub_strictify2: "contlub (\<lambda>x. if x = \<bottom> then \<bottom> else f\<cdot>x)"
|
huffman@16209
|
531 |
apply (rule contlubI)
|
huffman@27413
|
532 |
apply (case_tac "(\<Squnion>i. Y i) = \<bottom>")
|
huffman@16699
|
533 |
apply (drule (1) chain_UU_I)
|
huffman@18076
|
534 |
apply simp
|
huffman@17815
|
535 |
apply (simp del: if_image_distrib)
|
huffman@17815
|
536 |
apply (simp only: contlub_cfun_arg)
|
huffman@16085
|
537 |
apply (rule lub_equal2)
|
huffman@16085
|
538 |
apply (rule chain_mono2 [THEN exE])
|
huffman@16085
|
539 |
apply (erule chain_UU_I_inverse2)
|
huffman@16085
|
540 |
apply (assumption)
|
huffman@17815
|
541 |
apply (rule_tac x=x in exI, clarsimp)
|
huffman@16085
|
542 |
apply (erule chain_monofun)
|
huffman@17815
|
543 |
apply (erule monofun_strictify2 [THEN ch2ch_monofun])
|
huffman@16085
|
544 |
done
|
huffman@16085
|
545 |
|
huffman@17815
|
546 |
lemmas cont_strictify2 =
|
huffman@17815
|
547 |
monocontlub2cont [OF monofun_strictify2 contlub_strictify2, standard]
|
huffman@17815
|
548 |
|
huffman@17815
|
549 |
lemma strictify_conv_if: "strictify\<cdot>f\<cdot>x = (if x = \<bottom> then \<bottom> else f\<cdot>x)"
|
huffman@29530
|
550 |
unfolding strictify_def
|
huffman@29530
|
551 |
by (simp add: cont_strictify1 cont_strictify2 cont2cont_LAM)
|
huffman@16085
|
552 |
|
huffman@16085
|
553 |
lemma strictify1 [simp]: "strictify\<cdot>f\<cdot>\<bottom> = \<bottom>"
|
huffman@17815
|
554 |
by (simp add: strictify_conv_if)
|
huffman@16085
|
555 |
|
huffman@16085
|
556 |
lemma strictify2 [simp]: "x \<noteq> \<bottom> \<Longrightarrow> strictify\<cdot>f\<cdot>x = f\<cdot>x"
|
huffman@17815
|
557 |
by (simp add: strictify_conv_if)
|
huffman@16085
|
558 |
|
huffman@17816
|
559 |
subsection {* Continuous let-bindings *}
|
huffman@17816
|
560 |
|
wenzelm@25131
|
561 |
definition
|
wenzelm@25131
|
562 |
CLet :: "'a \<rightarrow> ('a \<rightarrow> 'b) \<rightarrow> 'b" where
|
wenzelm@25131
|
563 |
"CLet = (\<Lambda> s f. f\<cdot>s)"
|
huffman@17816
|
564 |
|
huffman@17816
|
565 |
syntax
|
huffman@17816
|
566 |
"_CLet" :: "[letbinds, 'a] => 'a" ("(Let (_)/ in (_))" 10)
|
huffman@17816
|
567 |
|
huffman@17816
|
568 |
translations
|
huffman@17816
|
569 |
"_CLet (_binds b bs) e" == "_CLet b (_CLet bs e)"
|
wenzelm@25131
|
570 |
"Let x = a in e" == "CONST CLet\<cdot>a\<cdot>(\<Lambda> x. e)"
|
huffman@17816
|
571 |
|
huffman@15576
|
572 |
end
|