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