author | huffman |
Thu, 12 Jun 2008 22:41:03 +0200 | |
changeset 27186 | 416d66c36d8f |
parent 26962 | c8b20f615d6c |
child 27309 | c74270fd72a8 |
permissions | -rw-r--r-- |
25903 | 1 |
(* Title: HOLCF/Bifinite.thy |
2 |
ID: $Id$ |
|
3 |
Author: Brian Huffman |
|
4 |
*) |
|
5 |
||
6 |
header {* Bifinite domains and approximation *} |
|
7 |
||
8 |
theory Bifinite |
|
9 |
imports Cfun |
|
10 |
begin |
|
11 |
||
26407
562a1d615336
rename class bifinite_cpo to profinite; generalize powerdomains from bifinite to profinite
huffman
parents:
25923
diff
changeset
|
12 |
subsection {* Omega-profinite and bifinite domains *} |
25903 | 13 |
|
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26407
diff
changeset
|
14 |
class profinite = cpo + |
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26407
diff
changeset
|
15 |
fixes approx :: "nat \<Rightarrow> 'a \<rightarrow> 'a" |
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26407
diff
changeset
|
16 |
assumes chain_approx_app: "chain (\<lambda>i. approx i\<cdot>x)" |
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26407
diff
changeset
|
17 |
assumes lub_approx_app [simp]: "(\<Squnion>i. approx i\<cdot>x) = x" |
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26407
diff
changeset
|
18 |
assumes approx_idem: "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x" |
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26407
diff
changeset
|
19 |
assumes finite_fixes_approx: "finite {x. approx i\<cdot>x = x}" |
25903 | 20 |
|
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26407
diff
changeset
|
21 |
class bifinite = profinite + pcpo |
25909
6b96b9392873
add class bifinite_cpo for possibly-unpointed bifinite domains
huffman
parents:
25903
diff
changeset
|
22 |
|
25903 | 23 |
lemma finite_range_imp_finite_fixes: |
24 |
"finite {x. \<exists>y. x = f y} \<Longrightarrow> finite {x. f x = x}" |
|
25 |
apply (subgoal_tac "{x. f x = x} \<subseteq> {x. \<exists>y. x = f y}") |
|
26 |
apply (erule (1) finite_subset) |
|
27 |
apply (clarify, erule subst, rule exI, rule refl) |
|
28 |
done |
|
29 |
||
27186
416d66c36d8f
add lemma finite_image_approx; remove unnecessary sort annotations
huffman
parents:
26962
diff
changeset
|
30 |
lemma chain_approx [simp]: "chain approx" |
25903 | 31 |
apply (rule chainI) |
32 |
apply (rule less_cfun_ext) |
|
33 |
apply (rule chainE) |
|
34 |
apply (rule chain_approx_app) |
|
35 |
done |
|
36 |
||
27186
416d66c36d8f
add lemma finite_image_approx; remove unnecessary sort annotations
huffman
parents:
26962
diff
changeset
|
37 |
lemma lub_approx [simp]: "(\<Squnion>i. approx i) = (\<Lambda> x. x)" |
25903 | 38 |
by (rule ext_cfun, simp add: contlub_cfun_fun) |
39 |
||
27186
416d66c36d8f
add lemma finite_image_approx; remove unnecessary sort annotations
huffman
parents:
26962
diff
changeset
|
40 |
lemma approx_less: "approx i\<cdot>x \<sqsubseteq> x" |
25903 | 41 |
apply (subgoal_tac "approx i\<cdot>x \<sqsubseteq> (\<Squnion>i. approx i\<cdot>x)", simp) |
42 |
apply (rule is_ub_thelub, simp) |
|
43 |
done |
|
44 |
||
45 |
lemma approx_strict [simp]: "approx i\<cdot>(\<bottom>::'a::bifinite) = \<bottom>" |
|
46 |
by (rule UU_I, rule approx_less) |
|
47 |
||
48 |
lemma approx_approx1: |
|
27186
416d66c36d8f
add lemma finite_image_approx; remove unnecessary sort annotations
huffman
parents:
26962
diff
changeset
|
49 |
"i \<le> j \<Longrightarrow> approx i\<cdot>(approx j\<cdot>x) = approx i\<cdot>x" |
25903 | 50 |
apply (rule antisym_less) |
51 |
apply (rule monofun_cfun_arg [OF approx_less]) |
|
52 |
apply (rule sq_ord_eq_less_trans [OF approx_idem [symmetric]]) |
|
53 |
apply (rule monofun_cfun_arg) |
|
54 |
apply (rule monofun_cfun_fun) |
|
25922
cb04d05e95fb
rename lemma chain_mono3 -> chain_mono, chain_mono -> chain_mono_less
huffman
parents:
25909
diff
changeset
|
55 |
apply (erule chain_mono [OF chain_approx]) |
25903 | 56 |
done |
57 |
||
58 |
lemma approx_approx2: |
|
27186
416d66c36d8f
add lemma finite_image_approx; remove unnecessary sort annotations
huffman
parents:
26962
diff
changeset
|
59 |
"j \<le> i \<Longrightarrow> approx i\<cdot>(approx j\<cdot>x) = approx j\<cdot>x" |
25903 | 60 |
apply (rule antisym_less) |
61 |
apply (rule approx_less) |
|
62 |
apply (rule sq_ord_eq_less_trans [OF approx_idem [symmetric]]) |
|
63 |
apply (rule monofun_cfun_fun) |
|
25922
cb04d05e95fb
rename lemma chain_mono3 -> chain_mono, chain_mono -> chain_mono_less
huffman
parents:
25909
diff
changeset
|
64 |
apply (erule chain_mono [OF chain_approx]) |
25903 | 65 |
done |
66 |
||
67 |
lemma approx_approx [simp]: |
|
27186
416d66c36d8f
add lemma finite_image_approx; remove unnecessary sort annotations
huffman
parents:
26962
diff
changeset
|
68 |
"approx i\<cdot>(approx j\<cdot>x) = approx (min i j)\<cdot>x" |
25903 | 69 |
apply (rule_tac x=i and y=j in linorder_le_cases) |
70 |
apply (simp add: approx_approx1 min_def) |
|
71 |
apply (simp add: approx_approx2 min_def) |
|
72 |
done |
|
73 |
||
74 |
lemma idem_fixes_eq_range: |
|
75 |
"\<forall>x. f (f x) = f x \<Longrightarrow> {x. f x = x} = {y. \<exists>x. y = f x}" |
|
76 |
by (auto simp add: eq_sym_conv) |
|
77 |
||
27186
416d66c36d8f
add lemma finite_image_approx; remove unnecessary sort annotations
huffman
parents:
26962
diff
changeset
|
78 |
lemma finite_approx: "finite {y. \<exists>x. y = approx n\<cdot>x}" |
25903 | 79 |
using finite_fixes_approx by (simp add: idem_fixes_eq_range) |
80 |
||
27186
416d66c36d8f
add lemma finite_image_approx; remove unnecessary sort annotations
huffman
parents:
26962
diff
changeset
|
81 |
lemma finite_image_approx: "finite ((\<lambda>x. approx n\<cdot>x) ` A)" |
416d66c36d8f
add lemma finite_image_approx; remove unnecessary sort annotations
huffman
parents:
26962
diff
changeset
|
82 |
by (rule finite_subset [OF _ finite_fixes_approx [where i=n]]) auto |
25903 | 83 |
|
27186
416d66c36d8f
add lemma finite_image_approx; remove unnecessary sort annotations
huffman
parents:
26962
diff
changeset
|
84 |
lemma finite_range_approx: "finite (range (\<lambda>x. approx n\<cdot>x))" |
416d66c36d8f
add lemma finite_image_approx; remove unnecessary sort annotations
huffman
parents:
26962
diff
changeset
|
85 |
by (rule finite_image_approx) |
416d66c36d8f
add lemma finite_image_approx; remove unnecessary sort annotations
huffman
parents:
26962
diff
changeset
|
86 |
|
416d66c36d8f
add lemma finite_image_approx; remove unnecessary sort annotations
huffman
parents:
26962
diff
changeset
|
87 |
lemma compact_approx [simp]: "compact (approx n\<cdot>x)" |
25903 | 88 |
proof (rule compactI2) |
89 |
fix Y::"nat \<Rightarrow> 'a" |
|
90 |
assume Y: "chain Y" |
|
91 |
have "finite_chain (\<lambda>i. approx n\<cdot>(Y i))" |
|
92 |
proof (rule finite_range_imp_finch) |
|
93 |
show "chain (\<lambda>i. approx n\<cdot>(Y i))" |
|
94 |
using Y by simp |
|
95 |
have "range (\<lambda>i. approx n\<cdot>(Y i)) \<subseteq> {x. approx n\<cdot>x = x}" |
|
96 |
by clarsimp |
|
97 |
thus "finite (range (\<lambda>i. approx n\<cdot>(Y i)))" |
|
98 |
using finite_fixes_approx by (rule finite_subset) |
|
99 |
qed |
|
100 |
hence "\<exists>j. (\<Squnion>i. approx n\<cdot>(Y i)) = approx n\<cdot>(Y j)" |
|
101 |
by (simp add: finite_chain_def maxinch_is_thelub Y) |
|
102 |
then obtain j where j: "(\<Squnion>i. approx n\<cdot>(Y i)) = approx n\<cdot>(Y j)" .. |
|
103 |
||
104 |
assume "approx n\<cdot>x \<sqsubseteq> (\<Squnion>i. Y i)" |
|
105 |
hence "approx n\<cdot>(approx n\<cdot>x) \<sqsubseteq> approx n\<cdot>(\<Squnion>i. Y i)" |
|
106 |
by (rule monofun_cfun_arg) |
|
107 |
hence "approx n\<cdot>x \<sqsubseteq> (\<Squnion>i. approx n\<cdot>(Y i))" |
|
108 |
by (simp add: contlub_cfun_arg Y) |
|
109 |
hence "approx n\<cdot>x \<sqsubseteq> approx n\<cdot>(Y j)" |
|
110 |
using j by simp |
|
111 |
hence "approx n\<cdot>x \<sqsubseteq> Y j" |
|
112 |
using approx_less by (rule trans_less) |
|
113 |
thus "\<exists>j. approx n\<cdot>x \<sqsubseteq> Y j" .. |
|
114 |
qed |
|
115 |
||
116 |
lemma bifinite_compact_eq_approx: |
|
117 |
assumes x: "compact x" |
|
118 |
shows "\<exists>i. approx i\<cdot>x = x" |
|
119 |
proof - |
|
120 |
have chain: "chain (\<lambda>i. approx i\<cdot>x)" by simp |
|
121 |
have less: "x \<sqsubseteq> (\<Squnion>i. approx i\<cdot>x)" by simp |
|
122 |
obtain i where i: "x \<sqsubseteq> approx i\<cdot>x" |
|
123 |
using compactD2 [OF x chain less] .. |
|
124 |
with approx_less have "approx i\<cdot>x = x" |
|
125 |
by (rule antisym_less) |
|
126 |
thus "\<exists>i. approx i\<cdot>x = x" .. |
|
127 |
qed |
|
128 |
||
129 |
lemma bifinite_compact_iff: |
|
27186
416d66c36d8f
add lemma finite_image_approx; remove unnecessary sort annotations
huffman
parents:
26962
diff
changeset
|
130 |
"compact x \<longleftrightarrow> (\<exists>n. approx n\<cdot>x = x)" |
25903 | 131 |
apply (rule iffI) |
132 |
apply (erule bifinite_compact_eq_approx) |
|
133 |
apply (erule exE) |
|
134 |
apply (erule subst) |
|
135 |
apply (rule compact_approx) |
|
136 |
done |
|
137 |
||
138 |
lemma approx_induct: |
|
139 |
assumes adm: "adm P" and P: "\<And>n x. P (approx n\<cdot>x)" |
|
27186
416d66c36d8f
add lemma finite_image_approx; remove unnecessary sort annotations
huffman
parents:
26962
diff
changeset
|
140 |
shows "P x" |
25903 | 141 |
proof - |
142 |
have "P (\<Squnion>n. approx n\<cdot>x)" |
|
143 |
by (rule admD [OF adm], simp, simp add: P) |
|
144 |
thus "P x" by simp |
|
145 |
qed |
|
146 |
||
27186
416d66c36d8f
add lemma finite_image_approx; remove unnecessary sort annotations
huffman
parents:
26962
diff
changeset
|
147 |
lemma bifinite_less_ext: "(\<And>i. approx i\<cdot>x \<sqsubseteq> approx i\<cdot>y) \<Longrightarrow> x \<sqsubseteq> y" |
25903 | 148 |
apply (subgoal_tac "(\<Squnion>i. approx i\<cdot>x) \<sqsubseteq> (\<Squnion>i. approx i\<cdot>y)", simp) |
25923 | 149 |
apply (rule lub_mono, simp, simp, simp) |
25903 | 150 |
done |
151 |
||
152 |
subsection {* Instance for continuous function space *} |
|
153 |
||
154 |
lemma finite_range_lemma: |
|
155 |
fixes h :: "'a::cpo \<rightarrow> 'b::cpo" |
|
156 |
fixes k :: "'c::cpo \<rightarrow> 'd::cpo" |
|
157 |
shows "\<lbrakk>finite {y. \<exists>x. y = h\<cdot>x}; finite {y. \<exists>x. y = k\<cdot>x}\<rbrakk> |
|
158 |
\<Longrightarrow> finite {g. \<exists>f. g = (\<Lambda> x. k\<cdot>(f\<cdot>(h\<cdot>x)))}" |
|
159 |
apply (rule_tac f="\<lambda>g. {(h\<cdot>x, y) |x y. y = g\<cdot>x}" in finite_imageD) |
|
160 |
apply (rule_tac B="Pow ({y. \<exists>x. y = h\<cdot>x} \<times> {y. \<exists>x. y = k\<cdot>x})" |
|
161 |
in finite_subset) |
|
162 |
apply (rule image_subsetI) |
|
163 |
apply (clarsimp, fast) |
|
164 |
apply simp |
|
165 |
apply (rule inj_onI) |
|
166 |
apply (clarsimp simp add: expand_set_eq) |
|
167 |
apply (rule ext_cfun, simp) |
|
168 |
apply (drule_tac x="h\<cdot>x" in spec) |
|
169 |
apply (drule_tac x="k\<cdot>(f\<cdot>(h\<cdot>x))" in spec) |
|
170 |
apply (drule iffD1, fast) |
|
171 |
apply clarsimp |
|
172 |
done |
|
173 |
||
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26407
diff
changeset
|
174 |
instantiation "->" :: (profinite, profinite) profinite |
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26407
diff
changeset
|
175 |
begin |
25903 | 176 |
|
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26407
diff
changeset
|
177 |
definition |
25903 | 178 |
approx_cfun_def: |
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26407
diff
changeset
|
179 |
"approx = (\<lambda>n. \<Lambda> f x. approx n\<cdot>(f\<cdot>(approx n\<cdot>x)))" |
25903 | 180 |
|
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26407
diff
changeset
|
181 |
instance |
25903 | 182 |
apply (intro_classes, unfold approx_cfun_def) |
183 |
apply simp |
|
184 |
apply (simp add: lub_distribs eta_cfun) |
|
185 |
apply simp |
|
186 |
apply simp |
|
187 |
apply (rule finite_range_imp_finite_fixes) |
|
188 |
apply (intro finite_range_lemma finite_approx) |
|
189 |
done |
|
190 |
||
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26407
diff
changeset
|
191 |
end |
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26407
diff
changeset
|
192 |
|
26407
562a1d615336
rename class bifinite_cpo to profinite; generalize powerdomains from bifinite to profinite
huffman
parents:
25923
diff
changeset
|
193 |
instance "->" :: (profinite, bifinite) bifinite .. |
25909
6b96b9392873
add class bifinite_cpo for possibly-unpointed bifinite domains
huffman
parents:
25903
diff
changeset
|
194 |
|
25903 | 195 |
lemma approx_cfun: "approx n\<cdot>f\<cdot>x = approx n\<cdot>(f\<cdot>(approx n\<cdot>x))" |
196 |
by (simp add: approx_cfun_def) |
|
197 |
||
198 |
end |