25904
|
1 |
(* Title: HOLCF/CompactBasis.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Brian Huffman
|
|
4 |
*)
|
|
5 |
|
|
6 |
header {* Compact bases of domains *}
|
|
7 |
|
|
8 |
theory CompactBasis
|
|
9 |
imports Bifinite SetPcpo
|
|
10 |
begin
|
|
11 |
|
|
12 |
subsection {* Ideals over a preorder *}
|
|
13 |
|
|
14 |
locale preorder =
|
|
15 |
fixes r :: "'a::type \<Rightarrow> 'a \<Rightarrow> bool"
|
|
16 |
assumes refl: "r x x"
|
|
17 |
assumes trans: "\<lbrakk>r x y; r y z\<rbrakk> \<Longrightarrow> r x z"
|
|
18 |
begin
|
|
19 |
|
|
20 |
definition
|
|
21 |
ideal :: "'a set \<Rightarrow> bool" where
|
|
22 |
"ideal A = ((\<exists>x. x \<in> A) \<and> (\<forall>x\<in>A. \<forall>y\<in>A. \<exists>z\<in>A. r x z \<and> r y z) \<and>
|
|
23 |
(\<forall>x y. r x y \<longrightarrow> y \<in> A \<longrightarrow> x \<in> A))"
|
|
24 |
|
|
25 |
lemma idealI:
|
|
26 |
assumes "\<exists>x. x \<in> A"
|
|
27 |
assumes "\<And>x y. \<lbrakk>x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> \<exists>z\<in>A. r x z \<and> r y z"
|
|
28 |
assumes "\<And>x y. \<lbrakk>r x y; y \<in> A\<rbrakk> \<Longrightarrow> x \<in> A"
|
|
29 |
shows "ideal A"
|
|
30 |
unfolding ideal_def using prems by fast
|
|
31 |
|
|
32 |
lemma idealD1:
|
|
33 |
"ideal A \<Longrightarrow> \<exists>x. x \<in> A"
|
|
34 |
unfolding ideal_def by fast
|
|
35 |
|
|
36 |
lemma idealD2:
|
|
37 |
"\<lbrakk>ideal A; x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> \<exists>z\<in>A. r x z \<and> r y z"
|
|
38 |
unfolding ideal_def by fast
|
|
39 |
|
|
40 |
lemma idealD3:
|
|
41 |
"\<lbrakk>ideal A; r x y; y \<in> A\<rbrakk> \<Longrightarrow> x \<in> A"
|
|
42 |
unfolding ideal_def by fast
|
|
43 |
|
|
44 |
lemma ideal_directed_finite:
|
|
45 |
assumes A: "ideal A"
|
|
46 |
shows "\<lbrakk>finite U; U \<subseteq> A\<rbrakk> \<Longrightarrow> \<exists>z\<in>A. \<forall>x\<in>U. r x z"
|
|
47 |
apply (induct U set: finite)
|
|
48 |
apply (simp add: idealD1 [OF A])
|
|
49 |
apply (simp, clarify, rename_tac y)
|
|
50 |
apply (drule (1) idealD2 [OF A])
|
|
51 |
apply (clarify, erule_tac x=z in rev_bexI)
|
|
52 |
apply (fast intro: trans)
|
|
53 |
done
|
|
54 |
|
|
55 |
lemma ideal_principal: "ideal {x. r x z}"
|
|
56 |
apply (rule idealI)
|
|
57 |
apply (rule_tac x=z in exI)
|
|
58 |
apply (fast intro: refl)
|
|
59 |
apply (rule_tac x=z in bexI, fast)
|
|
60 |
apply (fast intro: refl)
|
|
61 |
apply (fast intro: trans)
|
|
62 |
done
|
|
63 |
|
|
64 |
lemma directed_image_ideal:
|
|
65 |
assumes A: "ideal A"
|
|
66 |
assumes f: "\<And>x y. r x y \<Longrightarrow> f x \<sqsubseteq> f y"
|
|
67 |
shows "directed (f ` A)"
|
|
68 |
apply (rule directedI)
|
|
69 |
apply (cut_tac idealD1 [OF A], fast)
|
|
70 |
apply (clarify, rename_tac a b)
|
|
71 |
apply (drule (1) idealD2 [OF A])
|
|
72 |
apply (clarify, rename_tac c)
|
|
73 |
apply (rule_tac x="f c" in rev_bexI)
|
|
74 |
apply (erule imageI)
|
|
75 |
apply (simp add: f)
|
|
76 |
done
|
|
77 |
|
|
78 |
lemma adm_ideal: "adm (\<lambda>A. ideal A)"
|
|
79 |
unfolding ideal_def by (intro adm_lemmas adm_set_lemmas)
|
|
80 |
|
|
81 |
end
|
|
82 |
|
|
83 |
subsection {* Defining functions in terms of basis elements *}
|
|
84 |
|
|
85 |
lemma (in preorder) lub_image_principal:
|
|
86 |
assumes f: "\<And>x y. r x y \<Longrightarrow> f x \<sqsubseteq> f y"
|
|
87 |
shows "(\<Squnion>x\<in>{x. r x y}. f x) = f y"
|
|
88 |
apply (rule thelubI)
|
|
89 |
apply (rule is_lub_maximal)
|
|
90 |
apply (rule ub_imageI)
|
|
91 |
apply (simp add: f)
|
|
92 |
apply (rule imageI)
|
|
93 |
apply (simp add: refl)
|
|
94 |
done
|
|
95 |
|
|
96 |
lemma finite_directed_has_lub: "\<lbrakk>finite S; directed S\<rbrakk> \<Longrightarrow> \<exists>u. S <<| u"
|
|
97 |
apply (drule (1) directed_finiteD, rule subset_refl)
|
|
98 |
apply (erule bexE)
|
|
99 |
apply (rule_tac x=z in exI)
|
|
100 |
apply (erule (1) is_lub_maximal)
|
|
101 |
done
|
|
102 |
|
|
103 |
lemma is_ub_thelub0: "\<lbrakk>\<exists>u. S <<| u; x \<in> S\<rbrakk> \<Longrightarrow> x \<sqsubseteq> lub S"
|
|
104 |
apply (erule exE, drule lubI)
|
|
105 |
apply (drule is_lubD1)
|
|
106 |
apply (erule (1) is_ubD)
|
|
107 |
done
|
|
108 |
|
|
109 |
lemma is_lub_thelub0: "\<lbrakk>\<exists>u. S <<| u; S <| x\<rbrakk> \<Longrightarrow> lub S \<sqsubseteq> x"
|
|
110 |
by (erule exE, drule lubI, erule is_lub_lub)
|
|
111 |
|
|
112 |
locale bifinite_basis = preorder +
|
|
113 |
fixes principal :: "'a::type \<Rightarrow> 'b::cpo"
|
|
114 |
fixes approxes :: "'b::cpo \<Rightarrow> 'a::type set"
|
|
115 |
assumes ideal_approxes: "\<And>x. preorder.ideal r (approxes x)"
|
|
116 |
assumes cont_approxes: "cont approxes"
|
|
117 |
assumes approxes_principal: "\<And>a. approxes (principal a) = {b. r b a}"
|
|
118 |
assumes subset_approxesD: "\<And>x y. approxes x \<subseteq> approxes y \<Longrightarrow> x \<sqsubseteq> y"
|
|
119 |
|
|
120 |
fixes take :: "nat \<Rightarrow> 'a::type \<Rightarrow> 'a"
|
|
121 |
assumes take_less: "r (take n a) a"
|
|
122 |
assumes take_take: "take n (take n a) = take n a"
|
|
123 |
assumes take_mono: "r a b \<Longrightarrow> r (take n a) (take n b)"
|
|
124 |
assumes take_chain: "r (take n a) (take (Suc n) a)"
|
|
125 |
assumes finite_range_take: "finite (range (take n))"
|
|
126 |
assumes take_covers: "\<exists>n. take n a = a"
|
|
127 |
begin
|
|
128 |
|
|
129 |
lemma finite_take_approxes: "finite (take n ` approxes x)"
|
|
130 |
by (rule finite_subset [OF image_mono [OF subset_UNIV] finite_range_take])
|
|
131 |
|
|
132 |
lemma basis_fun_lemma0:
|
|
133 |
fixes f :: "'a::type \<Rightarrow> 'c::cpo"
|
|
134 |
assumes f_mono: "\<And>a b. r a b \<Longrightarrow> f a \<sqsubseteq> f b"
|
|
135 |
shows "\<exists>u. f ` take i ` approxes x <<| u"
|
|
136 |
apply (rule finite_directed_has_lub)
|
|
137 |
apply (rule finite_imageI)
|
|
138 |
apply (rule finite_take_approxes)
|
|
139 |
apply (subst image_image)
|
|
140 |
apply (rule directed_image_ideal)
|
|
141 |
apply (rule ideal_approxes)
|
|
142 |
apply (rule f_mono)
|
|
143 |
apply (erule take_mono)
|
|
144 |
done
|
|
145 |
|
|
146 |
lemma basis_fun_lemma1:
|
|
147 |
fixes f :: "'a::type \<Rightarrow> 'c::cpo"
|
|
148 |
assumes f_mono: "\<And>a b. r a b \<Longrightarrow> f a \<sqsubseteq> f b"
|
|
149 |
shows "chain (\<lambda>i. lub (f ` take i ` approxes x))"
|
|
150 |
apply (rule chainI)
|
|
151 |
apply (rule is_lub_thelub0)
|
|
152 |
apply (rule basis_fun_lemma0, erule f_mono)
|
|
153 |
apply (rule is_ubI, clarsimp, rename_tac a)
|
|
154 |
apply (rule trans_less [OF f_mono [OF take_chain]])
|
|
155 |
apply (rule is_ub_thelub0)
|
|
156 |
apply (rule basis_fun_lemma0, erule f_mono)
|
|
157 |
apply simp
|
|
158 |
done
|
|
159 |
|
|
160 |
lemma basis_fun_lemma2:
|
|
161 |
fixes f :: "'a::type \<Rightarrow> 'c::cpo"
|
|
162 |
assumes f_mono: "\<And>a b. r a b \<Longrightarrow> f a \<sqsubseteq> f b"
|
|
163 |
shows "f ` approxes x <<| (\<Squnion>i. lub (f ` take i ` approxes x))"
|
|
164 |
apply (rule is_lubI)
|
|
165 |
apply (rule ub_imageI, rename_tac a)
|
|
166 |
apply (cut_tac a=a in take_covers, erule exE, rename_tac i)
|
|
167 |
apply (erule subst)
|
|
168 |
apply (rule rev_trans_less)
|
|
169 |
apply (rule_tac x=i in is_ub_thelub)
|
|
170 |
apply (rule basis_fun_lemma1, erule f_mono)
|
|
171 |
apply (rule is_ub_thelub0)
|
|
172 |
apply (rule basis_fun_lemma0, erule f_mono)
|
|
173 |
apply simp
|
|
174 |
apply (rule is_lub_thelub [OF _ ub_rangeI])
|
|
175 |
apply (rule basis_fun_lemma1, erule f_mono)
|
|
176 |
apply (rule is_lub_thelub0)
|
|
177 |
apply (rule basis_fun_lemma0, erule f_mono)
|
|
178 |
apply (rule is_ubI, clarsimp, rename_tac a)
|
|
179 |
apply (rule trans_less [OF f_mono [OF take_less]])
|
|
180 |
apply (erule (1) ub_imageD)
|
|
181 |
done
|
|
182 |
|
|
183 |
lemma basis_fun_lemma:
|
|
184 |
fixes f :: "'a::type \<Rightarrow> 'c::cpo"
|
|
185 |
assumes f_mono: "\<And>a b. r a b \<Longrightarrow> f a \<sqsubseteq> f b"
|
|
186 |
shows "\<exists>u. f ` approxes x <<| u"
|
|
187 |
by (rule exI, rule basis_fun_lemma2, erule f_mono)
|
|
188 |
|
|
189 |
lemma approxes_mono: "x \<sqsubseteq> y \<Longrightarrow> approxes x \<subseteq> approxes y"
|
|
190 |
apply (drule cont_approxes [THEN cont2mono, THEN monofunE])
|
|
191 |
apply (simp add: set_cpo_simps)
|
|
192 |
done
|
|
193 |
|
|
194 |
lemma approxes_contlub:
|
|
195 |
"chain Y \<Longrightarrow> approxes (\<Squnion>i. Y i) = (\<Union>i. approxes (Y i))"
|
|
196 |
by (simp add: cont2contlubE [OF cont_approxes] set_cpo_simps)
|
|
197 |
|
|
198 |
lemma less_def: "(x \<sqsubseteq> y) = (approxes x \<subseteq> approxes y)"
|
|
199 |
by (rule iffI [OF approxes_mono subset_approxesD])
|
|
200 |
|
|
201 |
lemma approxes_eq: "approxes x = {a. principal a \<sqsubseteq> x}"
|
|
202 |
unfolding less_def approxes_principal
|
|
203 |
apply safe
|
|
204 |
apply (erule (1) idealD3 [OF ideal_approxes])
|
|
205 |
apply (erule subsetD, simp add: refl)
|
|
206 |
done
|
|
207 |
|
|
208 |
lemma mem_approxes_iff: "(a \<in> approxes x) = (principal a \<sqsubseteq> x)"
|
|
209 |
by (simp add: approxes_eq)
|
|
210 |
|
|
211 |
lemma principal_less_iff: "(principal a \<sqsubseteq> x) = (a \<in> approxes x)"
|
|
212 |
by (simp add: approxes_eq)
|
|
213 |
|
|
214 |
lemma approxesD: "a \<in> approxes x \<Longrightarrow> principal a \<sqsubseteq> x"
|
|
215 |
by (simp add: approxes_eq)
|
|
216 |
|
|
217 |
lemma principal_mono: "r a b \<Longrightarrow> principal a \<sqsubseteq> principal b"
|
|
218 |
by (rule approxesD, simp add: approxes_principal)
|
|
219 |
|
|
220 |
lemma lessI: "(\<And>a. principal a \<sqsubseteq> x \<Longrightarrow> principal a \<sqsubseteq> u) \<Longrightarrow> x \<sqsubseteq> u"
|
|
221 |
unfolding principal_less_iff
|
|
222 |
by (simp add: less_def subset_def)
|
|
223 |
|
|
224 |
lemma lub_principal_approxes: "principal ` approxes x <<| x"
|
|
225 |
apply (rule is_lubI)
|
|
226 |
apply (rule ub_imageI)
|
|
227 |
apply (erule approxesD)
|
|
228 |
apply (subst less_def)
|
|
229 |
apply (rule subsetI)
|
|
230 |
apply (drule (1) ub_imageD)
|
|
231 |
apply (simp add: approxes_eq)
|
|
232 |
done
|
|
233 |
|
|
234 |
definition
|
|
235 |
basis_fun :: "('a::type \<Rightarrow> 'c::cpo) \<Rightarrow> 'b \<rightarrow> 'c" where
|
|
236 |
"basis_fun = (\<lambda>f. (\<Lambda> x. lub (f ` approxes x)))"
|
|
237 |
|
|
238 |
lemma basis_fun_beta:
|
|
239 |
fixes f :: "'a::type \<Rightarrow> 'c::cpo"
|
|
240 |
assumes f_mono: "\<And>a b. r a b \<Longrightarrow> f a \<sqsubseteq> f b"
|
|
241 |
shows "basis_fun f\<cdot>x = lub (f ` approxes x)"
|
|
242 |
unfolding basis_fun_def
|
|
243 |
proof (rule beta_cfun)
|
|
244 |
have lub: "\<And>x. \<exists>u. f ` approxes x <<| u"
|
|
245 |
using f_mono by (rule basis_fun_lemma)
|
|
246 |
show cont: "cont (\<lambda>x. lub (f ` approxes x))"
|
|
247 |
apply (rule contI2)
|
|
248 |
apply (rule monofunI)
|
|
249 |
apply (rule is_lub_thelub0 [OF lub ub_imageI])
|
|
250 |
apply (rule is_ub_thelub0 [OF lub imageI])
|
|
251 |
apply (erule (1) subsetD [OF approxes_mono])
|
|
252 |
apply (rule is_lub_thelub0 [OF lub ub_imageI])
|
|
253 |
apply (simp add: approxes_contlub, clarify)
|
|
254 |
apply (erule rev_trans_less [OF is_ub_thelub])
|
|
255 |
apply (erule is_ub_thelub0 [OF lub imageI])
|
|
256 |
done
|
|
257 |
qed
|
|
258 |
|
|
259 |
lemma basis_fun_principal:
|
|
260 |
fixes f :: "'a::type \<Rightarrow> 'c::cpo"
|
|
261 |
assumes f_mono: "\<And>a b. r a b \<Longrightarrow> f a \<sqsubseteq> f b"
|
|
262 |
shows "basis_fun f\<cdot>(principal a) = f a"
|
|
263 |
apply (subst basis_fun_beta, erule f_mono)
|
|
264 |
apply (subst approxes_principal)
|
|
265 |
apply (rule lub_image_principal, erule f_mono)
|
|
266 |
done
|
|
267 |
|
|
268 |
lemma basis_fun_mono:
|
|
269 |
assumes f_mono: "\<And>a b. r a b \<Longrightarrow> f a \<sqsubseteq> f b"
|
|
270 |
assumes g_mono: "\<And>a b. r a b \<Longrightarrow> g a \<sqsubseteq> g b"
|
|
271 |
assumes less: "\<And>a. f a \<sqsubseteq> g a"
|
|
272 |
shows "basis_fun f \<sqsubseteq> basis_fun g"
|
|
273 |
apply (rule less_cfun_ext)
|
|
274 |
apply (simp only: basis_fun_beta f_mono g_mono)
|
|
275 |
apply (rule is_lub_thelub0)
|
|
276 |
apply (rule basis_fun_lemma, erule f_mono)
|
|
277 |
apply (rule ub_imageI, rename_tac a)
|
|
278 |
apply (rule trans_less [OF less])
|
|
279 |
apply (rule is_ub_thelub0)
|
|
280 |
apply (rule basis_fun_lemma, erule g_mono)
|
|
281 |
apply (erule imageI)
|
|
282 |
done
|
|
283 |
|
|
284 |
lemma compact_principal: "compact (principal a)"
|
|
285 |
by (rule compactI2, simp add: principal_less_iff approxes_contlub)
|
|
286 |
|
|
287 |
lemma chain_basis_fun_take:
|
|
288 |
"chain (\<lambda>i. basis_fun (\<lambda>a. principal (take i a)))"
|
|
289 |
apply (rule chainI)
|
|
290 |
apply (rule basis_fun_mono)
|
|
291 |
apply (erule principal_mono [OF take_mono])
|
|
292 |
apply (erule principal_mono [OF take_mono])
|
|
293 |
apply (rule principal_mono [OF take_chain])
|
|
294 |
done
|
|
295 |
|
|
296 |
lemma lub_basis_fun_take:
|
|
297 |
"(\<Squnion>i. basis_fun (\<lambda>a. principal (take i a))\<cdot>x) = x"
|
|
298 |
apply (simp add: basis_fun_beta principal_mono take_mono)
|
|
299 |
apply (subst image_image [where f=principal, symmetric])
|
|
300 |
apply (rule unique_lub [OF _ lub_principal_approxes])
|
|
301 |
apply (rule basis_fun_lemma2, erule principal_mono)
|
|
302 |
done
|
|
303 |
|
|
304 |
lemma finite_directed_contains_lub:
|
|
305 |
"\<lbrakk>finite S; directed S\<rbrakk> \<Longrightarrow> \<exists>u\<in>S. S <<| u"
|
|
306 |
apply (drule (1) directed_finiteD, rule subset_refl)
|
|
307 |
apply (erule bexE)
|
|
308 |
apply (rule rev_bexI, assumption)
|
|
309 |
apply (erule (1) is_lub_maximal)
|
|
310 |
done
|
|
311 |
|
|
312 |
lemma lub_finite_directed_in_self:
|
|
313 |
"\<lbrakk>finite S; directed S\<rbrakk> \<Longrightarrow> lub S \<in> S"
|
|
314 |
apply (drule (1) directed_finiteD, rule subset_refl)
|
|
315 |
apply (erule bexE)
|
|
316 |
apply (drule (1) is_lub_maximal)
|
|
317 |
apply (drule thelubI)
|
|
318 |
apply simp
|
|
319 |
done
|
|
320 |
|
|
321 |
lemma basis_fun_take_eq_principal:
|
|
322 |
"\<exists>a\<in>approxes x.
|
|
323 |
basis_fun (\<lambda>a. principal (take i a))\<cdot>x = principal (take i a)"
|
|
324 |
apply (simp add: basis_fun_beta principal_mono take_mono)
|
|
325 |
apply (subst image_image [where f=principal, symmetric])
|
|
326 |
apply (subgoal_tac "finite (principal ` take i ` approxes x)")
|
|
327 |
apply (subgoal_tac "directed (principal ` take i ` approxes x)")
|
|
328 |
apply (drule (1) lub_finite_directed_in_self, fast)
|
|
329 |
apply (subst image_image)
|
|
330 |
apply (rule directed_image_ideal)
|
|
331 |
apply (rule ideal_approxes)
|
|
332 |
apply (erule principal_mono [OF take_mono])
|
|
333 |
apply (rule finite_imageI)
|
|
334 |
apply (rule finite_take_approxes)
|
|
335 |
done
|
|
336 |
|
|
337 |
lemma principal_induct:
|
|
338 |
assumes adm: "adm P"
|
|
339 |
assumes P: "\<And>a. P (principal a)"
|
|
340 |
shows "P x"
|
|
341 |
apply (subgoal_tac "P (\<Squnion>i. basis_fun (\<lambda>a. principal (take i a))\<cdot>x)")
|
|
342 |
apply (simp add: lub_basis_fun_take)
|
|
343 |
apply (rule admD [rule_format, OF adm])
|
|
344 |
apply (simp add: chain_basis_fun_take)
|
|
345 |
apply (cut_tac x=x and i=i in basis_fun_take_eq_principal)
|
|
346 |
apply (clarify, simp add: P)
|
|
347 |
done
|
|
348 |
|
|
349 |
lemma finite_fixes_basis_fun_take:
|
|
350 |
"finite {x. basis_fun (\<lambda>a. principal (take i a))\<cdot>x = x}" (is "finite ?S")
|
|
351 |
apply (subgoal_tac "?S \<subseteq> principal ` range (take i)")
|
|
352 |
apply (erule finite_subset)
|
|
353 |
apply (rule finite_imageI)
|
|
354 |
apply (rule finite_range_take)
|
|
355 |
apply (clarify, erule subst)
|
|
356 |
apply (cut_tac x=x and i=i in basis_fun_take_eq_principal)
|
|
357 |
apply fast
|
|
358 |
done
|
|
359 |
|
|
360 |
end
|
|
361 |
|
|
362 |
|
|
363 |
subsection {* Compact bases of bifinite domains *}
|
|
364 |
|
|
365 |
defaultsort bifinite
|
|
366 |
|
|
367 |
typedef (open) 'a compact_basis = "{x::'a::bifinite. compact x}"
|
|
368 |
by (fast intro: compact_approx)
|
|
369 |
|
|
370 |
lemma compact_Rep_compact_basis [simp]: "compact (Rep_compact_basis a)"
|
|
371 |
by (rule Rep_compact_basis [simplified])
|
|
372 |
|
|
373 |
lemma Rep_Abs_compact_basis_approx [simp]:
|
|
374 |
"Rep_compact_basis (Abs_compact_basis (approx n\<cdot>x)) = approx n\<cdot>x"
|
|
375 |
by (rule Abs_compact_basis_inverse, simp)
|
|
376 |
|
|
377 |
lemma compact_imp_Rep_compact_basis:
|
|
378 |
"compact x \<Longrightarrow> \<exists>y. x = Rep_compact_basis y"
|
|
379 |
by (rule exI, rule Abs_compact_basis_inverse [symmetric], simp)
|
|
380 |
|
|
381 |
definition
|
|
382 |
compact_le :: "'a compact_basis \<Rightarrow> 'a compact_basis \<Rightarrow> bool" where
|
|
383 |
"compact_le = (\<lambda>x y. Rep_compact_basis x \<sqsubseteq> Rep_compact_basis y)"
|
|
384 |
|
|
385 |
lemma compact_le_refl: "compact_le x x"
|
|
386 |
unfolding compact_le_def by (rule refl_less)
|
|
387 |
|
|
388 |
lemma compact_le_trans: "\<lbrakk>compact_le x y; compact_le y z\<rbrakk> \<Longrightarrow> compact_le x z"
|
|
389 |
unfolding compact_le_def by (rule trans_less)
|
|
390 |
|
|
391 |
lemma compact_le_antisym: "\<lbrakk>compact_le x y; compact_le y x\<rbrakk> \<Longrightarrow> x = y"
|
|
392 |
unfolding compact_le_def
|
|
393 |
apply (rule Rep_compact_basis_inject [THEN iffD1])
|
|
394 |
apply (erule (1) antisym_less)
|
|
395 |
done
|
|
396 |
|
|
397 |
interpretation compact_le: preorder [compact_le]
|
|
398 |
by (rule preorder.intro, rule compact_le_refl, rule compact_le_trans)
|
|
399 |
|
|
400 |
text {* minimal compact element *}
|
|
401 |
|
|
402 |
definition
|
|
403 |
compact_bot :: "'a compact_basis" where
|
|
404 |
"compact_bot = Abs_compact_basis \<bottom>"
|
|
405 |
|
|
406 |
lemma Rep_compact_bot: "Rep_compact_basis compact_bot = \<bottom>"
|
|
407 |
unfolding compact_bot_def by (simp add: Abs_compact_basis_inverse)
|
|
408 |
|
|
409 |
lemma compact_minimal [simp]: "compact_le compact_bot a"
|
|
410 |
unfolding compact_le_def Rep_compact_bot by simp
|
|
411 |
|
|
412 |
text {* compacts *}
|
|
413 |
|
|
414 |
definition
|
|
415 |
compacts :: "'a \<Rightarrow> 'a compact_basis set" where
|
|
416 |
"compacts = (\<lambda>x. {a. Rep_compact_basis a \<sqsubseteq> x})"
|
|
417 |
|
|
418 |
lemma ideal_compacts: "compact_le.ideal (compacts w)"
|
|
419 |
unfolding compacts_def
|
|
420 |
apply (rule compact_le.idealI)
|
|
421 |
apply (rule_tac x="Abs_compact_basis (approx 0\<cdot>w)" in exI)
|
|
422 |
apply (simp add: approx_less)
|
|
423 |
apply simp
|
|
424 |
apply (cut_tac a=x in compact_Rep_compact_basis)
|
|
425 |
apply (cut_tac a=y in compact_Rep_compact_basis)
|
|
426 |
apply (drule bifinite_compact_eq_approx)
|
|
427 |
apply (drule bifinite_compact_eq_approx)
|
|
428 |
apply (clarify, rename_tac i j)
|
|
429 |
apply (rule_tac x="Abs_compact_basis (approx (max i j)\<cdot>w)" in exI)
|
|
430 |
apply (simp add: approx_less compact_le_def)
|
|
431 |
apply (erule subst, erule subst)
|
|
432 |
apply (simp add: monofun_cfun chain_mono3 [OF chain_approx])
|
|
433 |
apply (simp add: compact_le_def)
|
|
434 |
apply (erule (1) trans_less)
|
|
435 |
done
|
|
436 |
|
|
437 |
lemma compacts_Rep_compact_basis:
|
|
438 |
"compacts (Rep_compact_basis b) = {a. compact_le a b}"
|
|
439 |
unfolding compacts_def compact_le_def ..
|
|
440 |
|
|
441 |
lemma cont_compacts: "cont compacts"
|
|
442 |
unfolding compacts_def
|
|
443 |
apply (rule contI2)
|
|
444 |
apply (rule monofunI)
|
|
445 |
apply (simp add: set_cpo_simps)
|
|
446 |
apply (fast intro: trans_less)
|
|
447 |
apply (simp add: set_cpo_simps)
|
|
448 |
apply clarify
|
|
449 |
apply simp
|
|
450 |
apply (erule (1) compactD2 [OF compact_Rep_compact_basis])
|
|
451 |
done
|
|
452 |
|
|
453 |
lemma compacts_lessD: "compacts x \<subseteq> compacts y \<Longrightarrow> x \<sqsubseteq> y"
|
|
454 |
apply (subgoal_tac "(\<Squnion>i. approx i\<cdot>x) \<sqsubseteq> y", simp)
|
|
455 |
apply (rule admD [rule_format], simp, simp)
|
|
456 |
apply (drule_tac c="Abs_compact_basis (approx i\<cdot>x)" in subsetD)
|
|
457 |
apply (simp add: compacts_def Abs_compact_basis_inverse approx_less)
|
|
458 |
apply (simp add: compacts_def Abs_compact_basis_inverse)
|
|
459 |
done
|
|
460 |
|
|
461 |
lemma compacts_mono: "x \<sqsubseteq> y \<Longrightarrow> compacts x \<subseteq> compacts y"
|
|
462 |
unfolding compacts_def by (fast intro: trans_less)
|
|
463 |
|
|
464 |
lemma less_compact_basis_iff: "(x \<sqsubseteq> y) = (compacts x \<subseteq> compacts y)"
|
|
465 |
by (rule iffI [OF compacts_mono compacts_lessD])
|
|
466 |
|
|
467 |
lemma compact_basis_induct:
|
|
468 |
"\<lbrakk>adm P; \<And>a. P (Rep_compact_basis a)\<rbrakk> \<Longrightarrow> P x"
|
|
469 |
apply (erule approx_induct)
|
|
470 |
apply (drule_tac x="Abs_compact_basis (approx n\<cdot>x)" in meta_spec)
|
|
471 |
apply (simp add: Abs_compact_basis_inverse)
|
|
472 |
done
|
|
473 |
|
|
474 |
text {* approximation on compact bases *}
|
|
475 |
|
|
476 |
definition
|
|
477 |
compact_approx :: "nat \<Rightarrow> 'a compact_basis \<Rightarrow> 'a compact_basis" where
|
|
478 |
"compact_approx = (\<lambda>n a. Abs_compact_basis (approx n\<cdot>(Rep_compact_basis a)))"
|
|
479 |
|
|
480 |
lemma Rep_compact_approx:
|
|
481 |
"Rep_compact_basis (compact_approx n a) = approx n\<cdot>(Rep_compact_basis a)"
|
|
482 |
unfolding compact_approx_def
|
|
483 |
by (simp add: Abs_compact_basis_inverse)
|
|
484 |
|
|
485 |
lemmas approx_Rep_compact_basis = Rep_compact_approx [symmetric]
|
|
486 |
|
|
487 |
lemma compact_approx_le:
|
|
488 |
"compact_le (compact_approx n a) a"
|
|
489 |
unfolding compact_le_def
|
|
490 |
by (simp add: Rep_compact_approx approx_less)
|
|
491 |
|
|
492 |
lemma compact_approx_mono1:
|
|
493 |
"i \<le> j \<Longrightarrow> compact_le (compact_approx i a) (compact_approx j a)"
|
|
494 |
unfolding compact_le_def
|
|
495 |
apply (simp add: Rep_compact_approx)
|
|
496 |
apply (rule chain_mono3, simp, assumption)
|
|
497 |
done
|
|
498 |
|
|
499 |
lemma compact_approx_mono:
|
|
500 |
"compact_le a b \<Longrightarrow> compact_le (compact_approx n a) (compact_approx n b)"
|
|
501 |
unfolding compact_le_def
|
|
502 |
apply (simp add: Rep_compact_approx)
|
|
503 |
apply (erule monofun_cfun_arg)
|
|
504 |
done
|
|
505 |
|
|
506 |
lemma ex_compact_approx_eq: "\<exists>n. compact_approx n a = a"
|
|
507 |
apply (simp add: Rep_compact_basis_inject [symmetric])
|
|
508 |
apply (simp add: Rep_compact_approx)
|
|
509 |
apply (rule bifinite_compact_eq_approx)
|
|
510 |
apply (rule compact_Rep_compact_basis)
|
|
511 |
done
|
|
512 |
|
|
513 |
lemma compact_approx_idem:
|
|
514 |
"compact_approx n (compact_approx n a) = compact_approx n a"
|
|
515 |
apply (rule Rep_compact_basis_inject [THEN iffD1])
|
|
516 |
apply (simp add: Rep_compact_approx)
|
|
517 |
done
|
|
518 |
|
|
519 |
lemma finite_fixes_compact_approx: "finite {a. compact_approx n a = a}"
|
|
520 |
apply (subgoal_tac "finite (Rep_compact_basis ` {a. compact_approx n a = a})")
|
|
521 |
apply (erule finite_imageD)
|
|
522 |
apply (rule inj_onI, simp add: Rep_compact_basis_inject)
|
|
523 |
apply (rule finite_subset [OF _ finite_fixes_approx [where i=n]])
|
|
524 |
apply (rule subsetI, clarify, rename_tac a)
|
|
525 |
apply (simp add: Rep_compact_basis_inject [symmetric])
|
|
526 |
apply (simp add: Rep_compact_approx)
|
|
527 |
done
|
|
528 |
|
|
529 |
lemma finite_range_compact_approx: "finite (range (compact_approx n))"
|
|
530 |
apply (cut_tac n=n in finite_fixes_compact_approx)
|
|
531 |
apply (simp add: idem_fixes_eq_range compact_approx_idem)
|
|
532 |
apply (simp add: image_def)
|
|
533 |
done
|
|
534 |
|
|
535 |
interpretation compact_basis:
|
|
536 |
bifinite_basis [compact_le Rep_compact_basis compacts compact_approx]
|
|
537 |
apply unfold_locales
|
|
538 |
apply (rule ideal_compacts)
|
|
539 |
apply (rule cont_compacts)
|
|
540 |
apply (rule compacts_Rep_compact_basis)
|
|
541 |
apply (erule compacts_lessD)
|
|
542 |
apply (rule compact_approx_le)
|
|
543 |
apply (rule compact_approx_idem)
|
|
544 |
apply (erule compact_approx_mono)
|
|
545 |
apply (rule compact_approx_mono1, simp)
|
|
546 |
apply (rule finite_range_compact_approx)
|
|
547 |
apply (rule ex_compact_approx_eq)
|
|
548 |
done
|
|
549 |
|
|
550 |
|
|
551 |
subsection {* A compact basis for powerdomains *}
|
|
552 |
|
|
553 |
typedef 'a pd_basis =
|
|
554 |
"{S::'a::bifinite compact_basis set. finite S \<and> S \<noteq> {}}"
|
|
555 |
by (rule_tac x="{arbitrary}" in exI, simp)
|
|
556 |
|
|
557 |
lemma finite_Rep_pd_basis [simp]: "finite (Rep_pd_basis u)"
|
|
558 |
by (insert Rep_pd_basis [of u, unfolded pd_basis_def]) simp
|
|
559 |
|
|
560 |
lemma Rep_pd_basis_nonempty [simp]: "Rep_pd_basis u \<noteq> {}"
|
|
561 |
by (insert Rep_pd_basis [of u, unfolded pd_basis_def]) simp
|
|
562 |
|
|
563 |
text {* unit and plus *}
|
|
564 |
|
|
565 |
definition
|
|
566 |
PDUnit :: "'a compact_basis \<Rightarrow> 'a pd_basis" where
|
|
567 |
"PDUnit = (\<lambda>x. Abs_pd_basis {x})"
|
|
568 |
|
|
569 |
definition
|
|
570 |
PDPlus :: "'a pd_basis \<Rightarrow> 'a pd_basis \<Rightarrow> 'a pd_basis" where
|
|
571 |
"PDPlus t u = Abs_pd_basis (Rep_pd_basis t \<union> Rep_pd_basis u)"
|
|
572 |
|
|
573 |
lemma Rep_PDUnit:
|
|
574 |
"Rep_pd_basis (PDUnit x) = {x}"
|
|
575 |
unfolding PDUnit_def by (rule Abs_pd_basis_inverse) (simp add: pd_basis_def)
|
|
576 |
|
|
577 |
lemma Rep_PDPlus:
|
|
578 |
"Rep_pd_basis (PDPlus u v) = Rep_pd_basis u \<union> Rep_pd_basis v"
|
|
579 |
unfolding PDPlus_def by (rule Abs_pd_basis_inverse) (simp add: pd_basis_def)
|
|
580 |
|
|
581 |
lemma PDUnit_inject [simp]: "(PDUnit a = PDUnit b) = (a = b)"
|
|
582 |
unfolding Rep_pd_basis_inject [symmetric] Rep_PDUnit by simp
|
|
583 |
|
|
584 |
lemma PDPlus_assoc: "PDPlus (PDPlus t u) v = PDPlus t (PDPlus u v)"
|
|
585 |
unfolding Rep_pd_basis_inject [symmetric] Rep_PDPlus by (rule Un_assoc)
|
|
586 |
|
|
587 |
lemma PDPlus_commute: "PDPlus t u = PDPlus u t"
|
|
588 |
unfolding Rep_pd_basis_inject [symmetric] Rep_PDPlus by (rule Un_commute)
|
|
589 |
|
|
590 |
lemma PDPlus_absorb: "PDPlus t t = t"
|
|
591 |
unfolding Rep_pd_basis_inject [symmetric] Rep_PDPlus by (rule Un_absorb)
|
|
592 |
|
|
593 |
lemma pd_basis_induct1:
|
|
594 |
assumes PDUnit: "\<And>a. P (PDUnit a)"
|
|
595 |
assumes PDPlus: "\<And>a t. P t \<Longrightarrow> P (PDPlus (PDUnit a) t)"
|
|
596 |
shows "P x"
|
|
597 |
apply (induct x, unfold pd_basis_def, clarify)
|
|
598 |
apply (erule (1) finite_ne_induct)
|
|
599 |
apply (cut_tac a=x in PDUnit)
|
|
600 |
apply (simp add: PDUnit_def)
|
|
601 |
apply (drule_tac a=x in PDPlus)
|
|
602 |
apply (simp add: PDUnit_def PDPlus_def Abs_pd_basis_inverse [unfolded pd_basis_def])
|
|
603 |
done
|
|
604 |
|
|
605 |
lemma pd_basis_induct:
|
|
606 |
assumes PDUnit: "\<And>a. P (PDUnit a)"
|
|
607 |
assumes PDPlus: "\<And>t u. \<lbrakk>P t; P u\<rbrakk> \<Longrightarrow> P (PDPlus t u)"
|
|
608 |
shows "P x"
|
|
609 |
apply (induct x rule: pd_basis_induct1)
|
|
610 |
apply (rule PDUnit, erule PDPlus [OF PDUnit])
|
|
611 |
done
|
|
612 |
|
|
613 |
text {* fold-pd *}
|
|
614 |
|
|
615 |
definition
|
|
616 |
fold_pd ::
|
|
617 |
"('a compact_basis \<Rightarrow> 'b::type) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a pd_basis \<Rightarrow> 'b"
|
|
618 |
where "fold_pd g f t = fold1 f (g ` Rep_pd_basis t)"
|
|
619 |
|
|
620 |
lemma (in ACIf) fold_pd_PDUnit:
|
|
621 |
"fold_pd g f (PDUnit x) = g x"
|
|
622 |
unfolding fold_pd_def Rep_PDUnit by simp
|
|
623 |
|
|
624 |
lemma (in ACIf) fold_pd_PDPlus:
|
|
625 |
"fold_pd g f (PDPlus t u) = f (fold_pd g f t) (fold_pd g f u)"
|
|
626 |
unfolding fold_pd_def Rep_PDPlus by (simp add: image_Un fold1_Un2)
|
|
627 |
|
|
628 |
text {* approx-pd *}
|
|
629 |
|
|
630 |
definition
|
|
631 |
approx_pd :: "nat \<Rightarrow> 'a pd_basis \<Rightarrow> 'a pd_basis" where
|
|
632 |
"approx_pd n = (\<lambda>t. Abs_pd_basis (compact_approx n ` Rep_pd_basis t))"
|
|
633 |
|
|
634 |
lemma Rep_approx_pd:
|
|
635 |
"Rep_pd_basis (approx_pd n t) = compact_approx n ` Rep_pd_basis t"
|
|
636 |
unfolding approx_pd_def
|
|
637 |
apply (rule Abs_pd_basis_inverse)
|
|
638 |
apply (simp add: pd_basis_def)
|
|
639 |
done
|
|
640 |
|
|
641 |
lemma approx_pd_simps [simp]:
|
|
642 |
"approx_pd n (PDUnit a) = PDUnit (compact_approx n a)"
|
|
643 |
"approx_pd n (PDPlus t u) = PDPlus (approx_pd n t) (approx_pd n u)"
|
|
644 |
apply (simp_all add: Rep_pd_basis_inject [symmetric])
|
|
645 |
apply (simp_all add: Rep_approx_pd Rep_PDUnit Rep_PDPlus image_Un)
|
|
646 |
done
|
|
647 |
|
|
648 |
lemma approx_pd_idem: "approx_pd n (approx_pd n t) = approx_pd n t"
|
|
649 |
apply (induct t rule: pd_basis_induct)
|
|
650 |
apply (simp add: compact_approx_idem)
|
|
651 |
apply simp
|
|
652 |
done
|
|
653 |
|
|
654 |
lemma range_image_f: "range (image f) = Pow (range f)"
|
|
655 |
apply (safe, fast)
|
|
656 |
apply (rule_tac x="f -` x" in range_eqI)
|
|
657 |
apply (simp, fast)
|
|
658 |
done
|
|
659 |
|
|
660 |
lemma finite_range_approx_pd: "finite (range (approx_pd n))"
|
|
661 |
apply (subgoal_tac "finite (Rep_pd_basis ` range (approx_pd n))")
|
|
662 |
apply (erule finite_imageD)
|
|
663 |
apply (rule inj_onI, simp add: Rep_pd_basis_inject)
|
|
664 |
apply (subst image_image)
|
|
665 |
apply (subst Rep_approx_pd)
|
|
666 |
apply (simp only: range_composition)
|
|
667 |
apply (rule finite_subset [OF image_mono [OF subset_UNIV]])
|
|
668 |
apply (simp add: range_image_f)
|
|
669 |
apply (rule finite_range_compact_approx)
|
|
670 |
done
|
|
671 |
|
|
672 |
lemma ex_approx_pd_eq: "\<exists>n. approx_pd n t = t"
|
|
673 |
apply (subgoal_tac "\<exists>n. \<forall>m\<ge>n. approx_pd m t = t", fast)
|
|
674 |
apply (induct t rule: pd_basis_induct)
|
|
675 |
apply (cut_tac a=a in ex_compact_approx_eq)
|
|
676 |
apply (clarify, rule_tac x=n in exI)
|
|
677 |
apply (clarify, simp)
|
|
678 |
apply (rule compact_le_antisym)
|
|
679 |
apply (rule compact_approx_le)
|
|
680 |
apply (drule_tac a=a in compact_approx_mono1)
|
|
681 |
apply simp
|
|
682 |
apply (clarify, rename_tac i j)
|
|
683 |
apply (rule_tac x="max i j" in exI, simp)
|
|
684 |
done
|
|
685 |
|
|
686 |
end
|