author | wenzelm |
Fri, 06 Nov 2009 10:26:13 +0100 | |
changeset 33466 | 8f2e102f6628 |
parent 31076 | 99fe356cbbc2 |
child 33585 | 8d39394fe5cf |
permissions | -rw-r--r-- |
25904 | 1 |
(* Title: HOLCF/UpperPD.thy |
2 |
Author: Brian Huffman |
|
3 |
*) |
|
4 |
||
5 |
header {* Upper powerdomain *} |
|
6 |
||
7 |
theory UpperPD |
|
8 |
imports CompactBasis |
|
9 |
begin |
|
10 |
||
11 |
subsection {* Basis preorder *} |
|
12 |
||
13 |
definition |
|
14 |
upper_le :: "'a pd_basis \<Rightarrow> 'a pd_basis \<Rightarrow> bool" (infix "\<le>\<sharp>" 50) where |
|
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
15 |
"upper_le = (\<lambda>u v. \<forall>y\<in>Rep_pd_basis v. \<exists>x\<in>Rep_pd_basis u. x \<sqsubseteq> y)" |
25904 | 16 |
|
17 |
lemma upper_le_refl [simp]: "t \<le>\<sharp> t" |
|
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
18 |
unfolding upper_le_def by fast |
25904 | 19 |
|
20 |
lemma upper_le_trans: "\<lbrakk>t \<le>\<sharp> u; u \<le>\<sharp> v\<rbrakk> \<Longrightarrow> t \<le>\<sharp> v" |
|
21 |
unfolding upper_le_def |
|
22 |
apply (rule ballI) |
|
23 |
apply (drule (1) bspec, erule bexE) |
|
24 |
apply (drule (1) bspec, erule bexE) |
|
25 |
apply (erule rev_bexI) |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
26 |
apply (erule (1) below_trans) |
25904 | 27 |
done |
28 |
||
30729
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents:
29990
diff
changeset
|
29 |
interpretation upper_le: preorder upper_le |
25904 | 30 |
by (rule preorder.intro, rule upper_le_refl, rule upper_le_trans) |
31 |
||
32 |
lemma upper_le_minimal [simp]: "PDUnit compact_bot \<le>\<sharp> t" |
|
33 |
unfolding upper_le_def Rep_PDUnit by simp |
|
34 |
||
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
35 |
lemma PDUnit_upper_mono: "x \<sqsubseteq> y \<Longrightarrow> PDUnit x \<le>\<sharp> PDUnit y" |
25904 | 36 |
unfolding upper_le_def Rep_PDUnit by simp |
37 |
||
38 |
lemma PDPlus_upper_mono: "\<lbrakk>s \<le>\<sharp> t; u \<le>\<sharp> v\<rbrakk> \<Longrightarrow> PDPlus s u \<le>\<sharp> PDPlus t v" |
|
39 |
unfolding upper_le_def Rep_PDPlus by fast |
|
40 |
||
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
41 |
lemma PDPlus_upper_le: "PDPlus t u \<le>\<sharp> t" |
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
42 |
unfolding upper_le_def Rep_PDPlus by fast |
25904 | 43 |
|
44 |
lemma upper_le_PDUnit_PDUnit_iff [simp]: |
|
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
45 |
"(PDUnit a \<le>\<sharp> PDUnit b) = a \<sqsubseteq> b" |
25904 | 46 |
unfolding upper_le_def Rep_PDUnit by fast |
47 |
||
48 |
lemma upper_le_PDPlus_PDUnit_iff: |
|
49 |
"(PDPlus t u \<le>\<sharp> PDUnit a) = (t \<le>\<sharp> PDUnit a \<or> u \<le>\<sharp> PDUnit a)" |
|
50 |
unfolding upper_le_def Rep_PDPlus Rep_PDUnit by fast |
|
51 |
||
52 |
lemma upper_le_PDPlus_iff: "(t \<le>\<sharp> PDPlus u v) = (t \<le>\<sharp> u \<and> t \<le>\<sharp> v)" |
|
53 |
unfolding upper_le_def Rep_PDPlus by fast |
|
54 |
||
55 |
lemma upper_le_induct [induct set: upper_le]: |
|
56 |
assumes le: "t \<le>\<sharp> u" |
|
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
57 |
assumes 1: "\<And>a b. a \<sqsubseteq> b \<Longrightarrow> P (PDUnit a) (PDUnit b)" |
25904 | 58 |
assumes 2: "\<And>t u a. P t (PDUnit a) \<Longrightarrow> P (PDPlus t u) (PDUnit a)" |
59 |
assumes 3: "\<And>t u v. \<lbrakk>P t u; P t v\<rbrakk> \<Longrightarrow> P t (PDPlus u v)" |
|
60 |
shows "P t u" |
|
61 |
using le apply (induct u arbitrary: t rule: pd_basis_induct) |
|
62 |
apply (erule rev_mp) |
|
63 |
apply (induct_tac t rule: pd_basis_induct) |
|
64 |
apply (simp add: 1) |
|
65 |
apply (simp add: upper_le_PDPlus_PDUnit_iff) |
|
66 |
apply (simp add: 2) |
|
67 |
apply (subst PDPlus_commute) |
|
68 |
apply (simp add: 2) |
|
69 |
apply (simp add: upper_le_PDPlus_iff 3) |
|
70 |
done |
|
71 |
||
27405 | 72 |
lemma pd_take_upper_chain: |
73 |
"pd_take n t \<le>\<sharp> pd_take (Suc n) t" |
|
25904 | 74 |
apply (induct t rule: pd_basis_induct) |
27289 | 75 |
apply (simp add: compact_basis.take_chain) |
25904 | 76 |
apply (simp add: PDPlus_upper_mono) |
77 |
done |
|
78 |
||
27405 | 79 |
lemma pd_take_upper_le: "pd_take i t \<le>\<sharp> t" |
25904 | 80 |
apply (induct t rule: pd_basis_induct) |
27289 | 81 |
apply (simp add: compact_basis.take_less) |
25904 | 82 |
apply (simp add: PDPlus_upper_mono) |
83 |
done |
|
84 |
||
27405 | 85 |
lemma pd_take_upper_mono: |
86 |
"t \<le>\<sharp> u \<Longrightarrow> pd_take n t \<le>\<sharp> pd_take n u" |
|
25904 | 87 |
apply (erule upper_le_induct) |
27289 | 88 |
apply (simp add: compact_basis.take_mono) |
25904 | 89 |
apply (simp add: upper_le_PDPlus_PDUnit_iff) |
90 |
apply (simp add: upper_le_PDPlus_iff) |
|
91 |
done |
|
92 |
||
93 |
||
94 |
subsection {* Type definition *} |
|
95 |
||
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
96 |
typedef (open) 'a upper_pd = |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
97 |
"{S::'a pd_basis set. upper_le.ideal S}" |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
98 |
by (fast intro: upper_le.ideal_principal) |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
99 |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
100 |
instantiation upper_pd :: (profinite) below |
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
101 |
begin |
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 |
definition |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
104 |
"x \<sqsubseteq> y \<longleftrightarrow> Rep_upper_pd x \<subseteq> Rep_upper_pd y" |
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 |
instance .. |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
107 |
end |
25904 | 108 |
|
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
109 |
instance upper_pd :: (profinite) po |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
110 |
by (rule upper_le.typedef_ideal_po |
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
111 |
[OF type_definition_upper_pd below_upper_pd_def]) |
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
112 |
|
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
113 |
instance upper_pd :: (profinite) cpo |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
114 |
by (rule upper_le.typedef_ideal_cpo |
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
115 |
[OF type_definition_upper_pd below_upper_pd_def]) |
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
116 |
|
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
117 |
lemma Rep_upper_pd_lub: |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
118 |
"chain Y \<Longrightarrow> Rep_upper_pd (\<Squnion>i. Y i) = (\<Union>i. Rep_upper_pd (Y i))" |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
119 |
by (rule upper_le.typedef_ideal_rep_contlub |
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
120 |
[OF type_definition_upper_pd below_upper_pd_def]) |
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
121 |
|
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
122 |
lemma ideal_Rep_upper_pd: "upper_le.ideal (Rep_upper_pd xs)" |
26927 | 123 |
by (rule Rep_upper_pd [unfolded mem_Collect_eq]) |
25904 | 124 |
|
125 |
definition |
|
126 |
upper_principal :: "'a pd_basis \<Rightarrow> 'a upper_pd" where |
|
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
127 |
"upper_principal t = Abs_upper_pd {u. u \<le>\<sharp> t}" |
25904 | 128 |
|
129 |
lemma Rep_upper_principal: |
|
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
130 |
"Rep_upper_pd (upper_principal t) = {u. u \<le>\<sharp> t}" |
25904 | 131 |
unfolding upper_principal_def |
27297
2c42b1505f25
removed SetPcpo.thy and cpo instance for type bool;
huffman
parents:
27289
diff
changeset
|
132 |
by (simp add: Abs_upper_pd_inverse upper_le.ideal_principal) |
25904 | 133 |
|
30729
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents:
29990
diff
changeset
|
134 |
interpretation upper_pd: |
29237 | 135 |
ideal_completion upper_le pd_take upper_principal Rep_upper_pd |
25904 | 136 |
apply unfold_locales |
27405 | 137 |
apply (rule pd_take_upper_le) |
138 |
apply (rule pd_take_idem) |
|
139 |
apply (erule pd_take_upper_mono) |
|
140 |
apply (rule pd_take_upper_chain) |
|
141 |
apply (rule finite_range_pd_take) |
|
142 |
apply (rule pd_take_covers) |
|
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
143 |
apply (rule ideal_Rep_upper_pd) |
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
144 |
apply (erule Rep_upper_pd_lub) |
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
145 |
apply (rule Rep_upper_principal) |
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
146 |
apply (simp only: below_upper_pd_def) |
25904 | 147 |
done |
148 |
||
27289 | 149 |
text {* Upper powerdomain is pointed *} |
25904 | 150 |
|
151 |
lemma upper_pd_minimal: "upper_principal (PDUnit compact_bot) \<sqsubseteq> ys" |
|
152 |
by (induct ys rule: upper_pd.principal_induct, simp, simp) |
|
153 |
||
154 |
instance upper_pd :: (bifinite) pcpo |
|
26927 | 155 |
by intro_classes (fast intro: upper_pd_minimal) |
25904 | 156 |
|
157 |
lemma inst_upper_pd_pcpo: "\<bottom> = upper_principal (PDUnit compact_bot)" |
|
158 |
by (rule upper_pd_minimal [THEN UU_I, symmetric]) |
|
159 |
||
27289 | 160 |
text {* Upper powerdomain is profinite *} |
25904 | 161 |
|
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26927
diff
changeset
|
162 |
instantiation upper_pd :: (profinite) profinite |
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26927
diff
changeset
|
163 |
begin |
25904 | 164 |
|
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26927
diff
changeset
|
165 |
definition |
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26927
diff
changeset
|
166 |
approx_upper_pd_def: "approx = upper_pd.completion_approx" |
26927 | 167 |
|
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26927
diff
changeset
|
168 |
instance |
26927 | 169 |
apply (intro_classes, unfold approx_upper_pd_def) |
27310 | 170 |
apply (rule upper_pd.chain_completion_approx) |
26927 | 171 |
apply (rule upper_pd.lub_completion_approx) |
172 |
apply (rule upper_pd.completion_approx_idem) |
|
173 |
apply (rule upper_pd.finite_fixes_completion_approx) |
|
174 |
done |
|
175 |
||
26962
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26927
diff
changeset
|
176 |
end |
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
huffman
parents:
26927
diff
changeset
|
177 |
|
26927 | 178 |
instance upper_pd :: (bifinite) bifinite .. |
25904 | 179 |
|
180 |
lemma approx_upper_principal [simp]: |
|
27405 | 181 |
"approx n\<cdot>(upper_principal t) = upper_principal (pd_take n t)" |
25904 | 182 |
unfolding approx_upper_pd_def |
26927 | 183 |
by (rule upper_pd.completion_approx_principal) |
25904 | 184 |
|
185 |
lemma approx_eq_upper_principal: |
|
27405 | 186 |
"\<exists>t\<in>Rep_upper_pd xs. approx n\<cdot>xs = upper_principal (pd_take n t)" |
25904 | 187 |
unfolding approx_upper_pd_def |
26927 | 188 |
by (rule upper_pd.completion_approx_eq_principal) |
26407
562a1d615336
rename class bifinite_cpo to profinite; generalize powerdomains from bifinite to profinite
huffman
parents:
26041
diff
changeset
|
189 |
|
25904 | 190 |
|
26927 | 191 |
subsection {* Monadic unit and plus *} |
25904 | 192 |
|
193 |
definition |
|
194 |
upper_unit :: "'a \<rightarrow> 'a upper_pd" where |
|
195 |
"upper_unit = compact_basis.basis_fun (\<lambda>a. upper_principal (PDUnit a))" |
|
196 |
||
197 |
definition |
|
198 |
upper_plus :: "'a upper_pd \<rightarrow> 'a upper_pd \<rightarrow> 'a upper_pd" where |
|
199 |
"upper_plus = upper_pd.basis_fun (\<lambda>t. upper_pd.basis_fun (\<lambda>u. |
|
200 |
upper_principal (PDPlus t u)))" |
|
201 |
||
202 |
abbreviation |
|
203 |
upper_add :: "'a upper_pd \<Rightarrow> 'a upper_pd \<Rightarrow> 'a upper_pd" |
|
204 |
(infixl "+\<sharp>" 65) where |
|
205 |
"xs +\<sharp> ys == upper_plus\<cdot>xs\<cdot>ys" |
|
206 |
||
26927 | 207 |
syntax |
208 |
"_upper_pd" :: "args \<Rightarrow> 'a upper_pd" ("{_}\<sharp>") |
|
209 |
||
210 |
translations |
|
211 |
"{x,xs}\<sharp>" == "{x}\<sharp> +\<sharp> {xs}\<sharp>" |
|
212 |
"{x}\<sharp>" == "CONST upper_unit\<cdot>x" |
|
213 |
||
214 |
lemma upper_unit_Rep_compact_basis [simp]: |
|
215 |
"{Rep_compact_basis a}\<sharp> = upper_principal (PDUnit a)" |
|
216 |
unfolding upper_unit_def |
|
27289 | 217 |
by (simp add: compact_basis.basis_fun_principal PDUnit_upper_mono) |
26927 | 218 |
|
25904 | 219 |
lemma upper_plus_principal [simp]: |
26927 | 220 |
"upper_principal t +\<sharp> upper_principal u = upper_principal (PDPlus t u)" |
25904 | 221 |
unfolding upper_plus_def |
222 |
by (simp add: upper_pd.basis_fun_principal |
|
223 |
upper_pd.basis_fun_mono PDPlus_upper_mono) |
|
224 |
||
26927 | 225 |
lemma approx_upper_unit [simp]: |
226 |
"approx n\<cdot>{x}\<sharp> = {approx n\<cdot>x}\<sharp>" |
|
27289 | 227 |
apply (induct x rule: compact_basis.principal_induct, simp) |
26927 | 228 |
apply (simp add: approx_Rep_compact_basis) |
229 |
done |
|
230 |
||
25904 | 231 |
lemma approx_upper_plus [simp]: |
26927 | 232 |
"approx n\<cdot>(xs +\<sharp> ys) = (approx n\<cdot>xs) +\<sharp> (approx n\<cdot>ys)" |
27289 | 233 |
by (induct xs ys rule: upper_pd.principal_induct2, simp, simp, simp) |
25904 | 234 |
|
26927 | 235 |
lemma upper_plus_assoc: "(xs +\<sharp> ys) +\<sharp> zs = xs +\<sharp> (ys +\<sharp> zs)" |
27289 | 236 |
apply (induct xs ys arbitrary: zs rule: upper_pd.principal_induct2, simp, simp) |
237 |
apply (rule_tac x=zs in upper_pd.principal_induct, simp) |
|
25904 | 238 |
apply (simp add: PDPlus_assoc) |
239 |
done |
|
240 |
||
26927 | 241 |
lemma upper_plus_commute: "xs +\<sharp> ys = ys +\<sharp> xs" |
27289 | 242 |
apply (induct xs ys rule: upper_pd.principal_induct2, simp, simp) |
26927 | 243 |
apply (simp add: PDPlus_commute) |
244 |
done |
|
245 |
||
29990
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
246 |
lemma upper_plus_absorb [simp]: "xs +\<sharp> xs = xs" |
27289 | 247 |
apply (induct xs rule: upper_pd.principal_induct, simp) |
25904 | 248 |
apply (simp add: PDPlus_absorb) |
249 |
done |
|
250 |
||
29990
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
251 |
lemma upper_plus_left_commute: "xs +\<sharp> (ys +\<sharp> zs) = ys +\<sharp> (xs +\<sharp> zs)" |
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
252 |
by (rule mk_left_commute [of "op +\<sharp>", OF upper_plus_assoc upper_plus_commute]) |
26927 | 253 |
|
29990
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
254 |
lemma upper_plus_left_absorb [simp]: "xs +\<sharp> (xs +\<sharp> ys) = xs +\<sharp> ys" |
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
255 |
by (simp only: upper_plus_assoc [symmetric] upper_plus_absorb) |
26927 | 256 |
|
29990
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
257 |
text {* Useful for @{text "simp add: upper_plus_ac"} *} |
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
258 |
lemmas upper_plus_ac = |
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
259 |
upper_plus_assoc upper_plus_commute upper_plus_left_commute |
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
260 |
|
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
261 |
text {* Useful for @{text "simp only: upper_plus_aci"} *} |
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
262 |
lemmas upper_plus_aci = |
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
263 |
upper_plus_ac upper_plus_absorb upper_plus_left_absorb |
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
264 |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
265 |
lemma upper_plus_below1: "xs +\<sharp> ys \<sqsubseteq> xs" |
27289 | 266 |
apply (induct xs ys rule: upper_pd.principal_induct2, simp, simp) |
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
267 |
apply (simp add: PDPlus_upper_le) |
25904 | 268 |
done |
269 |
||
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
270 |
lemma upper_plus_below2: "xs +\<sharp> ys \<sqsubseteq> ys" |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
271 |
by (subst upper_plus_commute, rule upper_plus_below1) |
25904 | 272 |
|
26927 | 273 |
lemma upper_plus_greatest: "\<lbrakk>xs \<sqsubseteq> ys; xs \<sqsubseteq> zs\<rbrakk> \<Longrightarrow> xs \<sqsubseteq> ys +\<sharp> zs" |
25904 | 274 |
apply (subst upper_plus_absorb [of xs, symmetric]) |
275 |
apply (erule (1) monofun_cfun [OF monofun_cfun_arg]) |
|
276 |
done |
|
277 |
||
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
278 |
lemma upper_below_plus_iff: |
26927 | 279 |
"xs \<sqsubseteq> ys +\<sharp> zs \<longleftrightarrow> xs \<sqsubseteq> ys \<and> xs \<sqsubseteq> zs" |
25904 | 280 |
apply safe |
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
281 |
apply (erule below_trans [OF _ upper_plus_below1]) |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
282 |
apply (erule below_trans [OF _ upper_plus_below2]) |
25904 | 283 |
apply (erule (1) upper_plus_greatest) |
284 |
done |
|
285 |
||
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
286 |
lemma upper_plus_below_unit_iff: |
26927 | 287 |
"xs +\<sharp> ys \<sqsubseteq> {z}\<sharp> \<longleftrightarrow> xs \<sqsubseteq> {z}\<sharp> \<or> ys \<sqsubseteq> {z}\<sharp>" |
25904 | 288 |
apply (rule iffI) |
289 |
apply (subgoal_tac |
|
26927 | 290 |
"adm (\<lambda>f. f\<cdot>xs \<sqsubseteq> f\<cdot>{z}\<sharp> \<or> f\<cdot>ys \<sqsubseteq> f\<cdot>{z}\<sharp>)") |
25925 | 291 |
apply (drule admD, rule chain_approx) |
25904 | 292 |
apply (drule_tac f="approx i" in monofun_cfun_arg) |
27289 | 293 |
apply (cut_tac x="approx i\<cdot>xs" in upper_pd.compact_imp_principal, simp) |
294 |
apply (cut_tac x="approx i\<cdot>ys" in upper_pd.compact_imp_principal, simp) |
|
295 |
apply (cut_tac x="approx i\<cdot>z" in compact_basis.compact_imp_principal, simp) |
|
25904 | 296 |
apply (clarify, simp add: upper_le_PDPlus_PDUnit_iff) |
297 |
apply simp |
|
298 |
apply simp |
|
299 |
apply (erule disjE) |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
300 |
apply (erule below_trans [OF upper_plus_below1]) |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
301 |
apply (erule below_trans [OF upper_plus_below2]) |
25904 | 302 |
done |
303 |
||
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
304 |
lemma upper_unit_below_iff [simp]: "{x}\<sharp> \<sqsubseteq> {y}\<sharp> \<longleftrightarrow> x \<sqsubseteq> y" |
26927 | 305 |
apply (rule iffI) |
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
306 |
apply (rule profinite_below_ext) |
26927 | 307 |
apply (drule_tac f="approx i" in monofun_cfun_arg, simp) |
27289 | 308 |
apply (cut_tac x="approx i\<cdot>x" in compact_basis.compact_imp_principal, simp) |
309 |
apply (cut_tac x="approx i\<cdot>y" in compact_basis.compact_imp_principal, simp) |
|
310 |
apply clarsimp |
|
26927 | 311 |
apply (erule monofun_cfun_arg) |
312 |
done |
|
313 |
||
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
314 |
lemmas upper_pd_below_simps = |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
315 |
upper_unit_below_iff |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
316 |
upper_below_plus_iff |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
317 |
upper_plus_below_unit_iff |
25904 | 318 |
|
26927 | 319 |
lemma upper_unit_eq_iff [simp]: "{x}\<sharp> = {y}\<sharp> \<longleftrightarrow> x = y" |
320 |
unfolding po_eq_conv by simp |
|
321 |
||
322 |
lemma upper_unit_strict [simp]: "{\<bottom>}\<sharp> = \<bottom>" |
|
323 |
unfolding inst_upper_pd_pcpo Rep_compact_bot [symmetric] by simp |
|
324 |
||
325 |
lemma upper_plus_strict1 [simp]: "\<bottom> +\<sharp> ys = \<bottom>" |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
326 |
by (rule UU_I, rule upper_plus_below1) |
26927 | 327 |
|
328 |
lemma upper_plus_strict2 [simp]: "xs +\<sharp> \<bottom> = \<bottom>" |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
329 |
by (rule UU_I, rule upper_plus_below2) |
26927 | 330 |
|
331 |
lemma upper_unit_strict_iff [simp]: "{x}\<sharp> = \<bottom> \<longleftrightarrow> x = \<bottom>" |
|
332 |
unfolding upper_unit_strict [symmetric] by (rule upper_unit_eq_iff) |
|
333 |
||
334 |
lemma upper_plus_strict_iff [simp]: |
|
335 |
"xs +\<sharp> ys = \<bottom> \<longleftrightarrow> xs = \<bottom> \<or> ys = \<bottom>" |
|
336 |
apply (rule iffI) |
|
337 |
apply (erule rev_mp) |
|
27289 | 338 |
apply (rule upper_pd.principal_induct2 [where x=xs and y=ys], simp, simp) |
339 |
apply (simp add: inst_upper_pd_pcpo upper_pd.principal_eq_iff |
|
26927 | 340 |
upper_le_PDPlus_PDUnit_iff) |
341 |
apply auto |
|
342 |
done |
|
343 |
||
344 |
lemma compact_upper_unit_iff [simp]: "compact {x}\<sharp> \<longleftrightarrow> compact x" |
|
27309 | 345 |
unfolding profinite_compact_iff by simp |
26927 | 346 |
|
347 |
lemma compact_upper_plus [simp]: |
|
348 |
"\<lbrakk>compact xs; compact ys\<rbrakk> \<Longrightarrow> compact (xs +\<sharp> ys)" |
|
27289 | 349 |
by (auto dest!: upper_pd.compact_imp_principal) |
26927 | 350 |
|
25904 | 351 |
|
352 |
subsection {* Induction rules *} |
|
353 |
||
354 |
lemma upper_pd_induct1: |
|
355 |
assumes P: "adm P" |
|
26927 | 356 |
assumes unit: "\<And>x. P {x}\<sharp>" |
357 |
assumes insert: "\<And>x ys. \<lbrakk>P {x}\<sharp>; P ys\<rbrakk> \<Longrightarrow> P ({x}\<sharp> +\<sharp> ys)" |
|
25904 | 358 |
shows "P (xs::'a upper_pd)" |
27289 | 359 |
apply (induct xs rule: upper_pd.principal_induct, rule P) |
360 |
apply (induct_tac a rule: pd_basis_induct1) |
|
25904 | 361 |
apply (simp only: upper_unit_Rep_compact_basis [symmetric]) |
362 |
apply (rule unit) |
|
363 |
apply (simp only: upper_unit_Rep_compact_basis [symmetric] |
|
364 |
upper_plus_principal [symmetric]) |
|
365 |
apply (erule insert [OF unit]) |
|
366 |
done |
|
367 |
||
368 |
lemma upper_pd_induct: |
|
369 |
assumes P: "adm P" |
|
26927 | 370 |
assumes unit: "\<And>x. P {x}\<sharp>" |
371 |
assumes plus: "\<And>xs ys. \<lbrakk>P xs; P ys\<rbrakk> \<Longrightarrow> P (xs +\<sharp> ys)" |
|
25904 | 372 |
shows "P (xs::'a upper_pd)" |
27289 | 373 |
apply (induct xs rule: upper_pd.principal_induct, rule P) |
374 |
apply (induct_tac a rule: pd_basis_induct) |
|
25904 | 375 |
apply (simp only: upper_unit_Rep_compact_basis [symmetric] unit) |
376 |
apply (simp only: upper_plus_principal [symmetric] plus) |
|
377 |
done |
|
378 |
||
379 |
||
380 |
subsection {* Monadic bind *} |
|
381 |
||
382 |
definition |
|
383 |
upper_bind_basis :: |
|
384 |
"'a pd_basis \<Rightarrow> ('a \<rightarrow> 'b upper_pd) \<rightarrow> 'b upper_pd" where |
|
385 |
"upper_bind_basis = fold_pd |
|
386 |
(\<lambda>a. \<Lambda> f. f\<cdot>(Rep_compact_basis a)) |
|
26927 | 387 |
(\<lambda>x y. \<Lambda> f. x\<cdot>f +\<sharp> y\<cdot>f)" |
25904 | 388 |
|
26927 | 389 |
lemma ACI_upper_bind: |
390 |
"ab_semigroup_idem_mult (\<lambda>x y. \<Lambda> f. x\<cdot>f +\<sharp> y\<cdot>f)" |
|
25904 | 391 |
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
|
392 |
apply (simp add: upper_plus_assoc) |
25904 | 393 |
apply (simp add: upper_plus_commute) |
29990
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
394 |
apply (simp add: eta_cfun) |
25904 | 395 |
done |
396 |
||
397 |
lemma upper_bind_basis_simps [simp]: |
|
398 |
"upper_bind_basis (PDUnit a) = |
|
399 |
(\<Lambda> f. f\<cdot>(Rep_compact_basis a))" |
|
400 |
"upper_bind_basis (PDPlus t u) = |
|
26927 | 401 |
(\<Lambda> f. upper_bind_basis t\<cdot>f +\<sharp> upper_bind_basis u\<cdot>f)" |
25904 | 402 |
unfolding upper_bind_basis_def |
403 |
apply - |
|
26927 | 404 |
apply (rule fold_pd_PDUnit [OF ACI_upper_bind]) |
405 |
apply (rule fold_pd_PDPlus [OF ACI_upper_bind]) |
|
25904 | 406 |
done |
407 |
||
408 |
lemma upper_bind_basis_mono: |
|
409 |
"t \<le>\<sharp> u \<Longrightarrow> upper_bind_basis t \<sqsubseteq> upper_bind_basis u" |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
410 |
unfolding expand_cfun_below |
25904 | 411 |
apply (erule upper_le_induct, safe) |
27289 | 412 |
apply (simp add: monofun_cfun) |
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
413 |
apply (simp add: below_trans [OF upper_plus_below1]) |
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
30729
diff
changeset
|
414 |
apply (simp add: upper_below_plus_iff) |
25904 | 415 |
done |
416 |
||
417 |
definition |
|
418 |
upper_bind :: "'a upper_pd \<rightarrow> ('a \<rightarrow> 'b upper_pd) \<rightarrow> 'b upper_pd" where |
|
419 |
"upper_bind = upper_pd.basis_fun upper_bind_basis" |
|
420 |
||
421 |
lemma upper_bind_principal [simp]: |
|
422 |
"upper_bind\<cdot>(upper_principal t) = upper_bind_basis t" |
|
423 |
unfolding upper_bind_def |
|
424 |
apply (rule upper_pd.basis_fun_principal) |
|
425 |
apply (erule upper_bind_basis_mono) |
|
426 |
done |
|
427 |
||
428 |
lemma upper_bind_unit [simp]: |
|
26927 | 429 |
"upper_bind\<cdot>{x}\<sharp>\<cdot>f = f\<cdot>x" |
27289 | 430 |
by (induct x rule: compact_basis.principal_induct, simp, simp) |
25904 | 431 |
|
432 |
lemma upper_bind_plus [simp]: |
|
26927 | 433 |
"upper_bind\<cdot>(xs +\<sharp> ys)\<cdot>f = upper_bind\<cdot>xs\<cdot>f +\<sharp> upper_bind\<cdot>ys\<cdot>f" |
27289 | 434 |
by (induct xs ys rule: upper_pd.principal_induct2, simp, simp, simp) |
25904 | 435 |
|
436 |
lemma upper_bind_strict [simp]: "upper_bind\<cdot>\<bottom>\<cdot>f = f\<cdot>\<bottom>" |
|
437 |
unfolding upper_unit_strict [symmetric] by (rule upper_bind_unit) |
|
438 |
||
439 |
||
440 |
subsection {* Map and join *} |
|
441 |
||
442 |
definition |
|
443 |
upper_map :: "('a \<rightarrow> 'b) \<rightarrow> 'a upper_pd \<rightarrow> 'b upper_pd" where |
|
26927 | 444 |
"upper_map = (\<Lambda> f xs. upper_bind\<cdot>xs\<cdot>(\<Lambda> x. {f\<cdot>x}\<sharp>))" |
25904 | 445 |
|
446 |
definition |
|
447 |
upper_join :: "'a upper_pd upper_pd \<rightarrow> 'a upper_pd" where |
|
448 |
"upper_join = (\<Lambda> xss. upper_bind\<cdot>xss\<cdot>(\<Lambda> xs. xs))" |
|
449 |
||
450 |
lemma upper_map_unit [simp]: |
|
26927 | 451 |
"upper_map\<cdot>f\<cdot>{x}\<sharp> = {f\<cdot>x}\<sharp>" |
25904 | 452 |
unfolding upper_map_def by simp |
453 |
||
454 |
lemma upper_map_plus [simp]: |
|
26927 | 455 |
"upper_map\<cdot>f\<cdot>(xs +\<sharp> ys) = upper_map\<cdot>f\<cdot>xs +\<sharp> upper_map\<cdot>f\<cdot>ys" |
25904 | 456 |
unfolding upper_map_def by simp |
457 |
||
458 |
lemma upper_join_unit [simp]: |
|
26927 | 459 |
"upper_join\<cdot>{xs}\<sharp> = xs" |
25904 | 460 |
unfolding upper_join_def by simp |
461 |
||
462 |
lemma upper_join_plus [simp]: |
|
26927 | 463 |
"upper_join\<cdot>(xss +\<sharp> yss) = upper_join\<cdot>xss +\<sharp> upper_join\<cdot>yss" |
25904 | 464 |
unfolding upper_join_def by simp |
465 |
||
466 |
lemma upper_map_ident: "upper_map\<cdot>(\<Lambda> x. x)\<cdot>xs = xs" |
|
467 |
by (induct xs rule: upper_pd_induct, simp_all) |
|
468 |
||
469 |
lemma upper_map_map: |
|
470 |
"upper_map\<cdot>f\<cdot>(upper_map\<cdot>g\<cdot>xs) = upper_map\<cdot>(\<Lambda> x. f\<cdot>(g\<cdot>x))\<cdot>xs" |
|
471 |
by (induct xs rule: upper_pd_induct, simp_all) |
|
472 |
||
473 |
lemma upper_join_map_unit: |
|
474 |
"upper_join\<cdot>(upper_map\<cdot>upper_unit\<cdot>xs) = xs" |
|
475 |
by (induct xs rule: upper_pd_induct, simp_all) |
|
476 |
||
477 |
lemma upper_join_map_join: |
|
478 |
"upper_join\<cdot>(upper_map\<cdot>upper_join\<cdot>xsss) = upper_join\<cdot>(upper_join\<cdot>xsss)" |
|
479 |
by (induct xsss rule: upper_pd_induct, simp_all) |
|
480 |
||
481 |
lemma upper_join_map_map: |
|
482 |
"upper_join\<cdot>(upper_map\<cdot>(upper_map\<cdot>f)\<cdot>xss) = |
|
483 |
upper_map\<cdot>f\<cdot>(upper_join\<cdot>xss)" |
|
484 |
by (induct xss rule: upper_pd_induct, simp_all) |
|
485 |
||
486 |
lemma upper_map_approx: "upper_map\<cdot>(approx n)\<cdot>xs = approx n\<cdot>xs" |
|
487 |
by (induct xs rule: upper_pd_induct, simp_all) |
|
488 |
||
489 |
end |