author | wenzelm |
Tue, 08 Jul 2008 17:52:28 +0200 | |
changeset 27495 | d2bb5d61b392 |
parent 27405 | 785f5dbec8f4 |
child 29138 | 661a8db7e647 |
child 29237 | e90d9d51106b |
permissions | -rw-r--r-- |
25904 | 1 |
(* Title: HOLCF/LowerPD.thy |
2 |
ID: $Id$ |
|
3 |
Author: Brian Huffman |
|
4 |
*) |
|
5 |
||
6 |
header {* Lower powerdomain *} |
|
7 |
||
8 |
theory LowerPD |
|
9 |
imports CompactBasis |
|
10 |
begin |
|
11 |
||
12 |
subsection {* Basis preorder *} |
|
13 |
||
14 |
definition |
|
15 |
lower_le :: "'a pd_basis \<Rightarrow> 'a pd_basis \<Rightarrow> bool" (infix "\<le>\<flat>" 50) where |
|
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
16 |
"lower_le = (\<lambda>u v. \<forall>x\<in>Rep_pd_basis u. \<exists>y\<in>Rep_pd_basis v. x \<sqsubseteq> y)" |
25904 | 17 |
|
18 |
lemma lower_le_refl [simp]: "t \<le>\<flat> t" |
|
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
19 |
unfolding lower_le_def by fast |
25904 | 20 |
|
21 |
lemma lower_le_trans: "\<lbrakk>t \<le>\<flat> u; u \<le>\<flat> v\<rbrakk> \<Longrightarrow> t \<le>\<flat> v" |
|
22 |
unfolding lower_le_def |
|
23 |
apply (rule ballI) |
|
24 |
apply (drule (1) bspec, erule bexE) |
|
25 |
apply (drule (1) bspec, erule bexE) |
|
26 |
apply (erule rev_bexI) |
|
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
27 |
apply (erule (1) trans_less) |
25904 | 28 |
done |
29 |
||
30 |
interpretation lower_le: preorder [lower_le] |
|
31 |
by (rule preorder.intro, rule lower_le_refl, rule lower_le_trans) |
|
32 |
||
33 |
lemma lower_le_minimal [simp]: "PDUnit compact_bot \<le>\<flat> t" |
|
34 |
unfolding lower_le_def Rep_PDUnit |
|
35 |
by (simp, rule Rep_pd_basis_nonempty [folded ex_in_conv]) |
|
36 |
||
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
37 |
lemma PDUnit_lower_mono: "x \<sqsubseteq> y \<Longrightarrow> PDUnit x \<le>\<flat> PDUnit y" |
25904 | 38 |
unfolding lower_le_def Rep_PDUnit by fast |
39 |
||
40 |
lemma PDPlus_lower_mono: "\<lbrakk>s \<le>\<flat> t; u \<le>\<flat> v\<rbrakk> \<Longrightarrow> PDPlus s u \<le>\<flat> PDPlus t v" |
|
41 |
unfolding lower_le_def Rep_PDPlus by fast |
|
42 |
||
43 |
lemma PDPlus_lower_less: "t \<le>\<flat> PDPlus t u" |
|
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
44 |
unfolding lower_le_def Rep_PDPlus by fast |
25904 | 45 |
|
46 |
lemma lower_le_PDUnit_PDUnit_iff [simp]: |
|
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
47 |
"(PDUnit a \<le>\<flat> PDUnit b) = a \<sqsubseteq> b" |
25904 | 48 |
unfolding lower_le_def Rep_PDUnit by fast |
49 |
||
50 |
lemma lower_le_PDUnit_PDPlus_iff: |
|
51 |
"(PDUnit a \<le>\<flat> PDPlus t u) = (PDUnit a \<le>\<flat> t \<or> PDUnit a \<le>\<flat> u)" |
|
52 |
unfolding lower_le_def Rep_PDPlus Rep_PDUnit by fast |
|
53 |
||
54 |
lemma lower_le_PDPlus_iff: "(PDPlus t u \<le>\<flat> v) = (t \<le>\<flat> v \<and> u \<le>\<flat> v)" |
|
55 |
unfolding lower_le_def Rep_PDPlus by fast |
|
56 |
||
57 |
lemma lower_le_induct [induct set: lower_le]: |
|
58 |
assumes le: "t \<le>\<flat> u" |
|
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
59 |
assumes 1: "\<And>a b. a \<sqsubseteq> b \<Longrightarrow> P (PDUnit a) (PDUnit b)" |
25904 | 60 |
assumes 2: "\<And>t u a. P (PDUnit a) t \<Longrightarrow> P (PDUnit a) (PDPlus t u)" |
61 |
assumes 3: "\<And>t u v. \<lbrakk>P t v; P u v\<rbrakk> \<Longrightarrow> P (PDPlus t u) v" |
|
62 |
shows "P t u" |
|
63 |
using le |
|
64 |
apply (induct t arbitrary: u rule: pd_basis_induct) |
|
65 |
apply (erule rev_mp) |
|
66 |
apply (induct_tac u rule: pd_basis_induct) |
|
67 |
apply (simp add: 1) |
|
68 |
apply (simp add: lower_le_PDUnit_PDPlus_iff) |
|
69 |
apply (simp add: 2) |
|
70 |
apply (subst PDPlus_commute) |
|
71 |
apply (simp add: 2) |
|
72 |
apply (simp add: lower_le_PDPlus_iff 3) |
|
73 |
done |
|
74 |
||
27405 | 75 |
lemma pd_take_lower_chain: |
76 |
"pd_take n t \<le>\<flat> pd_take (Suc n) t" |
|
25904 | 77 |
apply (induct t rule: pd_basis_induct) |
27289 | 78 |
apply (simp add: compact_basis.take_chain) |
25904 | 79 |
apply (simp add: PDPlus_lower_mono) |
80 |
done |
|
81 |
||
27405 | 82 |
lemma pd_take_lower_le: "pd_take i t \<le>\<flat> t" |
25904 | 83 |
apply (induct t rule: pd_basis_induct) |
27289 | 84 |
apply (simp add: compact_basis.take_less) |
25904 | 85 |
apply (simp add: PDPlus_lower_mono) |
86 |
done |
|
87 |
||
27405 | 88 |
lemma pd_take_lower_mono: |
89 |
"t \<le>\<flat> u \<Longrightarrow> pd_take n t \<le>\<flat> pd_take n u" |
|
25904 | 90 |
apply (erule lower_le_induct) |
27289 | 91 |
apply (simp add: compact_basis.take_mono) |
25904 | 92 |
apply (simp add: lower_le_PDUnit_PDPlus_iff) |
93 |
apply (simp add: lower_le_PDPlus_iff) |
|
94 |
done |
|
95 |
||
96 |
||
97 |
subsection {* Type definition *} |
|
98 |
||
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
99 |
typedef (open) 'a lower_pd = |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
100 |
"{S::'a pd_basis set. lower_le.ideal S}" |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
101 |
by (fast intro: lower_le.ideal_principal) |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
102 |
|
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
103 |
instantiation lower_pd :: (profinite) sq_ord |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
104 |
begin |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
105 |
|
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
106 |
definition |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
107 |
"x \<sqsubseteq> y \<longleftrightarrow> Rep_lower_pd x \<subseteq> Rep_lower_pd y" |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
108 |
|
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
109 |
instance .. |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
110 |
end |
25904 | 111 |
|
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
112 |
instance lower_pd :: (profinite) po |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
113 |
by (rule lower_le.typedef_ideal_po |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
114 |
[OF type_definition_lower_pd sq_le_lower_pd_def]) |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
115 |
|
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
116 |
instance lower_pd :: (profinite) cpo |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
117 |
by (rule lower_le.typedef_ideal_cpo |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
118 |
[OF type_definition_lower_pd sq_le_lower_pd_def]) |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
119 |
|
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
120 |
lemma Rep_lower_pd_lub: |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
121 |
"chain Y \<Longrightarrow> Rep_lower_pd (\<Squnion>i. Y i) = (\<Union>i. Rep_lower_pd (Y i))" |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
122 |
by (rule lower_le.typedef_ideal_rep_contlub |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
123 |
[OF type_definition_lower_pd sq_le_lower_pd_def]) |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
124 |
|
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
125 |
lemma ideal_Rep_lower_pd: "lower_le.ideal (Rep_lower_pd xs)" |
26927 | 126 |
by (rule Rep_lower_pd [unfolded mem_Collect_eq]) |
25904 | 127 |
|
128 |
definition |
|
129 |
lower_principal :: "'a pd_basis \<Rightarrow> 'a lower_pd" where |
|
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
130 |
"lower_principal t = Abs_lower_pd {u. u \<le>\<flat> t}" |
25904 | 131 |
|
132 |
lemma Rep_lower_principal: |
|
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
133 |
"Rep_lower_pd (lower_principal t) = {u. u \<le>\<flat> t}" |
25904 | 134 |
unfolding lower_principal_def |
27297
2c42b1505f25
removed SetPcpo.thy and cpo instance for type bool;
huffman
parents:
27289
diff
changeset
|
135 |
by (simp add: Abs_lower_pd_inverse lower_le.ideal_principal) |
25904 | 136 |
|
137 |
interpretation lower_pd: |
|
27405 | 138 |
ideal_completion [lower_le pd_take lower_principal Rep_lower_pd] |
25904 | 139 |
apply unfold_locales |
27405 | 140 |
apply (rule pd_take_lower_le) |
141 |
apply (rule pd_take_idem) |
|
142 |
apply (erule pd_take_lower_mono) |
|
143 |
apply (rule pd_take_lower_chain) |
|
144 |
apply (rule finite_range_pd_take) |
|
145 |
apply (rule pd_take_covers) |
|
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
146 |
apply (rule ideal_Rep_lower_pd) |
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
147 |
apply (erule Rep_lower_pd_lub) |
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
148 |
apply (rule Rep_lower_principal) |
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
149 |
apply (simp only: sq_le_lower_pd_def) |
25904 | 150 |
done |
151 |
||
27289 | 152 |
text {* Lower powerdomain is pointed *} |
25904 | 153 |
|
154 |
lemma lower_pd_minimal: "lower_principal (PDUnit compact_bot) \<sqsubseteq> ys" |
|
155 |
by (induct ys rule: lower_pd.principal_induct, simp, simp) |
|
156 |
||
157 |
instance lower_pd :: (bifinite) pcpo |
|
26927 | 158 |
by intro_classes (fast intro: lower_pd_minimal) |
25904 | 159 |
|
160 |
lemma inst_lower_pd_pcpo: "\<bottom> = lower_principal (PDUnit compact_bot)" |
|
161 |
by (rule lower_pd_minimal [THEN UU_I, symmetric]) |
|
162 |
||
27289 | 163 |
text {* Lower powerdomain is profinite *} |
25904 | 164 |
|
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26927
diff
changeset
|
165 |
instantiation lower_pd :: (profinite) profinite |
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26927
diff
changeset
|
166 |
begin |
25904 | 167 |
|
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26927
diff
changeset
|
168 |
definition |
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26927
diff
changeset
|
169 |
approx_lower_pd_def: "approx = lower_pd.completion_approx" |
26927 | 170 |
|
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26927
diff
changeset
|
171 |
instance |
26927 | 172 |
apply (intro_classes, unfold approx_lower_pd_def) |
27310 | 173 |
apply (rule lower_pd.chain_completion_approx) |
26927 | 174 |
apply (rule lower_pd.lub_completion_approx) |
175 |
apply (rule lower_pd.completion_approx_idem) |
|
176 |
apply (rule lower_pd.finite_fixes_completion_approx) |
|
177 |
done |
|
178 |
||
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26927
diff
changeset
|
179 |
end |
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26927
diff
changeset
|
180 |
|
26927 | 181 |
instance lower_pd :: (bifinite) bifinite .. |
25904 | 182 |
|
183 |
lemma approx_lower_principal [simp]: |
|
27405 | 184 |
"approx n\<cdot>(lower_principal t) = lower_principal (pd_take n t)" |
25904 | 185 |
unfolding approx_lower_pd_def |
26927 | 186 |
by (rule lower_pd.completion_approx_principal) |
25904 | 187 |
|
188 |
lemma approx_eq_lower_principal: |
|
27405 | 189 |
"\<exists>t\<in>Rep_lower_pd xs. approx n\<cdot>xs = lower_principal (pd_take n t)" |
25904 | 190 |
unfolding approx_lower_pd_def |
26927 | 191 |
by (rule lower_pd.completion_approx_eq_principal) |
26407
562a1d615336
rename class bifinite_cpo to profinite; generalize powerdomains from bifinite to profinite
huffman
parents:
26041
diff
changeset
|
192 |
|
25904 | 193 |
|
26927 | 194 |
subsection {* Monadic unit and plus *} |
25904 | 195 |
|
196 |
definition |
|
197 |
lower_unit :: "'a \<rightarrow> 'a lower_pd" where |
|
198 |
"lower_unit = compact_basis.basis_fun (\<lambda>a. lower_principal (PDUnit a))" |
|
199 |
||
200 |
definition |
|
201 |
lower_plus :: "'a lower_pd \<rightarrow> 'a lower_pd \<rightarrow> 'a lower_pd" where |
|
202 |
"lower_plus = lower_pd.basis_fun (\<lambda>t. lower_pd.basis_fun (\<lambda>u. |
|
203 |
lower_principal (PDPlus t u)))" |
|
204 |
||
205 |
abbreviation |
|
206 |
lower_add :: "'a lower_pd \<Rightarrow> 'a lower_pd \<Rightarrow> 'a lower_pd" |
|
207 |
(infixl "+\<flat>" 65) where |
|
208 |
"xs +\<flat> ys == lower_plus\<cdot>xs\<cdot>ys" |
|
209 |
||
26927 | 210 |
syntax |
211 |
"_lower_pd" :: "args \<Rightarrow> 'a lower_pd" ("{_}\<flat>") |
|
212 |
||
213 |
translations |
|
214 |
"{x,xs}\<flat>" == "{x}\<flat> +\<flat> {xs}\<flat>" |
|
215 |
"{x}\<flat>" == "CONST lower_unit\<cdot>x" |
|
216 |
||
217 |
lemma lower_unit_Rep_compact_basis [simp]: |
|
218 |
"{Rep_compact_basis a}\<flat> = lower_principal (PDUnit a)" |
|
219 |
unfolding lower_unit_def |
|
27289 | 220 |
by (simp add: compact_basis.basis_fun_principal PDUnit_lower_mono) |
26927 | 221 |
|
25904 | 222 |
lemma lower_plus_principal [simp]: |
26927 | 223 |
"lower_principal t +\<flat> lower_principal u = lower_principal (PDPlus t u)" |
25904 | 224 |
unfolding lower_plus_def |
225 |
by (simp add: lower_pd.basis_fun_principal |
|
226 |
lower_pd.basis_fun_mono PDPlus_lower_mono) |
|
227 |
||
26927 | 228 |
lemma approx_lower_unit [simp]: |
229 |
"approx n\<cdot>{x}\<flat> = {approx n\<cdot>x}\<flat>" |
|
27289 | 230 |
apply (induct x rule: compact_basis.principal_induct, simp) |
26927 | 231 |
apply (simp add: approx_Rep_compact_basis) |
232 |
done |
|
233 |
||
25904 | 234 |
lemma approx_lower_plus [simp]: |
26927 | 235 |
"approx n\<cdot>(xs +\<flat> ys) = (approx n\<cdot>xs) +\<flat> (approx n\<cdot>ys)" |
27289 | 236 |
by (induct xs ys rule: lower_pd.principal_induct2, simp, simp, simp) |
25904 | 237 |
|
26927 | 238 |
lemma lower_plus_assoc: "(xs +\<flat> ys) +\<flat> zs = xs +\<flat> (ys +\<flat> zs)" |
27289 | 239 |
apply (induct xs ys arbitrary: zs rule: lower_pd.principal_induct2, simp, simp) |
240 |
apply (rule_tac x=zs in lower_pd.principal_induct, simp) |
|
25904 | 241 |
apply (simp add: PDPlus_assoc) |
242 |
done |
|
243 |
||
26927 | 244 |
lemma lower_plus_commute: "xs +\<flat> ys = ys +\<flat> xs" |
27289 | 245 |
apply (induct xs ys rule: lower_pd.principal_induct2, simp, simp) |
26927 | 246 |
apply (simp add: PDPlus_commute) |
247 |
done |
|
248 |
||
249 |
lemma lower_plus_absorb: "xs +\<flat> xs = xs" |
|
27289 | 250 |
apply (induct xs rule: lower_pd.principal_induct, simp) |
25904 | 251 |
apply (simp add: PDPlus_absorb) |
252 |
done |
|
253 |
||
26927 | 254 |
interpretation aci_lower_plus: ab_semigroup_idem_mult ["op +\<flat>"] |
255 |
by unfold_locales |
|
256 |
(rule lower_plus_assoc lower_plus_commute lower_plus_absorb)+ |
|
257 |
||
258 |
lemma lower_plus_left_commute: "xs +\<flat> (ys +\<flat> zs) = ys +\<flat> (xs +\<flat> zs)" |
|
259 |
by (rule aci_lower_plus.mult_left_commute) |
|
260 |
||
261 |
lemma lower_plus_left_absorb: "xs +\<flat> (xs +\<flat> ys) = xs +\<flat> ys" |
|
262 |
by (rule aci_lower_plus.mult_left_idem) |
|
263 |
||
264 |
lemmas lower_plus_aci = aci_lower_plus.mult_ac_idem |
|
265 |
||
266 |
lemma lower_plus_less1: "xs \<sqsubseteq> xs +\<flat> ys" |
|
27289 | 267 |
apply (induct xs ys rule: lower_pd.principal_induct2, simp, simp) |
25904 | 268 |
apply (simp add: PDPlus_lower_less) |
269 |
done |
|
270 |
||
26927 | 271 |
lemma lower_plus_less2: "ys \<sqsubseteq> xs +\<flat> ys" |
25904 | 272 |
by (subst lower_plus_commute, rule lower_plus_less1) |
273 |
||
26927 | 274 |
lemma lower_plus_least: "\<lbrakk>xs \<sqsubseteq> zs; ys \<sqsubseteq> zs\<rbrakk> \<Longrightarrow> xs +\<flat> ys \<sqsubseteq> zs" |
25904 | 275 |
apply (subst lower_plus_absorb [of zs, symmetric]) |
276 |
apply (erule (1) monofun_cfun [OF monofun_cfun_arg]) |
|
277 |
done |
|
278 |
||
279 |
lemma lower_plus_less_iff: |
|
26927 | 280 |
"xs +\<flat> ys \<sqsubseteq> zs \<longleftrightarrow> xs \<sqsubseteq> zs \<and> ys \<sqsubseteq> zs" |
25904 | 281 |
apply safe |
282 |
apply (erule trans_less [OF lower_plus_less1]) |
|
283 |
apply (erule trans_less [OF lower_plus_less2]) |
|
284 |
apply (erule (1) lower_plus_least) |
|
285 |
done |
|
286 |
||
287 |
lemma lower_unit_less_plus_iff: |
|
26927 | 288 |
"{x}\<flat> \<sqsubseteq> ys +\<flat> zs \<longleftrightarrow> {x}\<flat> \<sqsubseteq> ys \<or> {x}\<flat> \<sqsubseteq> zs" |
25904 | 289 |
apply (rule iffI) |
290 |
apply (subgoal_tac |
|
26927 | 291 |
"adm (\<lambda>f. f\<cdot>{x}\<flat> \<sqsubseteq> f\<cdot>ys \<or> f\<cdot>{x}\<flat> \<sqsubseteq> f\<cdot>zs)") |
25925 | 292 |
apply (drule admD, rule chain_approx) |
25904 | 293 |
apply (drule_tac f="approx i" in monofun_cfun_arg) |
27289 | 294 |
apply (cut_tac x="approx i\<cdot>x" in compact_basis.compact_imp_principal, simp) |
295 |
apply (cut_tac x="approx i\<cdot>ys" in lower_pd.compact_imp_principal, simp) |
|
296 |
apply (cut_tac x="approx i\<cdot>zs" in lower_pd.compact_imp_principal, simp) |
|
25904 | 297 |
apply (clarify, simp add: lower_le_PDUnit_PDPlus_iff) |
298 |
apply simp |
|
299 |
apply simp |
|
300 |
apply (erule disjE) |
|
301 |
apply (erule trans_less [OF _ lower_plus_less1]) |
|
302 |
apply (erule trans_less [OF _ lower_plus_less2]) |
|
303 |
done |
|
304 |
||
26927 | 305 |
lemma lower_unit_less_iff [simp]: "{x}\<flat> \<sqsubseteq> {y}\<flat> \<longleftrightarrow> x \<sqsubseteq> y" |
306 |
apply (rule iffI) |
|
27309 | 307 |
apply (rule profinite_less_ext) |
26927 | 308 |
apply (drule_tac f="approx i" in monofun_cfun_arg, simp) |
27289 | 309 |
apply (cut_tac x="approx i\<cdot>x" in compact_basis.compact_imp_principal, simp) |
310 |
apply (cut_tac x="approx i\<cdot>y" in compact_basis.compact_imp_principal, simp) |
|
311 |
apply clarsimp |
|
26927 | 312 |
apply (erule monofun_cfun_arg) |
313 |
done |
|
314 |
||
25904 | 315 |
lemmas lower_pd_less_simps = |
316 |
lower_unit_less_iff |
|
317 |
lower_plus_less_iff |
|
318 |
lower_unit_less_plus_iff |
|
319 |
||
27289 | 320 |
lemma fooble: |
321 |
fixes f :: "'a::po \<Rightarrow> 'b::po" |
|
322 |
assumes f: "\<And>x y. f x \<sqsubseteq> f y \<longleftrightarrow> x \<sqsubseteq> y" |
|
323 |
shows "f x = f y \<longleftrightarrow> x = y" |
|
324 |
unfolding po_eq_conv by (simp add: f) |
|
325 |
||
26927 | 326 |
lemma lower_unit_eq_iff [simp]: "{x}\<flat> = {y}\<flat> \<longleftrightarrow> x = y" |
27289 | 327 |
by (rule lower_unit_less_iff [THEN fooble]) |
26927 | 328 |
|
329 |
lemma lower_unit_strict [simp]: "{\<bottom>}\<flat> = \<bottom>" |
|
330 |
unfolding inst_lower_pd_pcpo Rep_compact_bot [symmetric] by simp |
|
331 |
||
332 |
lemma lower_unit_strict_iff [simp]: "{x}\<flat> = \<bottom> \<longleftrightarrow> x = \<bottom>" |
|
333 |
unfolding lower_unit_strict [symmetric] by (rule lower_unit_eq_iff) |
|
334 |
||
335 |
lemma lower_plus_strict_iff [simp]: |
|
336 |
"xs +\<flat> ys = \<bottom> \<longleftrightarrow> xs = \<bottom> \<and> ys = \<bottom>" |
|
337 |
apply safe |
|
338 |
apply (rule UU_I, erule subst, rule lower_plus_less1) |
|
339 |
apply (rule UU_I, erule subst, rule lower_plus_less2) |
|
340 |
apply (rule lower_plus_absorb) |
|
341 |
done |
|
342 |
||
343 |
lemma lower_plus_strict1 [simp]: "\<bottom> +\<flat> ys = ys" |
|
344 |
apply (rule antisym_less [OF _ lower_plus_less2]) |
|
345 |
apply (simp add: lower_plus_least) |
|
346 |
done |
|
347 |
||
348 |
lemma lower_plus_strict2 [simp]: "xs +\<flat> \<bottom> = xs" |
|
349 |
apply (rule antisym_less [OF _ lower_plus_less1]) |
|
350 |
apply (simp add: lower_plus_least) |
|
351 |
done |
|
352 |
||
353 |
lemma compact_lower_unit_iff [simp]: "compact {x}\<flat> \<longleftrightarrow> compact x" |
|
27309 | 354 |
unfolding profinite_compact_iff by simp |
26927 | 355 |
|
356 |
lemma compact_lower_plus [simp]: |
|
357 |
"\<lbrakk>compact xs; compact ys\<rbrakk> \<Longrightarrow> compact (xs +\<flat> ys)" |
|
27289 | 358 |
by (auto dest!: lower_pd.compact_imp_principal) |
26927 | 359 |
|
25904 | 360 |
|
361 |
subsection {* Induction rules *} |
|
362 |
||
363 |
lemma lower_pd_induct1: |
|
364 |
assumes P: "adm P" |
|
26927 | 365 |
assumes unit: "\<And>x. P {x}\<flat>" |
25904 | 366 |
assumes insert: |
26927 | 367 |
"\<And>x ys. \<lbrakk>P {x}\<flat>; P ys\<rbrakk> \<Longrightarrow> P ({x}\<flat> +\<flat> ys)" |
25904 | 368 |
shows "P (xs::'a lower_pd)" |
27289 | 369 |
apply (induct xs rule: lower_pd.principal_induct, rule P) |
370 |
apply (induct_tac a rule: pd_basis_induct1) |
|
25904 | 371 |
apply (simp only: lower_unit_Rep_compact_basis [symmetric]) |
372 |
apply (rule unit) |
|
373 |
apply (simp only: lower_unit_Rep_compact_basis [symmetric] |
|
374 |
lower_plus_principal [symmetric]) |
|
375 |
apply (erule insert [OF unit]) |
|
376 |
done |
|
377 |
||
378 |
lemma lower_pd_induct: |
|
379 |
assumes P: "adm P" |
|
26927 | 380 |
assumes unit: "\<And>x. P {x}\<flat>" |
381 |
assumes plus: "\<And>xs ys. \<lbrakk>P xs; P ys\<rbrakk> \<Longrightarrow> P (xs +\<flat> ys)" |
|
25904 | 382 |
shows "P (xs::'a lower_pd)" |
27289 | 383 |
apply (induct xs rule: lower_pd.principal_induct, rule P) |
384 |
apply (induct_tac a rule: pd_basis_induct) |
|
25904 | 385 |
apply (simp only: lower_unit_Rep_compact_basis [symmetric] unit) |
386 |
apply (simp only: lower_plus_principal [symmetric] plus) |
|
387 |
done |
|
388 |
||
389 |
||
390 |
subsection {* Monadic bind *} |
|
391 |
||
392 |
definition |
|
393 |
lower_bind_basis :: |
|
394 |
"'a pd_basis \<Rightarrow> ('a \<rightarrow> 'b lower_pd) \<rightarrow> 'b lower_pd" where |
|
395 |
"lower_bind_basis = fold_pd |
|
396 |
(\<lambda>a. \<Lambda> f. f\<cdot>(Rep_compact_basis a)) |
|
26927 | 397 |
(\<lambda>x y. \<Lambda> f. x\<cdot>f +\<flat> y\<cdot>f)" |
25904 | 398 |
|
26927 | 399 |
lemma ACI_lower_bind: |
400 |
"ab_semigroup_idem_mult (\<lambda>x y. \<Lambda> f. x\<cdot>f +\<flat> y\<cdot>f)" |
|
25904 | 401 |
apply unfold_locales |
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:
25925
diff
changeset
|
402 |
apply (simp add: lower_plus_assoc) |
25904 | 403 |
apply (simp add: lower_plus_commute) |
404 |
apply (simp add: lower_plus_absorb eta_cfun) |
|
405 |
done |
|
406 |
||
407 |
lemma lower_bind_basis_simps [simp]: |
|
408 |
"lower_bind_basis (PDUnit a) = |
|
409 |
(\<Lambda> f. f\<cdot>(Rep_compact_basis a))" |
|
410 |
"lower_bind_basis (PDPlus t u) = |
|
26927 | 411 |
(\<Lambda> f. lower_bind_basis t\<cdot>f +\<flat> lower_bind_basis u\<cdot>f)" |
25904 | 412 |
unfolding lower_bind_basis_def |
413 |
apply - |
|
26927 | 414 |
apply (rule fold_pd_PDUnit [OF ACI_lower_bind]) |
415 |
apply (rule fold_pd_PDPlus [OF ACI_lower_bind]) |
|
25904 | 416 |
done |
417 |
||
418 |
lemma lower_bind_basis_mono: |
|
419 |
"t \<le>\<flat> u \<Longrightarrow> lower_bind_basis t \<sqsubseteq> lower_bind_basis u" |
|
420 |
unfolding expand_cfun_less |
|
421 |
apply (erule lower_le_induct, safe) |
|
27289 | 422 |
apply (simp add: monofun_cfun) |
25904 | 423 |
apply (simp add: rev_trans_less [OF lower_plus_less1]) |
424 |
apply (simp add: lower_plus_less_iff) |
|
425 |
done |
|
426 |
||
427 |
definition |
|
428 |
lower_bind :: "'a lower_pd \<rightarrow> ('a \<rightarrow> 'b lower_pd) \<rightarrow> 'b lower_pd" where |
|
429 |
"lower_bind = lower_pd.basis_fun lower_bind_basis" |
|
430 |
||
431 |
lemma lower_bind_principal [simp]: |
|
432 |
"lower_bind\<cdot>(lower_principal t) = lower_bind_basis t" |
|
433 |
unfolding lower_bind_def |
|
434 |
apply (rule lower_pd.basis_fun_principal) |
|
435 |
apply (erule lower_bind_basis_mono) |
|
436 |
done |
|
437 |
||
438 |
lemma lower_bind_unit [simp]: |
|
26927 | 439 |
"lower_bind\<cdot>{x}\<flat>\<cdot>f = f\<cdot>x" |
27289 | 440 |
by (induct x rule: compact_basis.principal_induct, simp, simp) |
25904 | 441 |
|
442 |
lemma lower_bind_plus [simp]: |
|
26927 | 443 |
"lower_bind\<cdot>(xs +\<flat> ys)\<cdot>f = lower_bind\<cdot>xs\<cdot>f +\<flat> lower_bind\<cdot>ys\<cdot>f" |
27289 | 444 |
by (induct xs ys rule: lower_pd.principal_induct2, simp, simp, simp) |
25904 | 445 |
|
446 |
lemma lower_bind_strict [simp]: "lower_bind\<cdot>\<bottom>\<cdot>f = f\<cdot>\<bottom>" |
|
447 |
unfolding lower_unit_strict [symmetric] by (rule lower_bind_unit) |
|
448 |
||
449 |
||
450 |
subsection {* Map and join *} |
|
451 |
||
452 |
definition |
|
453 |
lower_map :: "('a \<rightarrow> 'b) \<rightarrow> 'a lower_pd \<rightarrow> 'b lower_pd" where |
|
26927 | 454 |
"lower_map = (\<Lambda> f xs. lower_bind\<cdot>xs\<cdot>(\<Lambda> x. {f\<cdot>x}\<flat>))" |
25904 | 455 |
|
456 |
definition |
|
457 |
lower_join :: "'a lower_pd lower_pd \<rightarrow> 'a lower_pd" where |
|
458 |
"lower_join = (\<Lambda> xss. lower_bind\<cdot>xss\<cdot>(\<Lambda> xs. xs))" |
|
459 |
||
460 |
lemma lower_map_unit [simp]: |
|
26927 | 461 |
"lower_map\<cdot>f\<cdot>{x}\<flat> = {f\<cdot>x}\<flat>" |
25904 | 462 |
unfolding lower_map_def by simp |
463 |
||
464 |
lemma lower_map_plus [simp]: |
|
26927 | 465 |
"lower_map\<cdot>f\<cdot>(xs +\<flat> ys) = lower_map\<cdot>f\<cdot>xs +\<flat> lower_map\<cdot>f\<cdot>ys" |
25904 | 466 |
unfolding lower_map_def by simp |
467 |
||
468 |
lemma lower_join_unit [simp]: |
|
26927 | 469 |
"lower_join\<cdot>{xs}\<flat> = xs" |
25904 | 470 |
unfolding lower_join_def by simp |
471 |
||
472 |
lemma lower_join_plus [simp]: |
|
26927 | 473 |
"lower_join\<cdot>(xss +\<flat> yss) = lower_join\<cdot>xss +\<flat> lower_join\<cdot>yss" |
25904 | 474 |
unfolding lower_join_def by simp |
475 |
||
476 |
lemma lower_map_ident: "lower_map\<cdot>(\<Lambda> x. x)\<cdot>xs = xs" |
|
477 |
by (induct xs rule: lower_pd_induct, simp_all) |
|
478 |
||
479 |
lemma lower_map_map: |
|
480 |
"lower_map\<cdot>f\<cdot>(lower_map\<cdot>g\<cdot>xs) = lower_map\<cdot>(\<Lambda> x. f\<cdot>(g\<cdot>x))\<cdot>xs" |
|
481 |
by (induct xs rule: lower_pd_induct, simp_all) |
|
482 |
||
483 |
lemma lower_join_map_unit: |
|
484 |
"lower_join\<cdot>(lower_map\<cdot>lower_unit\<cdot>xs) = xs" |
|
485 |
by (induct xs rule: lower_pd_induct, simp_all) |
|
486 |
||
487 |
lemma lower_join_map_join: |
|
488 |
"lower_join\<cdot>(lower_map\<cdot>lower_join\<cdot>xsss) = lower_join\<cdot>(lower_join\<cdot>xsss)" |
|
489 |
by (induct xsss rule: lower_pd_induct, simp_all) |
|
490 |
||
491 |
lemma lower_join_map_map: |
|
492 |
"lower_join\<cdot>(lower_map\<cdot>(lower_map\<cdot>f)\<cdot>xss) = |
|
493 |
lower_map\<cdot>f\<cdot>(lower_join\<cdot>xss)" |
|
494 |
by (induct xss rule: lower_pd_induct, simp_all) |
|
495 |
||
496 |
lemma lower_map_approx: "lower_map\<cdot>(approx n)\<cdot>xs = approx n\<cdot>xs" |
|
497 |
by (induct xs rule: lower_pd_induct, simp_all) |
|
498 |
||
499 |
end |