author | Fabian Huch <huch@in.tum.de> |
Thu, 30 Nov 2023 13:44:51 +0100 | |
changeset 79084 | dd689c4ab688 |
parent 78099 | 4d9349989d94 |
child 80768 | c7723cc15de8 |
permissions | -rw-r--r-- |
42151 | 1 |
(* Title: HOL/HOLCF/Pcpo.thy |
2640 | 2 |
Author: Franz Regensburger |
3 |
*) |
|
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
15563
diff
changeset
|
4 |
|
62175 | 5 |
section \<open>Classes cpo and pcpo\<close> |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
15563
diff
changeset
|
6 |
|
15577 | 7 |
theory Pcpo |
67312 | 8 |
imports Porder |
15577 | 9 |
begin |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
10 |
|
62175 | 11 |
subsection \<open>Complete partial orders\<close> |
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
12 |
|
62175 | 13 |
text \<open>The class cpo of chain complete partial orders\<close> |
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
14 |
|
29614
1f7b1b0df292
simplified handling of base sort, dropped axclass
haftmann
parents:
29138
diff
changeset
|
15 |
class cpo = po + |
31071 | 16 |
assumes cpo: "chain S \<Longrightarrow> \<exists>x. range S <<| x" |
17 |
begin |
|
2394 | 18 |
|
62175 | 19 |
text \<open>in cpo's everthing equal to THE lub has lub properties for every chain\<close> |
15563 | 20 |
|
31071 | 21 |
lemma cpo_lubI: "chain S \<Longrightarrow> range S <<| (\<Squnion>i. S i)" |
40771 | 22 |
by (fast dest: cpo elim: is_lub_lub) |
26026 | 23 |
|
31071 | 24 |
lemma thelubE: "\<lbrakk>chain S; (\<Squnion>i. S i) = l\<rbrakk> \<Longrightarrow> range S <<| l" |
40771 | 25 |
by (blast dest: cpo intro: is_lub_lub) |
15563 | 26 |
|
62175 | 27 |
text \<open>Properties of the lub\<close> |
15563 | 28 |
|
31071 | 29 |
lemma is_ub_thelub: "chain S \<Longrightarrow> S x \<sqsubseteq> (\<Squnion>i. S i)" |
40771 | 30 |
by (blast dest: cpo intro: is_lub_lub [THEN is_lub_rangeD1]) |
15563 | 31 |
|
67312 | 32 |
lemma is_lub_thelub: "\<lbrakk>chain S; range S <| x\<rbrakk> \<Longrightarrow> (\<Squnion>i. S i) \<sqsubseteq> x" |
40771 | 33 |
by (blast dest: cpo intro: is_lub_lub [THEN is_lubD2]) |
15563 | 34 |
|
39969 | 35 |
lemma lub_below_iff: "chain S \<Longrightarrow> (\<Squnion>i. S i) \<sqsubseteq> x \<longleftrightarrow> (\<forall>i. S i \<sqsubseteq> x)" |
36 |
by (simp add: is_lub_below_iff [OF cpo_lubI] is_ub_def) |
|
37 |
||
40500
ee9c8d36318e
add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
huffman
parents:
40498
diff
changeset
|
38 |
lemma lub_below: "\<lbrakk>chain S; \<And>i. S i \<sqsubseteq> x\<rbrakk> \<Longrightarrow> (\<Squnion>i. S i) \<sqsubseteq> x" |
ee9c8d36318e
add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
huffman
parents:
40498
diff
changeset
|
39 |
by (simp add: lub_below_iff) |
ee9c8d36318e
add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
huffman
parents:
40498
diff
changeset
|
40 |
|
ee9c8d36318e
add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
huffman
parents:
40498
diff
changeset
|
41 |
lemma below_lub: "\<lbrakk>chain S; x \<sqsubseteq> S i\<rbrakk> \<Longrightarrow> x \<sqsubseteq> (\<Squnion>i. S i)" |
ee9c8d36318e
add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
huffman
parents:
40498
diff
changeset
|
42 |
by (erule below_trans, erule is_ub_thelub) |
ee9c8d36318e
add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
huffman
parents:
40498
diff
changeset
|
43 |
|
67312 | 44 |
lemma lub_range_mono: "\<lbrakk>range X \<subseteq> range Y; chain Y; chain X\<rbrakk> \<Longrightarrow> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)" |
45 |
apply (erule lub_below) |
|
46 |
apply (subgoal_tac "\<exists>j. X i = Y j") |
|
47 |
apply clarsimp |
|
48 |
apply (erule is_ub_thelub) |
|
49 |
apply auto |
|
50 |
done |
|
15563 | 51 |
|
67312 | 52 |
lemma lub_range_shift: "chain Y \<Longrightarrow> (\<Squnion>i. Y (i + j)) = (\<Squnion>i. Y i)" |
53 |
apply (rule below_antisym) |
|
54 |
apply (rule lub_range_mono) |
|
55 |
apply fast |
|
56 |
apply assumption |
|
57 |
apply (erule chain_shift) |
|
58 |
apply (rule lub_below) |
|
59 |
apply assumption |
|
60 |
apply (rule_tac i="i" in below_lub) |
|
61 |
apply (erule chain_shift) |
|
62 |
apply (erule chain_mono) |
|
63 |
apply (rule le_add1) |
|
64 |
done |
|
15563 | 65 |
|
67312 | 66 |
lemma maxinch_is_thelub: "chain Y \<Longrightarrow> max_in_chain i Y = ((\<Squnion>i. Y i) = Y i)" |
67 |
apply (rule iffI) |
|
68 |
apply (fast intro!: lub_eqI lub_finch1) |
|
69 |
apply (unfold max_in_chain_def) |
|
70 |
apply (safe intro!: below_antisym) |
|
71 |
apply (fast elim!: chain_mono) |
|
72 |
apply (drule sym) |
|
73 |
apply (force elim!: is_ub_thelub) |
|
74 |
done |
|
15563 | 75 |
|
62175 | 76 |
text \<open>the \<open>\<sqsubseteq>\<close> relation between two chains is preserved by their lubs\<close> |
15563 | 77 |
|
67312 | 78 |
lemma lub_mono: "\<lbrakk>chain X; chain Y; \<And>i. X i \<sqsubseteq> Y i\<rbrakk> \<Longrightarrow> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)" |
79 |
by (fast elim: lub_below below_lub) |
|
15563 | 80 |
|
62175 | 81 |
text \<open>the = relation between two chains is preserved by their lubs\<close> |
15563 | 82 |
|
67312 | 83 |
lemma lub_eq: "(\<And>i. X i = Y i) \<Longrightarrow> (\<Squnion>i. X i) = (\<Squnion>i. Y i)" |
35492 | 84 |
by simp |
85 |
||
16203 | 86 |
lemma ch2ch_lub: |
87 |
assumes 1: "\<And>j. chain (\<lambda>i. Y i j)" |
|
88 |
assumes 2: "\<And>i. chain (\<lambda>j. Y i j)" |
|
89 |
shows "chain (\<lambda>i. \<Squnion>j. Y i j)" |
|
67312 | 90 |
apply (rule chainI) |
91 |
apply (rule lub_mono [OF 2 2]) |
|
92 |
apply (rule chainE [OF 1]) |
|
93 |
done |
|
16203 | 94 |
|
16201 | 95 |
lemma diag_lub: |
96 |
assumes 1: "\<And>j. chain (\<lambda>i. Y i j)" |
|
97 |
assumes 2: "\<And>i. chain (\<lambda>j. Y i j)" |
|
98 |
shows "(\<Squnion>i. \<Squnion>j. Y i j) = (\<Squnion>i. Y i i)" |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31071
diff
changeset
|
99 |
proof (rule below_antisym) |
16201 | 100 |
have 3: "chain (\<lambda>i. Y i i)" |
101 |
apply (rule chainI) |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31071
diff
changeset
|
102 |
apply (rule below_trans) |
67312 | 103 |
apply (rule chainE [OF 1]) |
16201 | 104 |
apply (rule chainE [OF 2]) |
105 |
done |
|
106 |
have 4: "chain (\<lambda>i. \<Squnion>j. Y i j)" |
|
16203 | 107 |
by (rule ch2ch_lub [OF 1 2]) |
16201 | 108 |
show "(\<Squnion>i. \<Squnion>j. Y i j) \<sqsubseteq> (\<Squnion>i. Y i i)" |
40500
ee9c8d36318e
add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
huffman
parents:
40498
diff
changeset
|
109 |
apply (rule lub_below [OF 4]) |
ee9c8d36318e
add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
huffman
parents:
40498
diff
changeset
|
110 |
apply (rule lub_below [OF 2]) |
ee9c8d36318e
add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
huffman
parents:
40498
diff
changeset
|
111 |
apply (rule below_lub [OF 3]) |
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31071
diff
changeset
|
112 |
apply (rule below_trans) |
67312 | 113 |
apply (rule chain_mono [OF 1 max.cobounded1]) |
54863
82acc20ded73
prefer more canonical names for lemmas on min/max
haftmann
parents:
42151
diff
changeset
|
114 |
apply (rule chain_mono [OF 2 max.cobounded2]) |
16201 | 115 |
done |
116 |
show "(\<Squnion>i. Y i i) \<sqsubseteq> (\<Squnion>i. \<Squnion>j. Y i j)" |
|
25923 | 117 |
apply (rule lub_mono [OF 3 4]) |
16201 | 118 |
apply (rule is_ub_thelub [OF 2]) |
119 |
done |
|
120 |
qed |
|
121 |
||
122 |
lemma ex_lub: |
|
123 |
assumes 1: "\<And>j. chain (\<lambda>i. Y i j)" |
|
124 |
assumes 2: "\<And>i. chain (\<lambda>j. Y i j)" |
|
125 |
shows "(\<Squnion>i. \<Squnion>j. Y i j) = (\<Squnion>j. \<Squnion>i. Y i j)" |
|
31071 | 126 |
by (simp add: diag_lub 1 2) |
16201 | 127 |
|
31071 | 128 |
end |
16201 | 129 |
|
67312 | 130 |
|
62175 | 131 |
subsection \<open>Pointed cpos\<close> |
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
132 |
|
62175 | 133 |
text \<open>The class pcpo of pointed cpos\<close> |
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
134 |
|
29614
1f7b1b0df292
simplified handling of base sort, dropped axclass
haftmann
parents:
29138
diff
changeset
|
135 |
class pcpo = cpo + |
1f7b1b0df292
simplified handling of base sort, dropped axclass
haftmann
parents:
29138
diff
changeset
|
136 |
assumes least: "\<exists>x. \<forall>y. x \<sqsubseteq> y" |
31071 | 137 |
begin |
25723 | 138 |
|
61998 | 139 |
definition bottom :: "'a" ("\<bottom>") |
41429
cf5f025bc3c7
renamed constant 'UU' to 'bottom', keeping 'UU' as alternative input syntax;
huffman
parents:
40774
diff
changeset
|
140 |
where "bottom = (THE x. \<forall>y. x \<sqsubseteq> y)" |
25723 | 141 |
|
41429
cf5f025bc3c7
renamed constant 'UU' to 'bottom', keeping 'UU' as alternative input syntax;
huffman
parents:
40774
diff
changeset
|
142 |
lemma minimal [iff]: "\<bottom> \<sqsubseteq> x" |
67312 | 143 |
unfolding bottom_def |
144 |
apply (rule the1I2) |
|
145 |
apply (rule ex_ex1I) |
|
146 |
apply (rule least) |
|
147 |
apply (blast intro: below_antisym) |
|
148 |
apply simp |
|
149 |
done |
|
25723 | 150 |
|
41429
cf5f025bc3c7
renamed constant 'UU' to 'bottom', keeping 'UU' as alternative input syntax;
huffman
parents:
40774
diff
changeset
|
151 |
end |
cf5f025bc3c7
renamed constant 'UU' to 'bottom', keeping 'UU' as alternative input syntax;
huffman
parents:
40774
diff
changeset
|
152 |
|
62175 | 153 |
text \<open>Old "UU" syntax:\<close> |
25723 | 154 |
|
41429
cf5f025bc3c7
renamed constant 'UU' to 'bottom', keeping 'UU' as alternative input syntax;
huffman
parents:
40774
diff
changeset
|
155 |
syntax UU :: logic |
67312 | 156 |
translations "UU" \<rightharpoonup> "CONST bottom" |
31071 | 157 |
|
69597 | 158 |
text \<open>Simproc to rewrite \<^term>\<open>\<bottom> = x\<close> to \<^term>\<open>x = \<bottom>\<close>.\<close> |
74305 | 159 |
setup \<open>Reorient_Proc.add (fn \<^Const_>\<open>bottom _\<close> => true | _ => false)\<close> |
78099
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
wenzelm
parents:
74305
diff
changeset
|
160 |
simproc_setup reorient_bottom ("\<bottom> = x") = \<open>K Reorient_Proc.proc\<close> |
25723 | 161 |
|
69597 | 162 |
text \<open>useful lemmas about \<^term>\<open>\<bottom>\<close>\<close> |
25723 | 163 |
|
67312 | 164 |
lemma below_bottom_iff [simp]: "x \<sqsubseteq> \<bottom> \<longleftrightarrow> x = \<bottom>" |
165 |
by (simp add: po_eq_conv) |
|
25723 | 166 |
|
67312 | 167 |
lemma eq_bottom_iff: "x = \<bottom> \<longleftrightarrow> x \<sqsubseteq> \<bottom>" |
168 |
by simp |
|
25723 | 169 |
|
41430
1aa23e9f2c87
change some lemma names containing 'UU' to 'bottom'
huffman
parents:
41429
diff
changeset
|
170 |
lemma bottomI: "x \<sqsubseteq> \<bottom> \<Longrightarrow> x = \<bottom>" |
67312 | 171 |
by (subst eq_bottom_iff) |
25723 | 172 |
|
40045 | 173 |
lemma lub_eq_bottom_iff: "chain Y \<Longrightarrow> (\<Squnion>i. Y i) = \<bottom> \<longleftrightarrow> (\<forall>i. Y i = \<bottom>)" |
67312 | 174 |
by (simp only: eq_bottom_iff lub_below_iff) |
175 |
||
40045 | 176 |
|
62175 | 177 |
subsection \<open>Chain-finite and flat cpos\<close> |
15563 | 178 |
|
62175 | 179 |
text \<open>further useful classes for HOLCF domains\<close> |
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
180 |
|
31071 | 181 |
class chfin = po + |
182 |
assumes chfin: "chain Y \<Longrightarrow> \<exists>n. max_in_chain n Y" |
|
183 |
begin |
|
25814 | 184 |
|
31071 | 185 |
subclass cpo |
67312 | 186 |
apply standard |
187 |
apply (frule chfin) |
|
188 |
apply (blast intro: lub_finch1) |
|
189 |
done |
|
15563 | 190 |
|
31071 | 191 |
lemma chfin2finch: "chain Y \<Longrightarrow> finite_chain Y" |
192 |
by (simp add: chfin finite_chain_def) |
|
193 |
||
194 |
end |
|
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
195 |
|
31071 | 196 |
class flat = pcpo + |
197 |
assumes ax_flat: "x \<sqsubseteq> y \<Longrightarrow> x = \<bottom> \<or> x = y" |
|
198 |
begin |
|
15640
2d1d6ea579a1
chfin now a subclass of po, proved instance chfin < cpo
huffman
parents:
15588
diff
changeset
|
199 |
|
31071 | 200 |
subclass chfin |
68369 | 201 |
proof |
202 |
fix Y |
|
203 |
assume *: "chain Y" |
|
204 |
show "\<exists>n. max_in_chain n Y" |
|
205 |
apply (unfold max_in_chain_def) |
|
206 |
apply (cases "\<forall>i. Y i = \<bottom>") |
|
207 |
apply simp |
|
208 |
apply simp |
|
209 |
apply (erule exE) |
|
210 |
apply (rule_tac x="i" in exI) |
|
211 |
apply clarify |
|
212 |
using * apply (blast dest: chain_mono ax_flat) |
|
213 |
done |
|
214 |
qed |
|
15563 | 215 |
|
67312 | 216 |
lemma flat_below_iff: "x \<sqsubseteq> y \<longleftrightarrow> x = \<bottom> \<or> x = y" |
31071 | 217 |
by (safe dest!: ax_flat) |
25826 | 218 |
|
31071 | 219 |
lemma flat_eq: "a \<noteq> \<bottom> \<Longrightarrow> a \<sqsubseteq> b = (a = b)" |
220 |
by (safe dest!: ax_flat) |
|
15563 | 221 |
|
31071 | 222 |
end |
15563 | 223 |
|
62175 | 224 |
subsection \<open>Discrete cpos\<close> |
26023 | 225 |
|
31076
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents:
31071
diff
changeset
|
226 |
class discrete_cpo = below + |
29614
1f7b1b0df292
simplified handling of base sort, dropped axclass
haftmann
parents:
29138
diff
changeset
|
227 |
assumes discrete_cpo [simp]: "x \<sqsubseteq> y \<longleftrightarrow> x = y" |
31071 | 228 |
begin |
26023 | 229 |
|
31071 | 230 |
subclass po |
67312 | 231 |
by standard simp_all |
26023 | 232 |
|
62175 | 233 |
text \<open>In a discrete cpo, every chain is constant\<close> |
26023 | 234 |
|
235 |
lemma discrete_chain_const: |
|
31071 | 236 |
assumes S: "chain S" |
26023 | 237 |
shows "\<exists>x. S = (\<lambda>i. x)" |
238 |
proof (intro exI ext) |
|
239 |
fix i :: nat |
|
67312 | 240 |
from S le0 have "S 0 \<sqsubseteq> S i" by (rule chain_mono) |
241 |
then have "S 0 = S i" by simp |
|
242 |
then show "S i = S 0" by (rule sym) |
|
26023 | 243 |
qed |
244 |
||
40091
1ca61fbd8a79
make discrete_cpo a subclass of chfin; remove chfin instances for fun, cfun
huffman
parents:
40089
diff
changeset
|
245 |
subclass chfin |
26023 | 246 |
proof |
247 |
fix S :: "nat \<Rightarrow> 'a" |
|
248 |
assume S: "chain S" |
|
67312 | 249 |
then have "\<exists>x. S = (\<lambda>i. x)" |
250 |
by (rule discrete_chain_const) |
|
251 |
then have "max_in_chain 0 S" |
|
252 |
by (auto simp: max_in_chain_def) |
|
253 |
then show "\<exists>i. max_in_chain i S" .. |
|
26023 | 254 |
qed |
255 |
||
31071 | 256 |
end |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
15563
diff
changeset
|
257 |
|
16626 | 258 |
end |