author | wenzelm |
Wed, 25 Jun 2025 16:35:25 +0200 | |
changeset 82768 | 8f866fd6fae1 |
parent 81620 | 2cb49d09f059 |
permissions | -rw-r--r-- |
42151 | 1 |
(* Title: HOL/HOLCF/ConvexPD.thy |
25904 | 2 |
Author: Brian Huffman |
3 |
*) |
|
4 |
||
62175 | 5 |
section \<open>Convex powerdomain\<close> |
25904 | 6 |
|
7 |
theory ConvexPD |
|
8 |
imports UpperPD LowerPD |
|
9 |
begin |
|
10 |
||
62175 | 11 |
subsection \<open>Basis preorder\<close> |
25904 | 12 |
|
13 |
definition |
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
14 |
convex_le :: "'a::bifinite pd_basis \<Rightarrow> 'a pd_basis \<Rightarrow> bool" (infix \<open>\<le>\<natural>\<close> 50) where |
25904 | 15 |
"convex_le = (\<lambda>u v. u \<le>\<sharp> v \<and> u \<le>\<flat> v)" |
16 |
||
17 |
lemma convex_le_refl [simp]: "t \<le>\<natural> t" |
|
18 |
unfolding convex_le_def by (fast intro: upper_le_refl lower_le_refl) |
|
19 |
||
20 |
lemma convex_le_trans: "\<lbrakk>t \<le>\<natural> u; u \<le>\<natural> v\<rbrakk> \<Longrightarrow> t \<le>\<natural> v" |
|
21 |
unfolding convex_le_def by (fast intro: upper_le_trans lower_le_trans) |
|
22 |
||
30729
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents:
29990
diff
changeset
|
23 |
interpretation convex_le: preorder convex_le |
25904 | 24 |
by (rule preorder.intro, rule convex_le_refl, rule convex_le_trans) |
25 |
||
26 |
lemma upper_le_minimal [simp]: "PDUnit compact_bot \<le>\<natural> t" |
|
27 |
unfolding convex_le_def Rep_PDUnit by simp |
|
28 |
||
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
29 |
lemma PDUnit_convex_mono: "x \<sqsubseteq> y \<Longrightarrow> PDUnit x \<le>\<natural> PDUnit y" |
25904 | 30 |
unfolding convex_le_def by (fast intro: PDUnit_upper_mono PDUnit_lower_mono) |
31 |
||
32 |
lemma PDPlus_convex_mono: "\<lbrakk>s \<le>\<natural> t; u \<le>\<natural> v\<rbrakk> \<Longrightarrow> PDPlus s u \<le>\<natural> PDPlus t v" |
|
33 |
unfolding convex_le_def by (fast intro: PDPlus_upper_mono PDPlus_lower_mono) |
|
34 |
||
35 |
lemma convex_le_PDUnit_PDUnit_iff [simp]: |
|
40436 | 36 |
"(PDUnit a \<le>\<natural> PDUnit b) = (a \<sqsubseteq> b)" |
25904 | 37 |
unfolding convex_le_def upper_le_def lower_le_def Rep_PDUnit by fast |
38 |
||
39 |
lemma convex_le_PDUnit_lemma1: |
|
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
40 |
"(PDUnit a \<le>\<natural> t) = (\<forall>b\<in>Rep_pd_basis t. a \<sqsubseteq> b)" |
25904 | 41 |
unfolding convex_le_def upper_le_def lower_le_def Rep_PDUnit |
42 |
using Rep_pd_basis_nonempty [of t, folded ex_in_conv] by fast |
|
43 |
||
44 |
lemma convex_le_PDUnit_PDPlus_iff [simp]: |
|
45 |
"(PDUnit a \<le>\<natural> PDPlus t u) = (PDUnit a \<le>\<natural> t \<and> PDUnit a \<le>\<natural> u)" |
|
46 |
unfolding convex_le_PDUnit_lemma1 Rep_PDPlus by fast |
|
47 |
||
48 |
lemma convex_le_PDUnit_lemma2: |
|
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
49 |
"(t \<le>\<natural> PDUnit b) = (\<forall>a\<in>Rep_pd_basis t. a \<sqsubseteq> b)" |
25904 | 50 |
unfolding convex_le_def upper_le_def lower_le_def Rep_PDUnit |
51 |
using Rep_pd_basis_nonempty [of t, folded ex_in_conv] by fast |
|
52 |
||
53 |
lemma convex_le_PDPlus_PDUnit_iff [simp]: |
|
54 |
"(PDPlus t u \<le>\<natural> PDUnit a) = (t \<le>\<natural> PDUnit a \<and> u \<le>\<natural> PDUnit a)" |
|
55 |
unfolding convex_le_PDUnit_lemma2 Rep_PDPlus by fast |
|
56 |
||
57 |
lemma convex_le_PDPlus_lemma: |
|
58 |
assumes z: "PDPlus t u \<le>\<natural> z" |
|
59 |
shows "\<exists>v w. z = PDPlus v w \<and> t \<le>\<natural> v \<and> u \<le>\<natural> w" |
|
60 |
proof (intro exI conjI) |
|
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
61 |
let ?A = "{b\<in>Rep_pd_basis z. \<exists>a\<in>Rep_pd_basis t. a \<sqsubseteq> b}" |
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
62 |
let ?B = "{b\<in>Rep_pd_basis z. \<exists>a\<in>Rep_pd_basis u. a \<sqsubseteq> b}" |
25904 | 63 |
let ?v = "Abs_pd_basis ?A" |
64 |
let ?w = "Abs_pd_basis ?B" |
|
65 |
have Rep_v: "Rep_pd_basis ?v = ?A" |
|
66 |
apply (rule Abs_pd_basis_inverse) |
|
67 |
apply (rule Rep_pd_basis_nonempty [of t, folded ex_in_conv, THEN exE]) |
|
68 |
apply (cut_tac z, simp only: convex_le_def lower_le_def, clarify) |
|
69 |
apply (drule_tac x=x in bspec, simp add: Rep_PDPlus, erule bexE) |
|
70 |
apply (simp add: pd_basis_def) |
|
71 |
apply fast |
|
72 |
done |
|
73 |
have Rep_w: "Rep_pd_basis ?w = ?B" |
|
74 |
apply (rule Abs_pd_basis_inverse) |
|
75 |
apply (rule Rep_pd_basis_nonempty [of u, folded ex_in_conv, THEN exE]) |
|
76 |
apply (cut_tac z, simp only: convex_le_def lower_le_def, clarify) |
|
77 |
apply (drule_tac x=x in bspec, simp add: Rep_PDPlus, erule bexE) |
|
78 |
apply (simp add: pd_basis_def) |
|
79 |
apply fast |
|
80 |
done |
|
81 |
show "z = PDPlus ?v ?w" |
|
82 |
apply (insert z) |
|
83 |
apply (simp add: convex_le_def, erule conjE) |
|
84 |
apply (simp add: Rep_pd_basis_inject [symmetric] Rep_PDPlus) |
|
85 |
apply (simp add: Rep_v Rep_w) |
|
86 |
apply (rule equalityI) |
|
87 |
apply (rule subsetI) |
|
88 |
apply (simp only: upper_le_def) |
|
89 |
apply (drule (1) bspec, erule bexE) |
|
90 |
apply (simp add: Rep_PDPlus) |
|
91 |
apply fast |
|
92 |
apply fast |
|
93 |
done |
|
94 |
show "t \<le>\<natural> ?v" "u \<le>\<natural> ?w" |
|
81620 | 95 |
using z by (simp_all add: convex_le_def upper_le_def lower_le_def Rep_PDPlus Rep_v Rep_w) fast+ |
25904 | 96 |
qed |
97 |
||
98 |
lemma convex_le_induct [induct set: convex_le]: |
|
99 |
assumes le: "t \<le>\<natural> u" |
|
100 |
assumes 2: "\<And>t u v. \<lbrakk>P t u; P u v\<rbrakk> \<Longrightarrow> P t v" |
|
26420
57a626f64875
make preorder locale into a superclass of class po
huffman
parents:
26407
diff
changeset
|
101 |
assumes 3: "\<And>a b. a \<sqsubseteq> b \<Longrightarrow> P (PDUnit a) (PDUnit b)" |
25904 | 102 |
assumes 4: "\<And>t u v w. \<lbrakk>P t v; P u w\<rbrakk> \<Longrightarrow> P (PDPlus t u) (PDPlus v w)" |
103 |
shows "P t u" |
|
81620 | 104 |
using le |
105 |
proof (induct t arbitrary: u rule: pd_basis_induct) |
|
106 |
case (PDUnit a) |
|
107 |
then show ?case |
|
108 |
proof (induct u rule: pd_basis_induct1) |
|
109 |
case (PDUnit b) |
|
110 |
then show ?case by (simp add: 3) |
|
111 |
next |
|
112 |
case (PDPlus b t) |
|
113 |
have "P (PDPlus (PDUnit a) (PDUnit a)) (PDPlus (PDUnit b) t)" |
|
114 |
by (rule 4 [OF 3]) (use PDPlus in simp_all) |
|
115 |
then show ?case by (simp add: PDPlus_absorb) |
|
116 |
qed |
|
117 |
next |
|
118 |
case PDPlus |
|
119 |
from PDPlus(1,2) show ?case |
|
120 |
using convex_le_PDPlus_lemma [OF PDPlus(3)] by (auto simp add: 4) |
|
121 |
qed |
|
25904 | 122 |
|
123 |
||
62175 | 124 |
subsection \<open>Type definition\<close> |
25904 | 125 |
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
126 |
typedef 'a::bifinite convex_pd (\<open>(\<open>notation=\<open>postfix convex_pd\<close>\<close>'(_')\<natural>)\<close>) = |
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
127 |
"{S::'a pd_basis set. convex_le.ideal S}" |
40888
28cd51cff70c
reformulate lemma preorder.ex_ideal, and use it for typedefs
huffman
parents:
40774
diff
changeset
|
128 |
by (rule convex_le.ex_ideal) |
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
129 |
|
41288
a19edebad961
powerdomain theories require class 'bifinite' instead of 'domain'
huffman
parents:
41287
diff
changeset
|
130 |
instantiation convex_pd :: (bifinite) below |
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
131 |
begin |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
132 |
|
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
133 |
definition |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
134 |
"x \<sqsubseteq> y \<longleftrightarrow> Rep_convex_pd x \<subseteq> Rep_convex_pd y" |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
135 |
|
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
136 |
instance .. |
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
137 |
end |
25904 | 138 |
|
41288
a19edebad961
powerdomain theories require class 'bifinite' instead of 'domain'
huffman
parents:
41287
diff
changeset
|
139 |
instance convex_pd :: (bifinite) po |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
140 |
using type_definition_convex_pd below_convex_pd_def |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
141 |
by (rule convex_le.typedef_ideal_po) |
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
142 |
|
41288
a19edebad961
powerdomain theories require class 'bifinite' instead of 'domain'
huffman
parents:
41287
diff
changeset
|
143 |
instance convex_pd :: (bifinite) cpo |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
144 |
using type_definition_convex_pd below_convex_pd_def |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
145 |
by (rule convex_le.typedef_ideal_cpo) |
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
146 |
|
25904 | 147 |
definition |
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
148 |
convex_principal :: "'a::bifinite pd_basis \<Rightarrow> 'a convex_pd" where |
27373
5794a0e3e26c
remove cset theory; define ideal completions using typedef instead of cpodef
huffman
parents:
27310
diff
changeset
|
149 |
"convex_principal t = Abs_convex_pd {u. u \<le>\<natural> t}" |
25904 | 150 |
|
30729
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents:
29990
diff
changeset
|
151 |
interpretation convex_pd: |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
152 |
ideal_completion convex_le convex_principal Rep_convex_pd |
39984 | 153 |
using type_definition_convex_pd below_convex_pd_def |
154 |
using convex_principal_def pd_basis_countable |
|
155 |
by (rule convex_le.typedef_ideal_completion) |
|
25904 | 156 |
|
62175 | 157 |
text \<open>Convex powerdomain is pointed\<close> |
25904 | 158 |
|
159 |
lemma convex_pd_minimal: "convex_principal (PDUnit compact_bot) \<sqsubseteq> ys" |
|
160 |
by (induct ys rule: convex_pd.principal_induct, simp, simp) |
|
161 |
||
41288
a19edebad961
powerdomain theories require class 'bifinite' instead of 'domain'
huffman
parents:
41287
diff
changeset
|
162 |
instance convex_pd :: (bifinite) pcpo |
26927 | 163 |
by intro_classes (fast intro: convex_pd_minimal) |
25904 | 164 |
|
165 |
lemma inst_convex_pd_pcpo: "\<bottom> = convex_principal (PDUnit compact_bot)" |
|
41430
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
huffman
parents:
41402
diff
changeset
|
166 |
by (rule convex_pd_minimal [THEN bottomI, symmetric]) |
25904 | 167 |
|
168 |
||
62175 | 169 |
subsection \<open>Monadic unit and plus\<close> |
25904 | 170 |
|
171 |
definition |
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
172 |
convex_unit :: "'a::bifinite \<rightarrow> 'a convex_pd" where |
41394
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents:
41289
diff
changeset
|
173 |
"convex_unit = compact_basis.extension (\<lambda>a. convex_principal (PDUnit a))" |
25904 | 174 |
|
175 |
definition |
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
176 |
convex_plus :: "'a::bifinite convex_pd \<rightarrow> 'a convex_pd \<rightarrow> 'a convex_pd" where |
41394
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents:
41289
diff
changeset
|
177 |
"convex_plus = convex_pd.extension (\<lambda>t. convex_pd.extension (\<lambda>u. |
25904 | 178 |
convex_principal (PDPlus t u)))" |
179 |
||
180 |
abbreviation |
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
181 |
convex_add :: "'a::bifinite convex_pd \<Rightarrow> 'a convex_pd \<Rightarrow> 'a convex_pd" |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80786
diff
changeset
|
182 |
(infixl \<open>\<union>\<natural>\<close> 65) where |
41399
ad093e4638e2
changed syntax of powerdomain binary union operators
huffman
parents:
41394
diff
changeset
|
183 |
"xs \<union>\<natural> ys == convex_plus\<cdot>xs\<cdot>ys" |
25904 | 184 |
|
26927 | 185 |
syntax |
81089
8042869c2072
clarified syntax: prefer nonterminal "args", use outer block (with indent);
wenzelm
parents:
80914
diff
changeset
|
186 |
"_convex_pd" :: "args \<Rightarrow> logic" (\<open>(\<open>indent=1 notation=\<open>mixfix convex_pd enumeration\<close>\<close>{_}\<natural>)\<close>) |
26927 | 187 |
translations |
41399
ad093e4638e2
changed syntax of powerdomain binary union operators
huffman
parents:
41394
diff
changeset
|
188 |
"{x,xs}\<natural>" == "{x}\<natural> \<union>\<natural> {xs}\<natural>" |
26927 | 189 |
"{x}\<natural>" == "CONST convex_unit\<cdot>x" |
190 |
||
191 |
lemma convex_unit_Rep_compact_basis [simp]: |
|
192 |
"{Rep_compact_basis a}\<natural> = convex_principal (PDUnit a)" |
|
193 |
unfolding convex_unit_def |
|
41394
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents:
41289
diff
changeset
|
194 |
by (simp add: compact_basis.extension_principal PDUnit_convex_mono) |
26927 | 195 |
|
25904 | 196 |
lemma convex_plus_principal [simp]: |
41399
ad093e4638e2
changed syntax of powerdomain binary union operators
huffman
parents:
41394
diff
changeset
|
197 |
"convex_principal t \<union>\<natural> convex_principal u = convex_principal (PDPlus t u)" |
25904 | 198 |
unfolding convex_plus_def |
41394
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents:
41289
diff
changeset
|
199 |
by (simp add: convex_pd.extension_principal |
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents:
41289
diff
changeset
|
200 |
convex_pd.extension_mono PDPlus_convex_mono) |
25904 | 201 |
|
37770
cddb3106adb8
avoid explicit mandatory prefix markers when prefixes are mandatory implicitly
haftmann
parents:
36635
diff
changeset
|
202 |
interpretation convex_add: semilattice convex_add proof |
34973
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
203 |
fix xs ys zs :: "'a convex_pd" |
41399
ad093e4638e2
changed syntax of powerdomain binary union operators
huffman
parents:
41394
diff
changeset
|
204 |
show "(xs \<union>\<natural> ys) \<union>\<natural> zs = xs \<union>\<natural> (ys \<union>\<natural> zs)" |
41402
b647212cee03
remove lemma ideal_completion.principal_induct2, use principal_induct twice instead
huffman
parents:
41399
diff
changeset
|
205 |
apply (induct xs rule: convex_pd.principal_induct, simp) |
b647212cee03
remove lemma ideal_completion.principal_induct2, use principal_induct twice instead
huffman
parents:
41399
diff
changeset
|
206 |
apply (induct ys rule: convex_pd.principal_induct, simp) |
b647212cee03
remove lemma ideal_completion.principal_induct2, use principal_induct twice instead
huffman
parents:
41399
diff
changeset
|
207 |
apply (induct zs rule: convex_pd.principal_induct, simp) |
34973
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
208 |
apply (simp add: PDPlus_assoc) |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
209 |
done |
41399
ad093e4638e2
changed syntax of powerdomain binary union operators
huffman
parents:
41394
diff
changeset
|
210 |
show "xs \<union>\<natural> ys = ys \<union>\<natural> xs" |
41402
b647212cee03
remove lemma ideal_completion.principal_induct2, use principal_induct twice instead
huffman
parents:
41399
diff
changeset
|
211 |
apply (induct xs rule: convex_pd.principal_induct, simp) |
b647212cee03
remove lemma ideal_completion.principal_induct2, use principal_induct twice instead
huffman
parents:
41399
diff
changeset
|
212 |
apply (induct ys rule: convex_pd.principal_induct, simp) |
34973
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
213 |
apply (simp add: PDPlus_commute) |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
214 |
done |
41399
ad093e4638e2
changed syntax of powerdomain binary union operators
huffman
parents:
41394
diff
changeset
|
215 |
show "xs \<union>\<natural> xs = xs" |
34973
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
216 |
apply (induct xs rule: convex_pd.principal_induct, simp) |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
217 |
apply (simp add: PDPlus_absorb) |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
218 |
done |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
219 |
qed |
26927 | 220 |
|
34973
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
221 |
lemmas convex_plus_assoc = convex_add.assoc |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
222 |
lemmas convex_plus_commute = convex_add.commute |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
223 |
lemmas convex_plus_absorb = convex_add.idem |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
224 |
lemmas convex_plus_left_commute = convex_add.left_commute |
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
haftmann
parents:
33808
diff
changeset
|
225 |
lemmas convex_plus_left_absorb = convex_add.left_idem |
26927 | 226 |
|
62175 | 227 |
text \<open>Useful for \<open>simp add: convex_plus_ac\<close>\<close> |
29990
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
228 |
lemmas convex_plus_ac = |
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
229 |
convex_plus_assoc convex_plus_commute convex_plus_left_commute |
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
230 |
|
62175 | 231 |
text \<open>Useful for \<open>simp only: convex_plus_aci\<close>\<close> |
29990
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
232 |
lemmas convex_plus_aci = |
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
233 |
convex_plus_ac convex_plus_absorb convex_plus_left_absorb |
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
234 |
|
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
|
235 |
lemma convex_unit_below_plus_iff [simp]: |
41399
ad093e4638e2
changed syntax of powerdomain binary union operators
huffman
parents:
41394
diff
changeset
|
236 |
"{x}\<natural> \<sqsubseteq> ys \<union>\<natural> zs \<longleftrightarrow> {x}\<natural> \<sqsubseteq> ys \<and> {x}\<natural> \<sqsubseteq> zs" |
39970
9023b897e67a
simplify proofs of powerdomain inequalities
Brian Huffman <brianh@cs.pdx.edu>
parents:
37770
diff
changeset
|
237 |
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
|
238 |
apply (induct ys rule: convex_pd.principal_induct, simp) |
9023b897e67a
simplify proofs of powerdomain inequalities
Brian Huffman <brianh@cs.pdx.edu>
parents:
37770
diff
changeset
|
239 |
apply (induct zs rule: convex_pd.principal_induct, simp) |
9023b897e67a
simplify proofs of powerdomain inequalities
Brian Huffman <brianh@cs.pdx.edu>
parents:
37770
diff
changeset
|
240 |
apply simp |
25904 | 241 |
done |
242 |
||
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
|
243 |
lemma convex_plus_below_unit_iff [simp]: |
41399
ad093e4638e2
changed syntax of powerdomain binary union operators
huffman
parents:
41394
diff
changeset
|
244 |
"xs \<union>\<natural> ys \<sqsubseteq> {z}\<natural> \<longleftrightarrow> xs \<sqsubseteq> {z}\<natural> \<and> ys \<sqsubseteq> {z}\<natural>" |
39970
9023b897e67a
simplify proofs of powerdomain inequalities
Brian Huffman <brianh@cs.pdx.edu>
parents:
37770
diff
changeset
|
245 |
apply (induct xs rule: convex_pd.principal_induct, simp) |
9023b897e67a
simplify proofs of powerdomain inequalities
Brian Huffman <brianh@cs.pdx.edu>
parents:
37770
diff
changeset
|
246 |
apply (induct ys rule: convex_pd.principal_induct, simp) |
9023b897e67a
simplify proofs of powerdomain inequalities
Brian Huffman <brianh@cs.pdx.edu>
parents:
37770
diff
changeset
|
247 |
apply (induct z rule: compact_basis.principal_induct, simp) |
9023b897e67a
simplify proofs of powerdomain inequalities
Brian Huffman <brianh@cs.pdx.edu>
parents:
37770
diff
changeset
|
248 |
apply simp |
25904 | 249 |
done |
250 |
||
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
|
251 |
lemma convex_unit_below_iff [simp]: "{x}\<natural> \<sqsubseteq> {y}\<natural> \<longleftrightarrow> x \<sqsubseteq> y" |
39970
9023b897e67a
simplify proofs of powerdomain inequalities
Brian Huffman <brianh@cs.pdx.edu>
parents:
37770
diff
changeset
|
252 |
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
|
253 |
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
|
254 |
apply simp |
26927 | 255 |
done |
256 |
||
257 |
lemma convex_unit_eq_iff [simp]: "{x}\<natural> = {y}\<natural> \<longleftrightarrow> x = y" |
|
258 |
unfolding po_eq_conv by simp |
|
259 |
||
260 |
lemma convex_unit_strict [simp]: "{\<bottom>}\<natural> = \<bottom>" |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
261 |
using convex_unit_Rep_compact_basis [of compact_bot] |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
262 |
by (simp add: inst_convex_pd_pcpo) |
26927 | 263 |
|
40321
d065b195ec89
rename lemmas *_defined_iff and *_strict_iff to *_bottom_iff
huffman
parents:
40002
diff
changeset
|
264 |
lemma convex_unit_bottom_iff [simp]: "{x}\<natural> = \<bottom> \<longleftrightarrow> x = \<bottom>" |
26927 | 265 |
unfolding convex_unit_strict [symmetric] by (rule convex_unit_eq_iff) |
266 |
||
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
267 |
lemma compact_convex_unit: "compact x \<Longrightarrow> compact {x}\<natural>" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
268 |
by (auto dest!: compact_basis.compact_imp_principal) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
269 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
270 |
lemma compact_convex_unit_iff [simp]: "compact {x}\<natural> \<longleftrightarrow> compact x" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
271 |
apply (safe elim!: compact_convex_unit) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
272 |
apply (simp only: compact_def convex_unit_below_iff [symmetric]) |
40327 | 273 |
apply (erule adm_subst [OF cont_Rep_cfun2]) |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
274 |
done |
26927 | 275 |
|
276 |
lemma compact_convex_plus [simp]: |
|
41399
ad093e4638e2
changed syntax of powerdomain binary union operators
huffman
parents:
41394
diff
changeset
|
277 |
"\<lbrakk>compact xs; compact ys\<rbrakk> \<Longrightarrow> compact (xs \<union>\<natural> ys)" |
27289 | 278 |
by (auto dest!: convex_pd.compact_imp_principal) |
26927 | 279 |
|
25904 | 280 |
|
62175 | 281 |
subsection \<open>Induction rules\<close> |
25904 | 282 |
|
283 |
lemma convex_pd_induct1: |
|
284 |
assumes P: "adm P" |
|
26927 | 285 |
assumes unit: "\<And>x. P {x}\<natural>" |
41399
ad093e4638e2
changed syntax of powerdomain binary union operators
huffman
parents:
41394
diff
changeset
|
286 |
assumes insert: "\<And>x ys. \<lbrakk>P {x}\<natural>; P ys\<rbrakk> \<Longrightarrow> P ({x}\<natural> \<union>\<natural> ys)" |
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
287 |
shows "P (xs::'a::bifinite convex_pd)" |
81620 | 288 |
proof (induct xs rule: convex_pd.principal_induct) |
289 |
show "P (convex_principal a)" for a |
|
290 |
proof (induct a rule: pd_basis_induct1) |
|
291 |
case PDUnit |
|
292 |
show ?case by (simp only: convex_unit_Rep_compact_basis [symmetric]) (rule unit) |
|
293 |
next |
|
294 |
case PDPlus |
|
295 |
show ?case |
|
296 |
by (simp only: convex_unit_Rep_compact_basis [symmetric] convex_plus_principal [symmetric]) |
|
297 |
(rule insert [OF unit PDPlus]) |
|
298 |
qed |
|
299 |
qed (rule P) |
|
25904 | 300 |
|
81620 | 301 |
lemma convex_pd_induct [case_names adm convex_unit convex_plus, induct type: convex_pd]: |
25904 | 302 |
assumes P: "adm P" |
26927 | 303 |
assumes unit: "\<And>x. P {x}\<natural>" |
41399
ad093e4638e2
changed syntax of powerdomain binary union operators
huffman
parents:
41394
diff
changeset
|
304 |
assumes plus: "\<And>xs ys. \<lbrakk>P xs; P ys\<rbrakk> \<Longrightarrow> P (xs \<union>\<natural> ys)" |
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
305 |
shows "P (xs::'a::bifinite convex_pd)" |
81620 | 306 |
proof (induct xs rule: convex_pd.principal_induct) |
307 |
show "P (convex_principal a)" for a |
|
308 |
proof (induct a rule: pd_basis_induct) |
|
309 |
case PDUnit |
|
310 |
then show ?case by (simp only: convex_unit_Rep_compact_basis [symmetric] unit) |
|
311 |
next |
|
312 |
case PDPlus |
|
313 |
then show ?case by (simp only: convex_plus_principal [symmetric] plus) |
|
314 |
qed |
|
315 |
qed (rule P) |
|
25904 | 316 |
|
317 |
||
62175 | 318 |
subsection \<open>Monadic bind\<close> |
25904 | 319 |
|
320 |
definition |
|
321 |
convex_bind_basis :: |
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
322 |
"'a::bifinite pd_basis \<Rightarrow> ('a \<rightarrow> 'b convex_pd) \<rightarrow> 'b::bifinite convex_pd" where |
25904 | 323 |
"convex_bind_basis = fold_pd |
324 |
(\<lambda>a. \<Lambda> f. f\<cdot>(Rep_compact_basis a)) |
|
41399
ad093e4638e2
changed syntax of powerdomain binary union operators
huffman
parents:
41394
diff
changeset
|
325 |
(\<lambda>x y. \<Lambda> f. x\<cdot>f \<union>\<natural> y\<cdot>f)" |
25904 | 326 |
|
26927 | 327 |
lemma ACI_convex_bind: |
51489 | 328 |
"semilattice (\<lambda>x y. \<Lambda> f. x\<cdot>f \<union>\<natural> y\<cdot>f)" |
25904 | 329 |
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
|
330 |
apply (simp add: convex_plus_assoc) |
25904 | 331 |
apply (simp add: convex_plus_commute) |
29990
b11793ea15a3
avoid using ab_semigroup_idem_mult locale for powerdomains
huffman
parents:
29672
diff
changeset
|
332 |
apply (simp add: eta_cfun) |
25904 | 333 |
done |
334 |
||
335 |
lemma convex_bind_basis_simps [simp]: |
|
336 |
"convex_bind_basis (PDUnit a) = |
|
337 |
(\<Lambda> f. f\<cdot>(Rep_compact_basis a))" |
|
338 |
"convex_bind_basis (PDPlus t u) = |
|
41399
ad093e4638e2
changed syntax of powerdomain binary union operators
huffman
parents:
41394
diff
changeset
|
339 |
(\<Lambda> f. convex_bind_basis t\<cdot>f \<union>\<natural> convex_bind_basis u\<cdot>f)" |
25904 | 340 |
unfolding convex_bind_basis_def |
341 |
apply - |
|
26927 | 342 |
apply (rule fold_pd_PDUnit [OF ACI_convex_bind]) |
343 |
apply (rule fold_pd_PDPlus [OF ACI_convex_bind]) |
|
25904 | 344 |
done |
345 |
||
346 |
lemma convex_bind_basis_mono: |
|
347 |
"t \<le>\<natural> u \<Longrightarrow> convex_bind_basis t \<sqsubseteq> convex_bind_basis u" |
|
348 |
apply (erule convex_le_induct) |
|
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
|
349 |
apply (erule (1) below_trans) |
27289 | 350 |
apply (simp add: monofun_LAM monofun_cfun) |
351 |
apply (simp add: monofun_LAM monofun_cfun) |
|
25904 | 352 |
done |
353 |
||
354 |
definition |
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
355 |
convex_bind :: "'a::bifinite convex_pd \<rightarrow> ('a \<rightarrow> 'b convex_pd) \<rightarrow> 'b::bifinite convex_pd" where |
41394
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents:
41289
diff
changeset
|
356 |
"convex_bind = convex_pd.extension convex_bind_basis" |
25904 | 357 |
|
41036
4acbacd6c5bc
add set-union-like syntax for powerdomain bind operators
huffman
parents:
40888
diff
changeset
|
358 |
syntax |
4acbacd6c5bc
add set-union-like syntax for powerdomain bind operators
huffman
parents:
40888
diff
changeset
|
359 |
"_convex_bind" :: "[logic, logic, logic] \<Rightarrow> logic" |
81095 | 360 |
(\<open>(\<open>indent=3 notation=\<open>binder convex_bind\<close>\<close>\<Union>\<natural>_\<in>_./ _)\<close> [0, 0, 10] 10) |
80768 | 361 |
|
41036
4acbacd6c5bc
add set-union-like syntax for powerdomain bind operators
huffman
parents:
40888
diff
changeset
|
362 |
translations |
4acbacd6c5bc
add set-union-like syntax for powerdomain bind operators
huffman
parents:
40888
diff
changeset
|
363 |
"\<Union>\<natural>x\<in>xs. e" == "CONST convex_bind\<cdot>xs\<cdot>(\<Lambda> x. e)" |
4acbacd6c5bc
add set-union-like syntax for powerdomain bind operators
huffman
parents:
40888
diff
changeset
|
364 |
|
25904 | 365 |
lemma convex_bind_principal [simp]: |
366 |
"convex_bind\<cdot>(convex_principal t) = convex_bind_basis t" |
|
367 |
unfolding convex_bind_def |
|
41394
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents:
41289
diff
changeset
|
368 |
apply (rule convex_pd.extension_principal) |
25904 | 369 |
apply (erule convex_bind_basis_mono) |
370 |
done |
|
371 |
||
372 |
lemma convex_bind_unit [simp]: |
|
26927 | 373 |
"convex_bind\<cdot>{x}\<natural>\<cdot>f = f\<cdot>x" |
27289 | 374 |
by (induct x rule: compact_basis.principal_induct, simp, simp) |
25904 | 375 |
|
376 |
lemma convex_bind_plus [simp]: |
|
41399
ad093e4638e2
changed syntax of powerdomain binary union operators
huffman
parents:
41394
diff
changeset
|
377 |
"convex_bind\<cdot>(xs \<union>\<natural> ys)\<cdot>f = convex_bind\<cdot>xs\<cdot>f \<union>\<natural> convex_bind\<cdot>ys\<cdot>f" |
41402
b647212cee03
remove lemma ideal_completion.principal_induct2, use principal_induct twice instead
huffman
parents:
41399
diff
changeset
|
378 |
by (induct xs rule: convex_pd.principal_induct, simp, |
b647212cee03
remove lemma ideal_completion.principal_induct2, use principal_induct twice instead
huffman
parents:
41399
diff
changeset
|
379 |
induct ys rule: convex_pd.principal_induct, simp, simp) |
25904 | 380 |
|
381 |
lemma convex_bind_strict [simp]: "convex_bind\<cdot>\<bottom>\<cdot>f = f\<cdot>\<bottom>" |
|
382 |
unfolding convex_unit_strict [symmetric] by (rule convex_bind_unit) |
|
383 |
||
40589 | 384 |
lemma convex_bind_bind: |
385 |
"convex_bind\<cdot>(convex_bind\<cdot>xs\<cdot>f)\<cdot>g = |
|
386 |
convex_bind\<cdot>xs\<cdot>(\<Lambda> x. convex_bind\<cdot>(f\<cdot>x)\<cdot>g)" |
|
387 |
by (induct xs, simp_all) |
|
388 |
||
25904 | 389 |
|
62175 | 390 |
subsection \<open>Map\<close> |
25904 | 391 |
|
392 |
definition |
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
393 |
convex_map :: "('a::bifinite \<rightarrow> 'b) \<rightarrow> 'a convex_pd \<rightarrow> 'b::bifinite convex_pd" where |
26927 | 394 |
"convex_map = (\<Lambda> f xs. convex_bind\<cdot>xs\<cdot>(\<Lambda> x. {f\<cdot>x}\<natural>))" |
25904 | 395 |
|
396 |
lemma convex_map_unit [simp]: |
|
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
397 |
"convex_map\<cdot>f\<cdot>{x}\<natural> = {f\<cdot>x}\<natural>" |
25904 | 398 |
unfolding convex_map_def by simp |
399 |
||
400 |
lemma convex_map_plus [simp]: |
|
41399
ad093e4638e2
changed syntax of powerdomain binary union operators
huffman
parents:
41394
diff
changeset
|
401 |
"convex_map\<cdot>f\<cdot>(xs \<union>\<natural> ys) = convex_map\<cdot>f\<cdot>xs \<union>\<natural> convex_map\<cdot>f\<cdot>ys" |
25904 | 402 |
unfolding convex_map_def by simp |
403 |
||
40577 | 404 |
lemma convex_map_bottom [simp]: "convex_map\<cdot>f\<cdot>\<bottom> = {f\<cdot>\<bottom>}\<natural>" |
405 |
unfolding convex_map_def by simp |
|
406 |
||
25904 | 407 |
lemma convex_map_ident: "convex_map\<cdot>(\<Lambda> x. x)\<cdot>xs = xs" |
408 |
by (induct xs rule: convex_pd_induct, simp_all) |
|
409 |
||
33808 | 410 |
lemma convex_map_ID: "convex_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
|
411 |
by (simp add: cfun_eq_iff ID_def convex_map_ident) |
33808 | 412 |
|
25904 | 413 |
lemma convex_map_map: |
414 |
"convex_map\<cdot>f\<cdot>(convex_map\<cdot>g\<cdot>xs) = convex_map\<cdot>(\<Lambda> x. f\<cdot>(g\<cdot>x))\<cdot>xs" |
|
415 |
by (induct xs rule: convex_pd_induct, simp_all) |
|
416 |
||
41110 | 417 |
lemma convex_bind_map: |
418 |
"convex_bind\<cdot>(convex_map\<cdot>f\<cdot>xs)\<cdot>g = convex_bind\<cdot>xs\<cdot>(\<Lambda> x. g\<cdot>(f\<cdot>x))" |
|
419 |
by (simp add: convex_map_def convex_bind_bind) |
|
420 |
||
421 |
lemma convex_map_bind: |
|
422 |
"convex_map\<cdot>f\<cdot>(convex_bind\<cdot>xs\<cdot>g) = convex_bind\<cdot>xs\<cdot>(\<Lambda> x. convex_map\<cdot>f\<cdot>(g\<cdot>x))" |
|
423 |
by (simp add: convex_map_def convex_bind_bind) |
|
424 |
||
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
425 |
lemma ep_pair_convex_map: "ep_pair e p \<Longrightarrow> ep_pair (convex_map\<cdot>e) (convex_map\<cdot>p)" |
61169 | 426 |
apply standard |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
427 |
apply (induct_tac x rule: convex_pd_induct, simp_all add: ep_pair.e_inverse) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
428 |
apply (induct_tac y rule: convex_pd_induct) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
429 |
apply (simp_all add: ep_pair.e_p_below monofun_cfun) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
430 |
done |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
431 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
432 |
lemma deflation_convex_map: "deflation d \<Longrightarrow> deflation (convex_map\<cdot>d)" |
61169 | 433 |
apply standard |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
434 |
apply (induct_tac x rule: convex_pd_induct, simp_all add: deflation.idem) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
435 |
apply (induct_tac x rule: convex_pd_induct) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
436 |
apply (simp_all add: deflation.below monofun_cfun) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
437 |
done |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
438 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
439 |
(* FIXME: long proof! *) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
440 |
lemma finite_deflation_convex_map: |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
441 |
assumes "finite_deflation d" shows "finite_deflation (convex_map\<cdot>d)" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
442 |
proof (rule finite_deflation_intro) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
443 |
interpret d: finite_deflation d by fact |
67682
00c436488398
tuned proofs -- prefer explicit names for facts from 'interpret';
wenzelm
parents:
62175
diff
changeset
|
444 |
from d.deflation_axioms show "deflation (convex_map\<cdot>d)" |
00c436488398
tuned proofs -- prefer explicit names for facts from 'interpret';
wenzelm
parents:
62175
diff
changeset
|
445 |
by (rule deflation_convex_map) |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
446 |
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
|
447 |
hence "finite (Rep_compact_basis -` range (\<lambda>x. d\<cdot>x))" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
448 |
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
|
449 |
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
|
450 |
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
|
451 |
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
|
452 |
hence *: "finite (convex_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
|
453 |
hence "finite (range (\<lambda>xs. convex_map\<cdot>d\<cdot>xs))" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
454 |
apply (rule rev_finite_subset) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
455 |
apply clarsimp |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
456 |
apply (induct_tac xs rule: convex_pd.principal_induct) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
457 |
apply (simp add: adm_mem_finite *) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
458 |
apply (rename_tac t, induct_tac t rule: pd_basis_induct) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
459 |
apply (simp only: convex_unit_Rep_compact_basis [symmetric] convex_map_unit) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
460 |
apply simp |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
461 |
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
|
462 |
apply clarsimp |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
463 |
apply (rule imageI) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
464 |
apply (rule vimageI2) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
465 |
apply (simp add: Rep_PDUnit) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
466 |
apply (rule range_eqI) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
467 |
apply (erule sym) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
468 |
apply (rule exI) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
469 |
apply (rule Abs_compact_basis_inverse [symmetric]) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
470 |
apply (simp add: d.compact) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
471 |
apply (simp only: convex_plus_principal [symmetric] convex_map_plus) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
472 |
apply clarsimp |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
473 |
apply (rule imageI) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
474 |
apply (rule vimageI2) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
475 |
apply (simp add: Rep_PDPlus) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
476 |
done |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
477 |
thus "finite {xs. convex_map\<cdot>d\<cdot>xs = xs}" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
478 |
by (rule finite_range_imp_finite_fixes) |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
479 |
qed |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
480 |
|
81577 | 481 |
|
62175 | 482 |
subsection \<open>Convex powerdomain is bifinite\<close> |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
483 |
|
41286
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41111
diff
changeset
|
484 |
lemma approx_chain_convex_map: |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41111
diff
changeset
|
485 |
assumes "approx_chain a" |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41111
diff
changeset
|
486 |
shows "approx_chain (\<lambda>i. convex_map\<cdot>(a i))" |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41111
diff
changeset
|
487 |
using assms unfolding approx_chain_def |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41111
diff
changeset
|
488 |
by (simp add: lub_APP convex_map_ID finite_deflation_convex_map) |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41111
diff
changeset
|
489 |
|
41288
a19edebad961
powerdomain theories require class 'bifinite' instead of 'domain'
huffman
parents:
41287
diff
changeset
|
490 |
instance convex_pd :: (bifinite) bifinite |
41286
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41111
diff
changeset
|
491 |
proof |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41111
diff
changeset
|
492 |
show "\<exists>(a::nat \<Rightarrow> 'a convex_pd \<rightarrow> 'a convex_pd). approx_chain a" |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41111
diff
changeset
|
493 |
using bifinite [where 'a='a] |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41111
diff
changeset
|
494 |
by (fast intro!: approx_chain_convex_map) |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41111
diff
changeset
|
495 |
qed |
3d7685a4a5ff
reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents:
41111
diff
changeset
|
496 |
|
81577 | 497 |
|
62175 | 498 |
subsection \<open>Join\<close> |
39974
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 |
definition |
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
501 |
convex_join :: "'a::bifinite convex_pd convex_pd \<rightarrow> 'a convex_pd" where |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
502 |
"convex_join = (\<Lambda> xss. convex_bind\<cdot>xss\<cdot>(\<Lambda> xs. xs))" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
503 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
504 |
lemma convex_join_unit [simp]: |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
505 |
"convex_join\<cdot>{xs}\<natural> = xs" |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
506 |
unfolding convex_join_def by simp |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
507 |
|
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
508 |
lemma convex_join_plus [simp]: |
41399
ad093e4638e2
changed syntax of powerdomain binary union operators
huffman
parents:
41394
diff
changeset
|
509 |
"convex_join\<cdot>(xss \<union>\<natural> yss) = convex_join\<cdot>xss \<union>\<natural> convex_join\<cdot>yss" |
39974
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
510 |
unfolding convex_join_def by simp |
b525988432e9
major reorganization/simplification of HOLCF type classes:
huffman
parents:
39970
diff
changeset
|
511 |
|
40577 | 512 |
lemma convex_join_bottom [simp]: "convex_join\<cdot>\<bottom> = \<bottom>" |
513 |
unfolding convex_join_def by simp |
|
514 |
||
25904 | 515 |
lemma convex_join_map_unit: |
516 |
"convex_join\<cdot>(convex_map\<cdot>convex_unit\<cdot>xs) = xs" |
|
517 |
by (induct xs rule: convex_pd_induct, simp_all) |
|
518 |
||
519 |
lemma convex_join_map_join: |
|
520 |
"convex_join\<cdot>(convex_map\<cdot>convex_join\<cdot>xsss) = convex_join\<cdot>(convex_join\<cdot>xsss)" |
|
521 |
by (induct xsss rule: convex_pd_induct, simp_all) |
|
522 |
||
523 |
lemma convex_join_map_map: |
|
524 |
"convex_join\<cdot>(convex_map\<cdot>(convex_map\<cdot>f)\<cdot>xss) = |
|
525 |
convex_map\<cdot>f\<cdot>(convex_join\<cdot>xss)" |
|
526 |
by (induct xss rule: convex_pd_induct, simp_all) |
|
527 |
||
528 |
||
62175 | 529 |
subsection \<open>Conversions to other powerdomains\<close> |
25904 | 530 |
|
62175 | 531 |
text \<open>Convex to upper\<close> |
25904 | 532 |
|
533 |
lemma convex_le_imp_upper_le: "t \<le>\<natural> u \<Longrightarrow> t \<le>\<sharp> u" |
|
534 |
unfolding convex_le_def by simp |
|
535 |
||
536 |
definition |
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
537 |
convex_to_upper :: "'a::bifinite convex_pd \<rightarrow> 'a upper_pd" where |
41394
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents:
41289
diff
changeset
|
538 |
"convex_to_upper = convex_pd.extension upper_principal" |
25904 | 539 |
|
540 |
lemma convex_to_upper_principal [simp]: |
|
541 |
"convex_to_upper\<cdot>(convex_principal t) = upper_principal t" |
|
542 |
unfolding convex_to_upper_def |
|
41394
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents:
41289
diff
changeset
|
543 |
apply (rule convex_pd.extension_principal) |
27289 | 544 |
apply (rule upper_pd.principal_mono) |
25904 | 545 |
apply (erule convex_le_imp_upper_le) |
546 |
done |
|
547 |
||
548 |
lemma convex_to_upper_unit [simp]: |
|
26927 | 549 |
"convex_to_upper\<cdot>{x}\<natural> = {x}\<sharp>" |
27289 | 550 |
by (induct x rule: compact_basis.principal_induct, simp, simp) |
25904 | 551 |
|
552 |
lemma convex_to_upper_plus [simp]: |
|
41399
ad093e4638e2
changed syntax of powerdomain binary union operators
huffman
parents:
41394
diff
changeset
|
553 |
"convex_to_upper\<cdot>(xs \<union>\<natural> ys) = convex_to_upper\<cdot>xs \<union>\<sharp> convex_to_upper\<cdot>ys" |
41402
b647212cee03
remove lemma ideal_completion.principal_induct2, use principal_induct twice instead
huffman
parents:
41399
diff
changeset
|
554 |
by (induct xs rule: convex_pd.principal_induct, simp, |
b647212cee03
remove lemma ideal_completion.principal_induct2, use principal_induct twice instead
huffman
parents:
41399
diff
changeset
|
555 |
induct ys rule: convex_pd.principal_induct, simp, simp) |
25904 | 556 |
|
27289 | 557 |
lemma convex_to_upper_bind [simp]: |
558 |
"convex_to_upper\<cdot>(convex_bind\<cdot>xs\<cdot>f) = |
|
559 |
upper_bind\<cdot>(convex_to_upper\<cdot>xs)\<cdot>(convex_to_upper oo f)" |
|
560 |
by (induct xs rule: convex_pd_induct, simp, simp, simp) |
|
561 |
||
562 |
lemma convex_to_upper_map [simp]: |
|
563 |
"convex_to_upper\<cdot>(convex_map\<cdot>f\<cdot>xs) = upper_map\<cdot>f\<cdot>(convex_to_upper\<cdot>xs)" |
|
564 |
by (simp add: convex_map_def upper_map_def cfcomp_LAM) |
|
565 |
||
566 |
lemma convex_to_upper_join [simp]: |
|
567 |
"convex_to_upper\<cdot>(convex_join\<cdot>xss) = |
|
568 |
upper_bind\<cdot>(convex_to_upper\<cdot>xss)\<cdot>convex_to_upper" |
|
569 |
by (simp add: convex_join_def upper_join_def cfcomp_LAM eta_cfun) |
|
570 |
||
62175 | 571 |
text \<open>Convex to lower\<close> |
25904 | 572 |
|
573 |
lemma convex_le_imp_lower_le: "t \<le>\<natural> u \<Longrightarrow> t \<le>\<flat> u" |
|
574 |
unfolding convex_le_def by simp |
|
575 |
||
576 |
definition |
|
81583
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
wenzelm
parents:
81577
diff
changeset
|
577 |
convex_to_lower :: "'a::bifinite convex_pd \<rightarrow> 'a lower_pd" where |
41394
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents:
41289
diff
changeset
|
578 |
"convex_to_lower = convex_pd.extension lower_principal" |
25904 | 579 |
|
580 |
lemma convex_to_lower_principal [simp]: |
|
581 |
"convex_to_lower\<cdot>(convex_principal t) = lower_principal t" |
|
582 |
unfolding convex_to_lower_def |
|
41394
51c866d1b53b
rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents:
41289
diff
changeset
|
583 |
apply (rule convex_pd.extension_principal) |
27289 | 584 |
apply (rule lower_pd.principal_mono) |
25904 | 585 |
apply (erule convex_le_imp_lower_le) |
586 |
done |
|
587 |
||
588 |
lemma convex_to_lower_unit [simp]: |
|
26927 | 589 |
"convex_to_lower\<cdot>{x}\<natural> = {x}\<flat>" |
27289 | 590 |
by (induct x rule: compact_basis.principal_induct, simp, simp) |
25904 | 591 |
|
592 |
lemma convex_to_lower_plus [simp]: |
|
41399
ad093e4638e2
changed syntax of powerdomain binary union operators
huffman
parents:
41394
diff
changeset
|
593 |
"convex_to_lower\<cdot>(xs \<union>\<natural> ys) = convex_to_lower\<cdot>xs \<union>\<flat> convex_to_lower\<cdot>ys" |
41402
b647212cee03
remove lemma ideal_completion.principal_induct2, use principal_induct twice instead
huffman
parents:
41399
diff
changeset
|
594 |
by (induct xs rule: convex_pd.principal_induct, simp, |
b647212cee03
remove lemma ideal_completion.principal_induct2, use principal_induct twice instead
huffman
parents:
41399
diff
changeset
|
595 |
induct ys rule: convex_pd.principal_induct, simp, simp) |
25904 | 596 |
|
27289 | 597 |
lemma convex_to_lower_bind [simp]: |
598 |
"convex_to_lower\<cdot>(convex_bind\<cdot>xs\<cdot>f) = |
|
599 |
lower_bind\<cdot>(convex_to_lower\<cdot>xs)\<cdot>(convex_to_lower oo f)" |
|
600 |
by (induct xs rule: convex_pd_induct, simp, simp, simp) |
|
601 |
||
602 |
lemma convex_to_lower_map [simp]: |
|
603 |
"convex_to_lower\<cdot>(convex_map\<cdot>f\<cdot>xs) = lower_map\<cdot>f\<cdot>(convex_to_lower\<cdot>xs)" |
|
604 |
by (simp add: convex_map_def lower_map_def cfcomp_LAM) |
|
605 |
||
606 |
lemma convex_to_lower_join [simp]: |
|
607 |
"convex_to_lower\<cdot>(convex_join\<cdot>xss) = |
|
608 |
lower_bind\<cdot>(convex_to_lower\<cdot>xss)\<cdot>convex_to_lower" |
|
609 |
by (simp add: convex_join_def lower_join_def cfcomp_LAM eta_cfun) |
|
610 |
||
62175 | 611 |
text \<open>Ordering property\<close> |
25904 | 612 |
|
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
|
613 |
lemma convex_pd_below_iff: |
25904 | 614 |
"(xs \<sqsubseteq> ys) = |
615 |
(convex_to_upper\<cdot>xs \<sqsubseteq> convex_to_upper\<cdot>ys \<and> |
|
616 |
convex_to_lower\<cdot>xs \<sqsubseteq> convex_to_lower\<cdot>ys)" |
|
39970
9023b897e67a
simplify proofs of powerdomain inequalities
Brian Huffman <brianh@cs.pdx.edu>
parents:
37770
diff
changeset
|
617 |
apply (induct xs rule: convex_pd.principal_induct, simp) |
9023b897e67a
simplify proofs of powerdomain inequalities
Brian Huffman <brianh@cs.pdx.edu>
parents:
37770
diff
changeset
|
618 |
apply (induct ys rule: convex_pd.principal_induct, simp) |
9023b897e67a
simplify proofs of powerdomain inequalities
Brian Huffman <brianh@cs.pdx.edu>
parents:
37770
diff
changeset
|
619 |
apply (simp add: convex_le_def) |
25904 | 620 |
done |
621 |
||
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
|
622 |
lemmas convex_plus_below_plus_iff = |
45606 | 623 |
convex_pd_below_iff [where xs="xs \<union>\<natural> ys" and ys="zs \<union>\<natural> ws"] |
624 |
for xs ys zs ws |
|
26927 | 625 |
|
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
|
626 |
lemmas convex_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
|
627 |
convex_unit_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
|
628 |
convex_plus_below_unit_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
|
629 |
convex_plus_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
|
630 |
convex_unit_below_iff |
26927 | 631 |
convex_to_upper_unit |
632 |
convex_to_upper_plus |
|
633 |
convex_to_lower_unit |
|
634 |
convex_to_lower_plus |
|
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
|
635 |
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
|
636 |
lower_pd_below_simps |
26927 | 637 |
|
25904 | 638 |
end |