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