author | huffman |
Tue, 16 Nov 2010 12:08:28 -0800 | |
changeset 40577 | 5c6225a1c2c0 |
parent 40576 | fa5e0cacd5e7 |
child 40589 | 0e77e45d2ffc |
permissions | -rw-r--r-- |
25904 | 1 |
(* Title: HOLCF/LowerPD.thy |
2 |
Author: Brian Huffman |
|
3 |
*) |
|
4 |
||
5 |
header {* Lower powerdomain *} |
|
6 |
||
7 |
theory LowerPD |
|
8 |
imports CompactBasis |
|
9 |
begin |
|
10 |
||
11 |
subsection {* Basis preorder *} |
|
12 |
||
13 |
definition |
|
14 |
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
|
15 |
"lower_le = (\<lambda>u v. \<forall>x\<in>Rep_pd_basis u. \<exists>y\<in>Rep_pd_basis v. x \<sqsubseteq> y)" |
25904 | 16 |
|
17 |
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
|
18 |
unfolding lower_le_def by fast |
25904 | 19 |
|
20 |
lemma lower_le_trans: "\<lbrakk>t \<le>\<flat> u; u \<le>\<flat> v\<rbrakk> \<Longrightarrow> t \<le>\<flat> v" |
|
21 |
unfolding lower_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 lower_le: preorder lower_le |
25904 | 30 |
by (rule preorder.intro, rule lower_le_refl, rule lower_le_trans) |
31 |
||
32 |
lemma lower_le_minimal [simp]: "PDUnit compact_bot \<le>\<flat> t" |
|
33 |
unfolding lower_le_def Rep_PDUnit |
|
34 |
by (simp, rule Rep_pd_basis_nonempty [folded ex_in_conv]) |
|
35 |
||
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
36 |
lemma PDUnit_lower_mono: "x \<sqsubseteq> y \<Longrightarrow> PDUnit x \<le>\<flat> PDUnit y" |
25904 | 37 |
unfolding lower_le_def Rep_PDUnit by fast |
38 |
||
39 |
lemma PDPlus_lower_mono: "\<lbrakk>s \<le>\<flat> t; u \<le>\<flat> v\<rbrakk> \<Longrightarrow> PDPlus s u \<le>\<flat> PDPlus t v" |
|
40 |
unfolding lower_le_def Rep_PDPlus by fast |
|
41 |
||
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
|
42 |
lemma PDPlus_lower_le: "t \<le>\<flat> PDPlus t u" |
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
43 |
unfolding lower_le_def Rep_PDPlus by fast |
25904 | 44 |
|
45 |
lemma lower_le_PDUnit_PDUnit_iff [simp]: |
|
40436 | 46 |
"(PDUnit a \<le>\<flat> PDUnit b) = (a \<sqsubseteq> b)" |
25904 | 47 |
unfolding lower_le_def Rep_PDUnit by fast |
48 |
||
49 |
lemma lower_le_PDUnit_PDPlus_iff: |
|
50 |
"(PDUnit a \<le>\<flat> PDPlus t u) = (PDUnit a \<le>\<flat> t \<or> PDUnit a \<le>\<flat> u)" |
|
51 |
unfolding lower_le_def Rep_PDPlus Rep_PDUnit by fast |
|
52 |
||
53 |
lemma lower_le_PDPlus_iff: "(PDPlus t u \<le>\<flat> v) = (t \<le>\<flat> v \<and> u \<le>\<flat> v)" |
|
54 |
unfolding lower_le_def Rep_PDPlus by fast |
|
55 |
||
56 |
lemma lower_le_induct [induct set: lower_le]: |
|
57 |
assumes le: "t \<le>\<flat> u" |
|
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
58 |
assumes 1: "\<And>a b. a \<sqsubseteq> b \<Longrightarrow> P (PDUnit a) (PDUnit b)" |
25904 | 59 |
assumes 2: "\<And>t u a. P (PDUnit a) t \<Longrightarrow> P (PDUnit a) (PDPlus t u)" |
60 |
assumes 3: "\<And>t u v. \<lbrakk>P t v; P u v\<rbrakk> \<Longrightarrow> P (PDPlus t u) v" |
|
61 |
shows "P t u" |
|
62 |
using le |
|
63 |
apply (induct t arbitrary: u rule: pd_basis_induct) |
|
64 |
apply (erule rev_mp) |
|
65 |
apply (induct_tac u rule: pd_basis_induct) |
|
66 |
apply (simp add: 1) |
|
67 |
apply (simp add: lower_le_PDUnit_PDPlus_iff) |
|
68 |
apply (simp add: 2) |
|
69 |
apply (subst PDPlus_commute) |
|
70 |
apply (simp add: 2) |
|
71 |
apply (simp add: lower_le_PDPlus_iff 3) |
|
72 |
done |
|
73 |
||
74 |
||
75 |
subsection {* Type definition *} |
|
76 |
||
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
77 |
typedef (open) 'a lower_pd = |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
78 |
"{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
|
79 |
by (fast intro: lower_le.ideal_principal) |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
80 |
|
40497 | 81 |
instantiation lower_pd :: ("domain") below |
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
82 |
begin |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
83 |
|
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
84 |
definition |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
85 |
"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
|
86 |
|
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
87 |
instance .. |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
88 |
end |
25904 | 89 |
|
40497 | 90 |
instance lower_pd :: ("domain") po |
39984 | 91 |
using type_definition_lower_pd below_lower_pd_def |
92 |
by (rule lower_le.typedef_ideal_po) |
|
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
93 |
|
40497 | 94 |
instance lower_pd :: ("domain") cpo |
39984 | 95 |
using type_definition_lower_pd below_lower_pd_def |
96 |
by (rule lower_le.typedef_ideal_cpo) |
|
25904 | 97 |
|
98 |
definition |
|
99 |
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
|
100 |
"lower_principal t = Abs_lower_pd {u. u \<le>\<flat> t}" |
25904 | 101 |
|
30729
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents:
29990
diff
changeset
|
102 |
interpretation lower_pd: |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
103 |
ideal_completion lower_le lower_principal Rep_lower_pd |
39984 | 104 |
using type_definition_lower_pd below_lower_pd_def |
105 |
using lower_principal_def pd_basis_countable |
|
106 |
by (rule lower_le.typedef_ideal_completion) |
|
25904 | 107 |
|
27289 | 108 |
text {* Lower powerdomain is pointed *} |
25904 | 109 |
|
110 |
lemma lower_pd_minimal: "lower_principal (PDUnit compact_bot) \<sqsubseteq> ys" |
|
111 |
by (induct ys rule: lower_pd.principal_induct, simp, simp) |
|
112 |
||
40497 | 113 |
instance lower_pd :: ("domain") pcpo |
26927 | 114 |
by intro_classes (fast intro: lower_pd_minimal) |
25904 | 115 |
|
116 |
lemma inst_lower_pd_pcpo: "\<bottom> = lower_principal (PDUnit compact_bot)" |
|
117 |
by (rule lower_pd_minimal [THEN UU_I, symmetric]) |
|
118 |
||
119 |
||
26927 | 120 |
subsection {* Monadic unit and plus *} |
25904 | 121 |
|
122 |
definition |
|
123 |
lower_unit :: "'a \<rightarrow> 'a lower_pd" where |
|
124 |
"lower_unit = compact_basis.basis_fun (\<lambda>a. lower_principal (PDUnit a))" |
|
125 |
||
126 |
definition |
|
127 |
lower_plus :: "'a lower_pd \<rightarrow> 'a lower_pd \<rightarrow> 'a lower_pd" where |
|
128 |
"lower_plus = lower_pd.basis_fun (\<lambda>t. lower_pd.basis_fun (\<lambda>u. |
|
129 |
lower_principal (PDPlus t u)))" |
|
130 |
||
131 |
abbreviation |
|
132 |
lower_add :: "'a lower_pd \<Rightarrow> 'a lower_pd \<Rightarrow> 'a lower_pd" |
|
133 |
(infixl "+\<flat>" 65) where |
|
134 |
"xs +\<flat> ys == lower_plus\<cdot>xs\<cdot>ys" |
|
135 |
||
26927 | 136 |
syntax |
137 |
"_lower_pd" :: "args \<Rightarrow> 'a lower_pd" ("{_}\<flat>") |
|
138 |
||
139 |
translations |
|
140 |
"{x,xs}\<flat>" == "{x}\<flat> +\<flat> {xs}\<flat>" |
|
141 |
"{x}\<flat>" == "CONST lower_unit\<cdot>x" |
|
142 |
||
143 |
lemma lower_unit_Rep_compact_basis [simp]: |
|
144 |
"{Rep_compact_basis a}\<flat> = lower_principal (PDUnit a)" |
|
145 |
unfolding lower_unit_def |
|
27289 | 146 |
by (simp add: compact_basis.basis_fun_principal PDUnit_lower_mono) |
26927 | 147 |
|
25904 | 148 |
lemma lower_plus_principal [simp]: |
26927 | 149 |
"lower_principal t +\<flat> lower_principal u = lower_principal (PDPlus t u)" |
25904 | 150 |
unfolding lower_plus_def |
151 |
by (simp add: lower_pd.basis_fun_principal |
|
152 |
lower_pd.basis_fun_mono PDPlus_lower_mono) |
|
153 |
||
37770
cddb3106adb8
avoid explicit mandatory prefix markers when prefixes are mandatory implicitly
haftmann
parents:
36635
diff
changeset
|
154 |
interpretation lower_add: semilattice lower_add proof |
34973
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
155 |
fix xs ys zs :: "'a lower_pd" |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
156 |
show "(xs +\<flat> ys) +\<flat> zs = xs +\<flat> (ys +\<flat> zs)" |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
157 |
apply (induct xs ys arbitrary: zs rule: lower_pd.principal_induct2, simp, simp) |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
158 |
apply (rule_tac x=zs in lower_pd.principal_induct, simp) |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
159 |
apply (simp add: PDPlus_assoc) |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
160 |
done |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
161 |
show "xs +\<flat> ys = ys +\<flat> xs" |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
162 |
apply (induct xs ys rule: lower_pd.principal_induct2, simp, simp) |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
163 |
apply (simp add: PDPlus_commute) |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
164 |
done |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
165 |
show "xs +\<flat> xs = xs" |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
166 |
apply (induct xs rule: lower_pd.principal_induct, simp) |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
167 |
apply (simp add: PDPlus_absorb) |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
168 |
done |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
169 |
qed |
26927 | 170 |
|
34973
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
171 |
lemmas lower_plus_assoc = lower_add.assoc |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
172 |
lemmas lower_plus_commute = lower_add.commute |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
173 |
lemmas lower_plus_absorb = lower_add.idem |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
174 |
lemmas lower_plus_left_commute = lower_add.left_commute |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
175 |
lemmas lower_plus_left_absorb = lower_add.left_idem |
26927 | 176 |
|
29990
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
177 |
text {* Useful for @{text "simp add: lower_plus_ac"} *} |
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
178 |
lemmas lower_plus_ac = |
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
179 |
lower_plus_assoc lower_plus_commute lower_plus_left_commute |
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
180 |
|
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
181 |
text {* Useful for @{text "simp only: lower_plus_aci"} *} |
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
182 |
lemmas lower_plus_aci = |
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
183 |
lower_plus_ac lower_plus_absorb lower_plus_left_absorb |
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
184 |
|
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
|
185 |
lemma lower_plus_below1: "xs \<sqsubseteq> xs +\<flat> ys" |
27289 | 186 |
apply (induct xs ys rule: lower_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
|
187 |
apply (simp add: PDPlus_lower_le) |
25904 | 188 |
done |
189 |
||
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
|
190 |
lemma lower_plus_below2: "ys \<sqsubseteq> xs +\<flat> 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
|
191 |
by (subst lower_plus_commute, rule lower_plus_below1) |
25904 | 192 |
|
26927 | 193 |
lemma lower_plus_least: "\<lbrakk>xs \<sqsubseteq> zs; ys \<sqsubseteq> zs\<rbrakk> \<Longrightarrow> xs +\<flat> ys \<sqsubseteq> zs" |
25904 | 194 |
apply (subst lower_plus_absorb [of zs, symmetric]) |
195 |
apply (erule (1) monofun_cfun [OF monofun_cfun_arg]) |
|
196 |
done |
|
197 |
||
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
|
198 |
lemma lower_plus_below_iff: |
26927 | 199 |
"xs +\<flat> ys \<sqsubseteq> zs \<longleftrightarrow> xs \<sqsubseteq> zs \<and> ys \<sqsubseteq> zs" |
25904 | 200 |
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
|
201 |
apply (erule below_trans [OF lower_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
|
202 |
apply (erule below_trans [OF lower_plus_below2]) |
25904 | 203 |
apply (erule (1) lower_plus_least) |
204 |
done |
|
205 |
||
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
|
206 |
lemma lower_unit_below_plus_iff: |
26927 | 207 |
"{x}\<flat> \<sqsubseteq> ys +\<flat> zs \<longleftrightarrow> {x}\<flat> \<sqsubseteq> ys \<or> {x}\<flat> \<sqsubseteq> zs" |
39970
9023b897e67a
simplify proofs of powerdomain inequalities
Brian Huffman <brianh@cs.pdx.edu>
parents:
37770
diff
changeset
|
208 |
apply (induct x rule: compact_basis.principal_induct, simp) |
9023b897e67a
simplify proofs of powerdomain inequalities
Brian Huffman <brianh@cs.pdx.edu>
parents:
37770
diff
changeset
|
209 |
apply (induct ys rule: lower_pd.principal_induct, simp) |
9023b897e67a
simplify proofs of powerdomain inequalities
Brian Huffman <brianh@cs.pdx.edu>
parents:
37770
diff
changeset
|
210 |
apply (induct zs rule: lower_pd.principal_induct, simp) |
9023b897e67a
simplify proofs of powerdomain inequalities
Brian Huffman <brianh@cs.pdx.edu>
parents:
37770
diff
changeset
|
211 |
apply (simp add: lower_le_PDUnit_PDPlus_iff) |
25904 | 212 |
done |
213 |
||
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
|
214 |
lemma lower_unit_below_iff [simp]: "{x}\<flat> \<sqsubseteq> {y}\<flat> \<longleftrightarrow> x \<sqsubseteq> y" |
39970
9023b897e67a
simplify proofs of powerdomain inequalities
Brian Huffman <brianh@cs.pdx.edu>
parents:
37770
diff
changeset
|
215 |
apply (induct x rule: compact_basis.principal_induct, simp) |
9023b897e67a
simplify proofs of powerdomain inequalities
Brian Huffman <brianh@cs.pdx.edu>
parents:
37770
diff
changeset
|
216 |
apply (induct y rule: compact_basis.principal_induct, simp) |
9023b897e67a
simplify proofs of powerdomain inequalities
Brian Huffman <brianh@cs.pdx.edu>
parents:
37770
diff
changeset
|
217 |
apply simp |
26927 | 218 |
done |
219 |
||
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
|
220 |
lemmas lower_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
|
221 |
lower_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
|
222 |
lower_plus_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
|
223 |
lower_unit_below_plus_iff |
25904 | 224 |
|
26927 | 225 |
lemma lower_unit_eq_iff [simp]: "{x}\<flat> = {y}\<flat> \<longleftrightarrow> x = y" |
29990
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
226 |
by (simp add: po_eq_conv) |
26927 | 227 |
|
228 |
lemma lower_unit_strict [simp]: "{\<bottom>}\<flat> = \<bottom>" |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
229 |
using lower_unit_Rep_compact_basis [of compact_bot] |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
230 |
by (simp add: inst_lower_pd_pcpo) |
26927 | 231 |
|
40321
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
huffman
parents:
40002
diff
changeset
|
232 |
lemma lower_unit_bottom_iff [simp]: "{x}\<flat> = \<bottom> \<longleftrightarrow> x = \<bottom>" |
26927 | 233 |
unfolding lower_unit_strict [symmetric] by (rule lower_unit_eq_iff) |
234 |
||
40321
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
huffman
parents:
40002
diff
changeset
|
235 |
lemma lower_plus_bottom_iff [simp]: |
26927 | 236 |
"xs +\<flat> ys = \<bottom> \<longleftrightarrow> xs = \<bottom> \<and> ys = \<bottom>" |
237 |
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
|
238 |
apply (rule UU_I, erule subst, rule lower_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
|
239 |
apply (rule UU_I, erule subst, rule lower_plus_below2) |
26927 | 240 |
apply (rule lower_plus_absorb) |
241 |
done |
|
242 |
||
243 |
lemma lower_plus_strict1 [simp]: "\<bottom> +\<flat> ys = ys" |
|
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
|
244 |
apply (rule below_antisym [OF _ lower_plus_below2]) |
26927 | 245 |
apply (simp add: lower_plus_least) |
246 |
done |
|
247 |
||
248 |
lemma lower_plus_strict2 [simp]: "xs +\<flat> \<bottom> = xs" |
|
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
|
249 |
apply (rule below_antisym [OF _ lower_plus_below1]) |
26927 | 250 |
apply (simp add: lower_plus_least) |
251 |
done |
|
252 |
||
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
253 |
lemma compact_lower_unit: "compact x \<Longrightarrow> compact {x}\<flat>" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
254 |
by (auto dest!: compact_basis.compact_imp_principal) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
255 |
|
26927 | 256 |
lemma compact_lower_unit_iff [simp]: "compact {x}\<flat> \<longleftrightarrow> compact x" |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
257 |
apply (safe elim!: compact_lower_unit) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
258 |
apply (simp only: compact_def lower_unit_below_iff [symmetric]) |
40327 | 259 |
apply (erule adm_subst [OF cont_Rep_cfun2]) |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
260 |
done |
26927 | 261 |
|
262 |
lemma compact_lower_plus [simp]: |
|
263 |
"\<lbrakk>compact xs; compact ys\<rbrakk> \<Longrightarrow> compact (xs +\<flat> ys)" |
|
27289 | 264 |
by (auto dest!: lower_pd.compact_imp_principal) |
26927 | 265 |
|
25904 | 266 |
|
267 |
subsection {* Induction rules *} |
|
268 |
||
269 |
lemma lower_pd_induct1: |
|
270 |
assumes P: "adm P" |
|
26927 | 271 |
assumes unit: "\<And>x. P {x}\<flat>" |
25904 | 272 |
assumes insert: |
26927 | 273 |
"\<And>x ys. \<lbrakk>P {x}\<flat>; P ys\<rbrakk> \<Longrightarrow> P ({x}\<flat> +\<flat> ys)" |
25904 | 274 |
shows "P (xs::'a lower_pd)" |
27289 | 275 |
apply (induct xs rule: lower_pd.principal_induct, rule P) |
276 |
apply (induct_tac a rule: pd_basis_induct1) |
|
25904 | 277 |
apply (simp only: lower_unit_Rep_compact_basis [symmetric]) |
278 |
apply (rule unit) |
|
279 |
apply (simp only: lower_unit_Rep_compact_basis [symmetric] |
|
280 |
lower_plus_principal [symmetric]) |
|
281 |
apply (erule insert [OF unit]) |
|
282 |
done |
|
283 |
||
40576
fa5e0cacd5e7
declare {upper,lower,convex}_pd_induct as default induction rules
huffman
parents:
40497
diff
changeset
|
284 |
lemma lower_pd_induct |
fa5e0cacd5e7
declare {upper,lower,convex}_pd_induct as default induction rules
huffman
parents:
40497
diff
changeset
|
285 |
[case_names adm lower_unit lower_plus, induct type: lower_pd]: |
25904 | 286 |
assumes P: "adm P" |
26927 | 287 |
assumes unit: "\<And>x. P {x}\<flat>" |
288 |
assumes plus: "\<And>xs ys. \<lbrakk>P xs; P ys\<rbrakk> \<Longrightarrow> P (xs +\<flat> ys)" |
|
25904 | 289 |
shows "P (xs::'a lower_pd)" |
27289 | 290 |
apply (induct xs rule: lower_pd.principal_induct, rule P) |
291 |
apply (induct_tac a rule: pd_basis_induct) |
|
25904 | 292 |
apply (simp only: lower_unit_Rep_compact_basis [symmetric] unit) |
293 |
apply (simp only: lower_plus_principal [symmetric] plus) |
|
294 |
done |
|
295 |
||
296 |
||
297 |
subsection {* Monadic bind *} |
|
298 |
||
299 |
definition |
|
300 |
lower_bind_basis :: |
|
301 |
"'a pd_basis \<Rightarrow> ('a \<rightarrow> 'b lower_pd) \<rightarrow> 'b lower_pd" where |
|
302 |
"lower_bind_basis = fold_pd |
|
303 |
(\<lambda>a. \<Lambda> f. f\<cdot>(Rep_compact_basis a)) |
|
26927 | 304 |
(\<lambda>x y. \<Lambda> f. x\<cdot>f +\<flat> y\<cdot>f)" |
25904 | 305 |
|
26927 | 306 |
lemma ACI_lower_bind: |
36635
080b755377c0
locale predicates of classes carry a mandatory "class" prefix
haftmann
parents:
35901
diff
changeset
|
307 |
"class.ab_semigroup_idem_mult (\<lambda>x y. \<Lambda> f. x\<cdot>f +\<flat> y\<cdot>f)" |
25904 | 308 |
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
|
309 |
apply (simp add: lower_plus_assoc) |
25904 | 310 |
apply (simp add: lower_plus_commute) |
29990
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
311 |
apply (simp add: eta_cfun) |
25904 | 312 |
done |
313 |
||
314 |
lemma lower_bind_basis_simps [simp]: |
|
315 |
"lower_bind_basis (PDUnit a) = |
|
316 |
(\<Lambda> f. f\<cdot>(Rep_compact_basis a))" |
|
317 |
"lower_bind_basis (PDPlus t u) = |
|
26927 | 318 |
(\<Lambda> f. lower_bind_basis t\<cdot>f +\<flat> lower_bind_basis u\<cdot>f)" |
25904 | 319 |
unfolding lower_bind_basis_def |
320 |
apply - |
|
26927 | 321 |
apply (rule fold_pd_PDUnit [OF ACI_lower_bind]) |
322 |
apply (rule fold_pd_PDPlus [OF ACI_lower_bind]) |
|
25904 | 323 |
done |
324 |
||
325 |
lemma lower_bind_basis_mono: |
|
326 |
"t \<le>\<flat> u \<Longrightarrow> lower_bind_basis t \<sqsubseteq> lower_bind_basis u" |
|
40002
c5b5f7a3a3b1
new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents:
39989
diff
changeset
|
327 |
unfolding cfun_below_iff |
25904 | 328 |
apply (erule lower_le_induct, safe) |
27289 | 329 |
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
|
330 |
apply (simp add: rev_below_trans [OF lower_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
|
331 |
apply (simp add: lower_plus_below_iff) |
25904 | 332 |
done |
333 |
||
334 |
definition |
|
335 |
lower_bind :: "'a lower_pd \<rightarrow> ('a \<rightarrow> 'b lower_pd) \<rightarrow> 'b lower_pd" where |
|
336 |
"lower_bind = lower_pd.basis_fun lower_bind_basis" |
|
337 |
||
338 |
lemma lower_bind_principal [simp]: |
|
339 |
"lower_bind\<cdot>(lower_principal t) = lower_bind_basis t" |
|
340 |
unfolding lower_bind_def |
|
341 |
apply (rule lower_pd.basis_fun_principal) |
|
342 |
apply (erule lower_bind_basis_mono) |
|
343 |
done |
|
344 |
||
345 |
lemma lower_bind_unit [simp]: |
|
26927 | 346 |
"lower_bind\<cdot>{x}\<flat>\<cdot>f = f\<cdot>x" |
27289 | 347 |
by (induct x rule: compact_basis.principal_induct, simp, simp) |
25904 | 348 |
|
349 |
lemma lower_bind_plus [simp]: |
|
26927 | 350 |
"lower_bind\<cdot>(xs +\<flat> ys)\<cdot>f = lower_bind\<cdot>xs\<cdot>f +\<flat> lower_bind\<cdot>ys\<cdot>f" |
27289 | 351 |
by (induct xs ys rule: lower_pd.principal_induct2, simp, simp, simp) |
25904 | 352 |
|
353 |
lemma lower_bind_strict [simp]: "lower_bind\<cdot>\<bottom>\<cdot>f = f\<cdot>\<bottom>" |
|
354 |
unfolding lower_unit_strict [symmetric] by (rule lower_bind_unit) |
|
355 |
||
356 |
||
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
357 |
subsection {* Map *} |
25904 | 358 |
|
359 |
definition |
|
360 |
lower_map :: "('a \<rightarrow> 'b) \<rightarrow> 'a lower_pd \<rightarrow> 'b lower_pd" where |
|
26927 | 361 |
"lower_map = (\<Lambda> f xs. lower_bind\<cdot>xs\<cdot>(\<Lambda> x. {f\<cdot>x}\<flat>))" |
25904 | 362 |
|
363 |
lemma lower_map_unit [simp]: |
|
26927 | 364 |
"lower_map\<cdot>f\<cdot>{x}\<flat> = {f\<cdot>x}\<flat>" |
25904 | 365 |
unfolding lower_map_def by simp |
366 |
||
367 |
lemma lower_map_plus [simp]: |
|
26927 | 368 |
"lower_map\<cdot>f\<cdot>(xs +\<flat> ys) = lower_map\<cdot>f\<cdot>xs +\<flat> lower_map\<cdot>f\<cdot>ys" |
25904 | 369 |
unfolding lower_map_def by simp |
370 |
||
40577 | 371 |
lemma lower_map_bottom [simp]: "lower_map\<cdot>f\<cdot>\<bottom> = {f\<cdot>\<bottom>}\<flat>" |
372 |
unfolding lower_map_def by simp |
|
373 |
||
25904 | 374 |
lemma lower_map_ident: "lower_map\<cdot>(\<Lambda> x. x)\<cdot>xs = xs" |
375 |
by (induct xs rule: lower_pd_induct, simp_all) |
|
376 |
||
33808 | 377 |
lemma lower_map_ID: "lower_map\<cdot>ID = ID" |
40002
c5b5f7a3a3b1
new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents:
39989
diff
changeset
|
378 |
by (simp add: cfun_eq_iff ID_def lower_map_ident) |
33808 | 379 |
|
25904 | 380 |
lemma lower_map_map: |
381 |
"lower_map\<cdot>f\<cdot>(lower_map\<cdot>g\<cdot>xs) = lower_map\<cdot>(\<Lambda> x. f\<cdot>(g\<cdot>x))\<cdot>xs" |
|
382 |
by (induct xs rule: lower_pd_induct, simp_all) |
|
383 |
||
33585
8d39394fe5cf
ep_pair and deflation lemmas for powerdomain map functions
huffman
parents:
31076
diff
changeset
|
384 |
lemma ep_pair_lower_map: "ep_pair e p \<Longrightarrow> ep_pair (lower_map\<cdot>e) (lower_map\<cdot>p)" |
8d39394fe5cf
ep_pair and deflation lemmas for powerdomain map functions
huffman
parents:
31076
diff
changeset
|
385 |
apply default |
8d39394fe5cf
ep_pair and deflation lemmas for powerdomain map functions
huffman
parents:
31076
diff
changeset
|
386 |
apply (induct_tac x rule: lower_pd_induct, simp_all add: ep_pair.e_inverse) |
35901
12f09bf2c77f
fix LaTeX overfull hbox warnings in HOLCF document
huffman
parents:
34973
diff
changeset
|
387 |
apply (induct_tac y rule: lower_pd_induct) |
12f09bf2c77f
fix LaTeX overfull hbox warnings in HOLCF document
huffman
parents:
34973
diff
changeset
|
388 |
apply (simp_all add: ep_pair.e_p_below monofun_cfun) |
33585
8d39394fe5cf
ep_pair and deflation lemmas for powerdomain map functions
huffman
parents:
31076
diff
changeset
|
389 |
done |
8d39394fe5cf
ep_pair and deflation lemmas for powerdomain map functions
huffman
parents:
31076
diff
changeset
|
390 |
|
8d39394fe5cf
ep_pair and deflation lemmas for powerdomain map functions
huffman
parents:
31076
diff
changeset
|
391 |
lemma deflation_lower_map: "deflation d \<Longrightarrow> deflation (lower_map\<cdot>d)" |
8d39394fe5cf
ep_pair and deflation lemmas for powerdomain map functions
huffman
parents:
31076
diff
changeset
|
392 |
apply default |
8d39394fe5cf
ep_pair and deflation lemmas for powerdomain map functions
huffman
parents:
31076
diff
changeset
|
393 |
apply (induct_tac x rule: lower_pd_induct, simp_all add: deflation.idem) |
35901
12f09bf2c77f
fix LaTeX overfull hbox warnings in HOLCF document
huffman
parents:
34973
diff
changeset
|
394 |
apply (induct_tac x rule: lower_pd_induct) |
12f09bf2c77f
fix LaTeX overfull hbox warnings in HOLCF document
huffman
parents:
34973
diff
changeset
|
395 |
apply (simp_all add: deflation.below monofun_cfun) |
33585
8d39394fe5cf
ep_pair and deflation lemmas for powerdomain map functions
huffman
parents:
31076
diff
changeset
|
396 |
done |
8d39394fe5cf
ep_pair and deflation lemmas for powerdomain map functions
huffman
parents:
31076
diff
changeset
|
397 |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
398 |
(* FIXME: long proof! *) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
399 |
lemma finite_deflation_lower_map: |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
400 |
assumes "finite_deflation d" shows "finite_deflation (lower_map\<cdot>d)" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
401 |
proof (rule finite_deflation_intro) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
402 |
interpret d: finite_deflation d by fact |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
403 |
have "deflation d" by fact |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
404 |
thus "deflation (lower_map\<cdot>d)" by (rule deflation_lower_map) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
405 |
have "finite (range (\<lambda>x. d\<cdot>x))" by (rule d.finite_range) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
406 |
hence "finite (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
407 |
by (rule finite_vimageI, simp add: inj_on_def Rep_compact_basis_inject) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
408 |
hence "finite (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x)))" by simp |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
409 |
hence "finite (Rep_pd_basis -` (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))))" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
410 |
by (rule finite_vimageI, simp add: inj_on_def Rep_pd_basis_inject) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
411 |
hence *: "finite (lower_principal ` Rep_pd_basis -` (Pow (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))))" by simp |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
412 |
hence "finite (range (\<lambda>xs. lower_map\<cdot>d\<cdot>xs))" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
413 |
apply (rule rev_finite_subset) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
414 |
apply clarsimp |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
415 |
apply (induct_tac xs rule: lower_pd.principal_induct) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
416 |
apply (simp add: adm_mem_finite *) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
417 |
apply (rename_tac t, induct_tac t rule: pd_basis_induct) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
418 |
apply (simp only: lower_unit_Rep_compact_basis [symmetric] lower_map_unit) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
419 |
apply simp |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
420 |
apply (subgoal_tac "\<exists>b. d\<cdot>(Rep_compact_basis a) = Rep_compact_basis b") |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
421 |
apply clarsimp |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
422 |
apply (rule imageI) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
423 |
apply (rule vimageI2) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
424 |
apply (simp add: Rep_PDUnit) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
425 |
apply (rule range_eqI) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
426 |
apply (erule sym) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
427 |
apply (rule exI) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
428 |
apply (rule Abs_compact_basis_inverse [symmetric]) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
429 |
apply (simp add: d.compact) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
430 |
apply (simp only: lower_plus_principal [symmetric] lower_map_plus) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
431 |
apply clarsimp |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
432 |
apply (rule imageI) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
433 |
apply (rule vimageI2) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
434 |
apply (simp add: Rep_PDPlus) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
435 |
done |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
436 |
thus "finite {xs. lower_map\<cdot>d\<cdot>xs = xs}" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
437 |
by (rule finite_range_imp_finite_fixes) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
438 |
qed |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
439 |
|
40497 | 440 |
subsection {* Lower powerdomain is a domain *} |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
441 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
442 |
definition |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
443 |
lower_approx :: "nat \<Rightarrow> udom lower_pd \<rightarrow> udom lower_pd" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
444 |
where |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
445 |
"lower_approx = (\<lambda>i. lower_map\<cdot>(udom_approx i))" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
446 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
447 |
lemma lower_approx: "approx_chain lower_approx" |
40484
768f7e264e2b
reorganize Bifinite.thy; simplify some proofs related to bifinite class instances
huffman
parents:
40436
diff
changeset
|
448 |
using lower_map_ID finite_deflation_lower_map |
768f7e264e2b
reorganize Bifinite.thy; simplify some proofs related to bifinite class instances
huffman
parents:
40436
diff
changeset
|
449 |
unfolding lower_approx_def by (rule approx_chain_lemma1) |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
450 |
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39986
diff
changeset
|
451 |
definition lower_defl :: "defl \<rightarrow> defl" |
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39986
diff
changeset
|
452 |
where "lower_defl = defl_fun1 lower_approx lower_map" |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
453 |
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39986
diff
changeset
|
454 |
lemma cast_lower_defl: |
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39986
diff
changeset
|
455 |
"cast\<cdot>(lower_defl\<cdot>A) = |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
456 |
udom_emb lower_approx oo lower_map\<cdot>(cast\<cdot>A) oo udom_prj lower_approx" |
40484
768f7e264e2b
reorganize Bifinite.thy; simplify some proofs related to bifinite class instances
huffman
parents:
40436
diff
changeset
|
457 |
using lower_approx finite_deflation_lower_map |
768f7e264e2b
reorganize Bifinite.thy; simplify some proofs related to bifinite class instances
huffman
parents:
40436
diff
changeset
|
458 |
unfolding lower_defl_def by (rule cast_defl_fun1) |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
459 |
|
40497 | 460 |
instantiation lower_pd :: ("domain") liftdomain |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
461 |
begin |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
462 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
463 |
definition |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
464 |
"emb = udom_emb lower_approx oo lower_map\<cdot>emb" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
465 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
466 |
definition |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
467 |
"prj = lower_map\<cdot>prj oo udom_prj lower_approx" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
468 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
469 |
definition |
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39986
diff
changeset
|
470 |
"defl (t::'a lower_pd itself) = lower_defl\<cdot>DEFL('a)" |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
471 |
|
40491
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40484
diff
changeset
|
472 |
definition |
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40484
diff
changeset
|
473 |
"(liftemb :: 'a lower_pd u \<rightarrow> udom) = udom_emb u_approx oo u_map\<cdot>emb" |
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40484
diff
changeset
|
474 |
|
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40484
diff
changeset
|
475 |
definition |
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40484
diff
changeset
|
476 |
"(liftprj :: udom \<rightarrow> 'a lower_pd u) = u_map\<cdot>prj oo udom_prj u_approx" |
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40484
diff
changeset
|
477 |
|
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40484
diff
changeset
|
478 |
definition |
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40484
diff
changeset
|
479 |
"liftdefl (t::'a lower_pd itself) = u_defl\<cdot>DEFL('a lower_pd)" |
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40484
diff
changeset
|
480 |
|
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40484
diff
changeset
|
481 |
instance |
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
huffman
parents:
40484
diff
changeset
|
482 |
using liftemb_lower_pd_def liftprj_lower_pd_def liftdefl_lower_pd_def |
40494
db8a09daba7b
add class liftdomain, for bifinite domains where DEFL('a u) = u_defl('a)
huffman
parents:
40491
diff
changeset
|
483 |
proof (rule liftdomain_class_intro) |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
484 |
show "ep_pair emb (prj :: udom \<rightarrow> 'a lower_pd)" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
485 |
unfolding emb_lower_pd_def prj_lower_pd_def |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
486 |
using ep_pair_udom [OF lower_approx] |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
487 |
by (intro ep_pair_comp ep_pair_lower_map ep_pair_emb_prj) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
488 |
next |
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39986
diff
changeset
|
489 |
show "cast\<cdot>DEFL('a lower_pd) = emb oo (prj :: udom \<rightarrow> 'a lower_pd)" |
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39986
diff
changeset
|
490 |
unfolding emb_lower_pd_def prj_lower_pd_def defl_lower_pd_def cast_lower_defl |
40002
c5b5f7a3a3b1
new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents:
39989
diff
changeset
|
491 |
by (simp add: cast_DEFL oo_def cfun_eq_iff lower_map_map) |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
492 |
qed |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
493 |
|
25904 | 494 |
end |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
495 |
|
39989
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39986
diff
changeset
|
496 |
lemma DEFL_lower: "DEFL('a lower_pd) = lower_defl\<cdot>DEFL('a)" |
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
huffman
parents:
39986
diff
changeset
|
497 |
by (rule defl_lower_pd_def) |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
498 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
499 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
500 |
subsection {* Join *} |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
501 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
502 |
definition |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
503 |
lower_join :: "'a lower_pd lower_pd \<rightarrow> 'a lower_pd" where |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
504 |
"lower_join = (\<Lambda> xss. lower_bind\<cdot>xss\<cdot>(\<Lambda> xs. xs))" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
505 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
506 |
lemma lower_join_unit [simp]: |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
507 |
"lower_join\<cdot>{xs}\<flat> = xs" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
508 |
unfolding lower_join_def by simp |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
509 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
510 |
lemma lower_join_plus [simp]: |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
511 |
"lower_join\<cdot>(xss +\<flat> yss) = lower_join\<cdot>xss +\<flat> lower_join\<cdot>yss" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
512 |
unfolding lower_join_def by simp |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
513 |
|
40577 | 514 |
lemma lower_join_bottom [simp]: "lower_join\<cdot>\<bottom> = \<bottom>" |
515 |
unfolding lower_join_def by simp |
|
516 |
||
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
517 |
lemma lower_join_map_unit: |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
518 |
"lower_join\<cdot>(lower_map\<cdot>lower_unit\<cdot>xs) = xs" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
519 |
by (induct xs rule: lower_pd_induct, simp_all) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
520 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
521 |
lemma lower_join_map_join: |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
522 |
"lower_join\<cdot>(lower_map\<cdot>lower_join\<cdot>xsss) = lower_join\<cdot>(lower_join\<cdot>xsss)" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
523 |
by (induct xsss rule: lower_pd_induct, simp_all) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
524 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
525 |
lemma lower_join_map_map: |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
526 |
"lower_join\<cdot>(lower_map\<cdot>(lower_map\<cdot>f)\<cdot>xss) = |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
527 |
lower_map\<cdot>f\<cdot>(lower_join\<cdot>xss)" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
528 |
by (induct xss rule: lower_pd_induct, simp_all) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
529 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
530 |
end |