author | haftmann |
Tue, 28 Oct 2008 16:58:59 +0100 | |
changeset 28703 | aef727ef30e5 |
parent 28262 | aa7ca36d67fd |
child 29138 | 661a8db7e647 |
permissions | -rw-r--r-- |
2640 | 1 |
(* Title: HOLCF/Pcpo.thy |
2 |
ID: $Id$ |
|
3 |
Author: Franz Regensburger |
|
4 |
*) |
|
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
15563
diff
changeset
|
5 |
|
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
15563
diff
changeset
|
6 |
header {* Classes cpo and pcpo *} |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
15563
diff
changeset
|
7 |
|
15577 | 8 |
theory Pcpo |
9 |
imports Porder |
|
10 |
begin |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
11 |
|
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
12 |
subsection {* Complete partial orders *} |
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
13 |
|
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
14 |
text {* The class cpo of chain complete partial orders *} |
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
15 |
|
2640 | 16 |
axclass cpo < po |
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
17 |
-- {* class axiom: *} |
27413 | 18 |
cpo: "chain S \<Longrightarrow> \<exists>x. range S <<| x" |
2394 | 19 |
|
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
20 |
text {* in cpo's everthing equal to THE lub has lub properties for every chain *} |
15563 | 21 |
|
27413 | 22 |
lemma cpo_lubI: "chain (S::nat \<Rightarrow> 'a::cpo) \<Longrightarrow> range S <<| (\<Squnion>i. S i)" |
26026 | 23 |
by (fast dest: cpo elim: lubI) |
24 |
||
16626 | 25 |
lemma thelubE: "\<lbrakk>chain S; (\<Squnion>i. S i) = (l::'a::cpo)\<rbrakk> \<Longrightarrow> range S <<| l" |
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
26 |
by (blast dest: cpo intro: lubI) |
15563 | 27 |
|
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
28 |
text {* Properties of the lub *} |
15563 | 29 |
|
16626 | 30 |
lemma is_ub_thelub: "chain (S::nat \<Rightarrow> 'a::cpo) \<Longrightarrow> S x \<sqsubseteq> (\<Squnion>i. S i)" |
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
31 |
by (blast dest: cpo intro: lubI [THEN is_ub_lub]) |
15563 | 32 |
|
16626 | 33 |
lemma is_lub_thelub: |
34 |
"\<lbrakk>chain (S::nat \<Rightarrow> 'a::cpo); range S <| x\<rbrakk> \<Longrightarrow> (\<Squnion>i. S i) \<sqsubseteq> x" |
|
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
35 |
by (blast dest: cpo intro: lubI [THEN is_lub_lub]) |
15563 | 36 |
|
16626 | 37 |
lemma lub_range_mono: |
38 |
"\<lbrakk>range X \<subseteq> range Y; chain Y; chain (X::nat \<Rightarrow> 'a::cpo)\<rbrakk> |
|
39 |
\<Longrightarrow> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)" |
|
15563 | 40 |
apply (erule is_lub_thelub) |
41 |
apply (rule ub_rangeI) |
|
16626 | 42 |
apply (subgoal_tac "\<exists>j. X i = Y j") |
15563 | 43 |
apply clarsimp |
44 |
apply (erule is_ub_thelub) |
|
45 |
apply auto |
|
46 |
done |
|
47 |
||
16626 | 48 |
lemma lub_range_shift: |
49 |
"chain (Y::nat \<Rightarrow> 'a::cpo) \<Longrightarrow> (\<Squnion>i. Y (i + j)) = (\<Squnion>i. Y i)" |
|
15563 | 50 |
apply (rule antisym_less) |
51 |
apply (rule lub_range_mono) |
|
52 |
apply fast |
|
53 |
apply assumption |
|
54 |
apply (erule chain_shift) |
|
55 |
apply (rule is_lub_thelub) |
|
56 |
apply assumption |
|
57 |
apply (rule ub_rangeI) |
|
17813 | 58 |
apply (rule_tac y="Y (i + j)" in trans_less) |
25922
cb04d05e95fb
rename lemma chain_mono3 -> chain_mono, chain_mono -> chain_mono_less
huffman
parents:
25921
diff
changeset
|
59 |
apply (erule chain_mono) |
15563 | 60 |
apply (rule le_add1) |
17813 | 61 |
apply (rule is_ub_thelub) |
62 |
apply (erule chain_shift) |
|
15563 | 63 |
done |
64 |
||
16626 | 65 |
lemma maxinch_is_thelub: |
66 |
"chain Y \<Longrightarrow> max_in_chain i Y = ((\<Squnion>i. Y i) = ((Y i)::'a::cpo))" |
|
15563 | 67 |
apply (rule iffI) |
68 |
apply (fast intro!: thelubI lub_finch1) |
|
69 |
apply (unfold max_in_chain_def) |
|
70 |
apply (safe intro!: antisym_less) |
|
25922
cb04d05e95fb
rename lemma chain_mono3 -> chain_mono, chain_mono -> chain_mono_less
huffman
parents:
25921
diff
changeset
|
71 |
apply (fast elim!: chain_mono) |
15563 | 72 |
apply (drule sym) |
73 |
apply (force elim!: is_ub_thelub) |
|
74 |
done |
|
75 |
||
16626 | 76 |
text {* the @{text "\<sqsubseteq>"} relation between two chains is preserved by their lubs *} |
15563 | 77 |
|
16626 | 78 |
lemma lub_mono: |
25923 | 79 |
"\<lbrakk>chain (X::nat \<Rightarrow> 'a::cpo); chain Y; \<And>i. X i \<sqsubseteq> Y i\<rbrakk> |
16626 | 80 |
\<Longrightarrow> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)" |
15563 | 81 |
apply (erule is_lub_thelub) |
82 |
apply (rule ub_rangeI) |
|
83 |
apply (rule trans_less) |
|
25923 | 84 |
apply (erule meta_spec) |
15563 | 85 |
apply (erule is_ub_thelub) |
86 |
done |
|
87 |
||
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
88 |
text {* the = relation between two chains is preserved by their lubs *} |
15563 | 89 |
|
16626 | 90 |
lemma lub_equal: |
91 |
"\<lbrakk>chain (X::nat \<Rightarrow> 'a::cpo); chain Y; \<forall>k. X k = Y k\<rbrakk> |
|
92 |
\<Longrightarrow> (\<Squnion>i. X i) = (\<Squnion>i. Y i)" |
|
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
93 |
by (simp only: expand_fun_eq [symmetric]) |
15563 | 94 |
|
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
95 |
text {* more results about mono and = of lubs of chains *} |
3326 | 96 |
|
16626 | 97 |
lemma lub_mono2: |
17813 | 98 |
"\<lbrakk>\<exists>j. \<forall>i>j. X i = Y i; chain (X::nat \<Rightarrow> 'a::cpo); chain Y\<rbrakk> |
16626 | 99 |
\<Longrightarrow> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)" |
15563 | 100 |
apply (erule exE) |
17813 | 101 |
apply (subgoal_tac "(\<Squnion>i. X (i + Suc j)) \<sqsubseteq> (\<Squnion>i. Y (i + Suc j))") |
102 |
apply (thin_tac "\<forall>i>j. X i = Y i") |
|
103 |
apply (simp only: lub_range_shift) |
|
16626 | 104 |
apply simp |
15563 | 105 |
done |
106 |
||
16626 | 107 |
lemma lub_equal2: |
108 |
"\<lbrakk>\<exists>j. \<forall>i>j. X i = Y i; chain (X::nat \<Rightarrow> 'a::cpo); chain Y\<rbrakk> |
|
109 |
\<Longrightarrow> (\<Squnion>i. X i) = (\<Squnion>i. Y i)" |
|
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
110 |
by (blast intro: antisym_less lub_mono2 sym) |
15563 | 111 |
|
16626 | 112 |
lemma lub_mono3: |
113 |
"\<lbrakk>chain (Y::nat \<Rightarrow> 'a::cpo); chain X; \<forall>i. \<exists>j. Y i \<sqsubseteq> X j\<rbrakk> |
|
114 |
\<Longrightarrow> (\<Squnion>i. Y i) \<sqsubseteq> (\<Squnion>i. X i)" |
|
17813 | 115 |
apply (erule is_lub_thelub) |
15563 | 116 |
apply (rule ub_rangeI) |
117 |
apply (erule allE) |
|
118 |
apply (erule exE) |
|
16626 | 119 |
apply (erule trans_less) |
120 |
apply (erule is_ub_thelub) |
|
15563 | 121 |
done |
122 |
||
16203 | 123 |
lemma ch2ch_lub: |
124 |
fixes Y :: "nat \<Rightarrow> nat \<Rightarrow> 'a::cpo" |
|
125 |
assumes 1: "\<And>j. chain (\<lambda>i. Y i j)" |
|
126 |
assumes 2: "\<And>i. chain (\<lambda>j. Y i j)" |
|
127 |
shows "chain (\<lambda>i. \<Squnion>j. Y i j)" |
|
128 |
apply (rule chainI) |
|
25923 | 129 |
apply (rule lub_mono [OF 2 2]) |
16203 | 130 |
apply (rule chainE [OF 1]) |
131 |
done |
|
132 |
||
16201 | 133 |
lemma diag_lub: |
134 |
fixes Y :: "nat \<Rightarrow> nat \<Rightarrow> 'a::cpo" |
|
135 |
assumes 1: "\<And>j. chain (\<lambda>i. Y i j)" |
|
136 |
assumes 2: "\<And>i. chain (\<lambda>j. Y i j)" |
|
137 |
shows "(\<Squnion>i. \<Squnion>j. Y i j) = (\<Squnion>i. Y i i)" |
|
138 |
proof (rule antisym_less) |
|
139 |
have 3: "chain (\<lambda>i. Y i i)" |
|
140 |
apply (rule chainI) |
|
141 |
apply (rule trans_less) |
|
142 |
apply (rule chainE [OF 1]) |
|
143 |
apply (rule chainE [OF 2]) |
|
144 |
done |
|
145 |
have 4: "chain (\<lambda>i. \<Squnion>j. Y i j)" |
|
16203 | 146 |
by (rule ch2ch_lub [OF 1 2]) |
16201 | 147 |
show "(\<Squnion>i. \<Squnion>j. Y i j) \<sqsubseteq> (\<Squnion>i. Y i i)" |
148 |
apply (rule is_lub_thelub [OF 4]) |
|
149 |
apply (rule ub_rangeI) |
|
16203 | 150 |
apply (rule lub_mono3 [rule_format, OF 2 3]) |
16201 | 151 |
apply (rule exI) |
152 |
apply (rule trans_less) |
|
25922
cb04d05e95fb
rename lemma chain_mono3 -> chain_mono, chain_mono -> chain_mono_less
huffman
parents:
25921
diff
changeset
|
153 |
apply (rule chain_mono [OF 1 le_maxI1]) |
cb04d05e95fb
rename lemma chain_mono3 -> chain_mono, chain_mono -> chain_mono_less
huffman
parents:
25921
diff
changeset
|
154 |
apply (rule chain_mono [OF 2 le_maxI2]) |
16201 | 155 |
done |
156 |
show "(\<Squnion>i. Y i i) \<sqsubseteq> (\<Squnion>i. \<Squnion>j. Y i j)" |
|
25923 | 157 |
apply (rule lub_mono [OF 3 4]) |
16201 | 158 |
apply (rule is_ub_thelub [OF 2]) |
159 |
done |
|
160 |
qed |
|
161 |
||
162 |
lemma ex_lub: |
|
163 |
fixes Y :: "nat \<Rightarrow> nat \<Rightarrow> 'a::cpo" |
|
164 |
assumes 1: "\<And>j. chain (\<lambda>i. Y i j)" |
|
165 |
assumes 2: "\<And>i. chain (\<lambda>j. Y i j)" |
|
166 |
shows "(\<Squnion>i. \<Squnion>j. Y i j) = (\<Squnion>j. \<Squnion>i. Y i j)" |
|
167 |
by (simp add: diag_lub 1 2) |
|
168 |
||
169 |
||
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
170 |
subsection {* Pointed cpos *} |
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
171 |
|
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
172 |
text {* The class pcpo of pointed cpos *} |
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
173 |
|
25723 | 174 |
axclass pcpo < cpo |
175 |
least: "\<exists>x. \<forall>y. x \<sqsubseteq> y" |
|
176 |
||
177 |
definition |
|
178 |
UU :: "'a::pcpo" where |
|
179 |
"UU = (THE x. \<forall>y. x \<sqsubseteq> y)" |
|
180 |
||
181 |
notation (xsymbols) |
|
182 |
UU ("\<bottom>") |
|
183 |
||
184 |
text {* derive the old rule minimal *} |
|
185 |
||
186 |
lemma UU_least: "\<forall>z. \<bottom> \<sqsubseteq> z" |
|
187 |
apply (unfold UU_def) |
|
188 |
apply (rule theI') |
|
189 |
apply (rule ex_ex1I) |
|
190 |
apply (rule least) |
|
191 |
apply (blast intro: antisym_less) |
|
192 |
done |
|
193 |
||
194 |
lemma minimal [iff]: "\<bottom> \<sqsubseteq> x" |
|
195 |
by (rule UU_least [THEN spec]) |
|
196 |
||
197 |
lemma UU_reorient: "(\<bottom> = x) = (x = \<bottom>)" |
|
198 |
by auto |
|
16739 | 199 |
|
26480 | 200 |
ML {* |
25723 | 201 |
local |
202 |
val meta_UU_reorient = thm "UU_reorient" RS eq_reflection; |
|
203 |
fun reorient_proc sg _ (_ $ t $ u) = |
|
204 |
case u of |
|
205 |
Const("Pcpo.UU",_) => NONE |
|
206 |
| Const("HOL.zero", _) => NONE |
|
207 |
| Const("HOL.one", _) => NONE |
|
208 |
| Const("Numeral.number_of", _) $ _ => NONE |
|
209 |
| _ => SOME meta_UU_reorient; |
|
210 |
in |
|
211 |
val UU_reorient_simproc = |
|
28262
aa7ca36d67fd
back to dynamic the_context(), because static @{theory} is invalidated if ML environment changes within the same code block;
wenzelm
parents:
27415
diff
changeset
|
212 |
Simplifier.simproc (the_context ()) "UU_reorient_simproc" ["UU=x"] reorient_proc |
25723 | 213 |
end; |
214 |
||
215 |
Addsimprocs [UU_reorient_simproc]; |
|
216 |
*} |
|
217 |
||
218 |
text {* useful lemmas about @{term \<bottom>} *} |
|
219 |
||
220 |
lemma less_UU_iff [simp]: "(x \<sqsubseteq> \<bottom>) = (x = \<bottom>)" |
|
221 |
by (simp add: po_eq_conv) |
|
222 |
||
223 |
lemma eq_UU_iff: "(x = \<bottom>) = (x \<sqsubseteq> \<bottom>)" |
|
224 |
by simp |
|
225 |
||
226 |
lemma UU_I: "x \<sqsubseteq> \<bottom> \<Longrightarrow> x = \<bottom>" |
|
227 |
by (subst eq_UU_iff) |
|
228 |
||
229 |
lemma not_less2not_eq: "\<not> (x::'a::po) \<sqsubseteq> y \<Longrightarrow> x \<noteq> y" |
|
230 |
by auto |
|
231 |
||
232 |
lemma chain_UU_I: "\<lbrakk>chain Y; (\<Squnion>i. Y i) = \<bottom>\<rbrakk> \<Longrightarrow> \<forall>i. Y i = \<bottom>" |
|
15563 | 233 |
apply (rule allI) |
16626 | 234 |
apply (rule UU_I) |
15563 | 235 |
apply (erule subst) |
236 |
apply (erule is_ub_thelub) |
|
237 |
done |
|
238 |
||
16626 | 239 |
lemma chain_UU_I_inverse: "\<forall>i::nat. Y i = \<bottom> \<Longrightarrow> (\<Squnion>i. Y i) = \<bottom>" |
15563 | 240 |
apply (rule lub_chain_maxelem) |
241 |
apply (erule spec) |
|
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
242 |
apply simp |
15563 | 243 |
done |
244 |
||
16626 | 245 |
lemma chain_UU_I_inverse2: "(\<Squnion>i. Y i) \<noteq> \<bottom> \<Longrightarrow> \<exists>i::nat. Y i \<noteq> \<bottom>" |
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
246 |
by (blast intro: chain_UU_I_inverse) |
15563 | 247 |
|
16626 | 248 |
lemma notUU_I: "\<lbrakk>x \<sqsubseteq> y; x \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> y \<noteq> \<bottom>" |
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
249 |
by (blast intro: UU_I) |
15563 | 250 |
|
16627 | 251 |
lemma chain_mono2: "\<lbrakk>\<exists>j. Y j \<noteq> \<bottom>; chain Y\<rbrakk> \<Longrightarrow> \<exists>j. \<forall>i>j. Y i \<noteq> \<bottom>" |
25922
cb04d05e95fb
rename lemma chain_mono3 -> chain_mono, chain_mono -> chain_mono_less
huffman
parents:
25921
diff
changeset
|
252 |
by (blast dest: notUU_I chain_mono_less) |
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
253 |
|
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
254 |
subsection {* Chain-finite and flat cpos *} |
15563 | 255 |
|
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
256 |
text {* further useful classes for HOLCF domains *} |
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
257 |
|
25814 | 258 |
axclass finite_po < finite, po |
259 |
||
15640
2d1d6ea579a1
chfin now a subclass of po, proved instance chfin < cpo
huffman
parents:
15588
diff
changeset
|
260 |
axclass chfin < po |
25921 | 261 |
chfin: "chain Y \<Longrightarrow> \<exists>n. max_in_chain n Y" |
15563 | 262 |
|
25723 | 263 |
axclass flat < pcpo |
25920 | 264 |
ax_flat: "x \<sqsubseteq> y \<Longrightarrow> (x = \<bottom>) \<or> (x = y)" |
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
265 |
|
27415 | 266 |
text {* finite partial orders are chain-finite *} |
25814 | 267 |
|
268 |
instance finite_po < chfin |
|
25921 | 269 |
apply intro_classes |
25814 | 270 |
apply (drule finite_range_imp_finch) |
271 |
apply (rule finite) |
|
272 |
apply (simp add: finite_chain_def) |
|
273 |
done |
|
274 |
||
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
275 |
text {* some properties for chfin and flat *} |
15563 | 276 |
|
15640
2d1d6ea579a1
chfin now a subclass of po, proved instance chfin < cpo
huffman
parents:
15588
diff
changeset
|
277 |
text {* chfin types are cpo *} |
2d1d6ea579a1
chfin now a subclass of po, proved instance chfin < cpo
huffman
parents:
15588
diff
changeset
|
278 |
|
25921 | 279 |
instance chfin < cpo |
280 |
apply intro_classes |
|
281 |
apply (frule chfin) |
|
15640
2d1d6ea579a1
chfin now a subclass of po, proved instance chfin < cpo
huffman
parents:
15588
diff
changeset
|
282 |
apply (blast intro: lub_finch1) |
2d1d6ea579a1
chfin now a subclass of po, proved instance chfin < cpo
huffman
parents:
15588
diff
changeset
|
283 |
done |
2d1d6ea579a1
chfin now a subclass of po, proved instance chfin < cpo
huffman
parents:
15588
diff
changeset
|
284 |
|
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
285 |
text {* flat types are chfin *} |
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
286 |
|
25920 | 287 |
instance flat < chfin |
288 |
apply intro_classes |
|
15563 | 289 |
apply (unfold max_in_chain_def) |
16626 | 290 |
apply (case_tac "\<forall>i. Y i = \<bottom>") |
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
291 |
apply simp |
15563 | 292 |
apply simp |
293 |
apply (erule exE) |
|
16626 | 294 |
apply (rule_tac x="i" in exI) |
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
295 |
apply clarify |
25922
cb04d05e95fb
rename lemma chain_mono3 -> chain_mono, chain_mono -> chain_mono_less
huffman
parents:
25921
diff
changeset
|
296 |
apply (blast dest: chain_mono ax_flat) |
15563 | 297 |
done |
298 |
||
16627 | 299 |
text {* flat subclass of chfin; @{text adm_flat} not needed *} |
15563 | 300 |
|
25826 | 301 |
lemma flat_less_iff: |
302 |
fixes x y :: "'a::flat" |
|
303 |
shows "(x \<sqsubseteq> y) = (x = \<bottom> \<or> x = y)" |
|
25920 | 304 |
by (safe dest!: ax_flat) |
25826 | 305 |
|
16626 | 306 |
lemma flat_eq: "(a::'a::flat) \<noteq> \<bottom> \<Longrightarrow> a \<sqsubseteq> b = (a = b)" |
25920 | 307 |
by (safe dest!: ax_flat) |
15563 | 308 |
|
16626 | 309 |
lemma chfin2finch: "chain (Y::nat \<Rightarrow> 'a::chfin) \<Longrightarrow> finite_chain Y" |
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
310 |
by (simp add: chfin finite_chain_def) |
15563 | 311 |
|
26023 | 312 |
text {* Discrete cpos *} |
313 |
||
314 |
axclass discrete_cpo < sq_ord |
|
315 |
discrete_cpo [simp]: "x \<sqsubseteq> y \<longleftrightarrow> x = y" |
|
316 |
||
317 |
instance discrete_cpo < po |
|
318 |
by (intro_classes, simp_all) |
|
319 |
||
320 |
text {* In a discrete cpo, every chain is constant *} |
|
321 |
||
322 |
lemma discrete_chain_const: |
|
323 |
assumes S: "chain (S::nat \<Rightarrow> 'a::discrete_cpo)" |
|
324 |
shows "\<exists>x. S = (\<lambda>i. x)" |
|
325 |
proof (intro exI ext) |
|
326 |
fix i :: nat |
|
327 |
have "S 0 \<sqsubseteq> S i" using S le0 by (rule chain_mono) |
|
328 |
hence "S 0 = S i" by simp |
|
329 |
thus "S i = S 0" by (rule sym) |
|
330 |
qed |
|
331 |
||
332 |
instance discrete_cpo < cpo |
|
333 |
proof |
|
334 |
fix S :: "nat \<Rightarrow> 'a" |
|
335 |
assume S: "chain S" |
|
336 |
hence "\<exists>x. S = (\<lambda>i. x)" |
|
337 |
by (rule discrete_chain_const) |
|
338 |
thus "\<exists>x. range S <<| x" |
|
339 |
by (fast intro: lub_const) |
|
340 |
qed |
|
341 |
||
15588
14e3228f18cc
arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
342 |
text {* lemmata for improved admissibility introdution rule *} |
15563 | 343 |
|
344 |
lemma infinite_chain_adm_lemma: |
|
16626 | 345 |
"\<lbrakk>chain Y; \<forall>i. P (Y i); |
346 |
\<And>Y. \<lbrakk>chain Y; \<forall>i. P (Y i); \<not> finite_chain Y\<rbrakk> \<Longrightarrow> P (\<Squnion>i. Y i)\<rbrakk> |
|
347 |
\<Longrightarrow> P (\<Squnion>i. Y i)" |
|
15563 | 348 |
apply (case_tac "finite_chain Y") |
349 |
prefer 2 apply fast |
|
350 |
apply (unfold finite_chain_def) |
|
351 |
apply safe |
|
352 |
apply (erule lub_finch1 [THEN thelubI, THEN ssubst]) |
|
353 |
apply assumption |
|
354 |
apply (erule spec) |
|
355 |
done |
|
356 |
||
357 |
lemma increasing_chain_adm_lemma: |
|
16626 | 358 |
"\<lbrakk>chain Y; \<forall>i. P (Y i); \<And>Y. \<lbrakk>chain Y; \<forall>i. P (Y i); |
359 |
\<forall>i. \<exists>j>i. Y i \<noteq> Y j \<and> Y i \<sqsubseteq> Y j\<rbrakk> \<Longrightarrow> P (\<Squnion>i. Y i)\<rbrakk> |
|
360 |
\<Longrightarrow> P (\<Squnion>i. Y i)" |
|
15563 | 361 |
apply (erule infinite_chain_adm_lemma) |
362 |
apply assumption |
|
363 |
apply (erule thin_rl) |
|
364 |
apply (unfold finite_chain_def) |
|
365 |
apply (unfold max_in_chain_def) |
|
25922
cb04d05e95fb
rename lemma chain_mono3 -> chain_mono, chain_mono -> chain_mono_less
huffman
parents:
25921
diff
changeset
|
366 |
apply (fast dest: le_imp_less_or_eq elim: chain_mono_less) |
15563 | 367 |
done |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
15563
diff
changeset
|
368 |
|
16626 | 369 |
end |