| author | Brian Huffman <brianh@cs.pdx.edu> | 
| Tue, 05 Oct 2010 17:07:57 -0700 | |
| changeset 39970 | 9023b897e67a | 
| parent 39969 | 0b8e19f588a4 | 
| child 40045 | e0f372e18f3e | 
| 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: 
15563diff
changeset | 4 | |
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: 
15563diff
changeset | 5 | header {* Classes cpo and pcpo *}
 | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: 
15563diff
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: 
15577diff
changeset | 11 | subsection {* Complete partial orders *}
 | 
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 12 | |
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 13 | text {* The class cpo of chain complete partial orders *}
 | 
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 14 | |
| 29614 
1f7b1b0df292
simplified handling of base sort, dropped axclass
 haftmann parents: 
29138diff
changeset | 15 | class cpo = po + | 
| 31071 | 16 | assumes cpo: "chain S \<Longrightarrow> \<exists>x. range S <<| x" | 
| 17 | begin | |
| 2394 | 18 | |
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 19 | text {* in cpo's everthing equal to THE lub has lub properties for every chain *}
 | 
| 15563 | 20 | |
| 31071 | 21 | lemma cpo_lubI: "chain S \<Longrightarrow> range S <<| (\<Squnion>i. S i)" | 
| 22 | by (fast dest: cpo elim: lubI) | |
| 26026 | 23 | |
| 31071 | 24 | lemma thelubE: "\<lbrakk>chain S; (\<Squnion>i. S i) = l\<rbrakk> \<Longrightarrow> range S <<| l" | 
| 25 | by (blast dest: cpo intro: lubI) | |
| 15563 | 26 | |
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 27 | text {* Properties of the lub *}
 | 
| 15563 | 28 | |
| 31071 | 29 | lemma is_ub_thelub: "chain S \<Longrightarrow> S x \<sqsubseteq> (\<Squnion>i. S i)" | 
| 30 | by (blast dest: cpo intro: lubI [THEN is_ub_lub]) | |
| 15563 | 31 | |
| 16626 | 32 | lemma is_lub_thelub: | 
| 31071 | 33 | "\<lbrakk>chain S; range S <| x\<rbrakk> \<Longrightarrow> (\<Squnion>i. S i) \<sqsubseteq> x" | 
| 34 | by (blast dest: cpo intro: lubI [THEN is_lub_lub]) | |
| 15563 | 35 | |
| 39969 | 36 | lemma lub_below_iff: "chain S \<Longrightarrow> (\<Squnion>i. S i) \<sqsubseteq> x \<longleftrightarrow> (\<forall>i. S i \<sqsubseteq> x)" | 
| 37 | by (simp add: is_lub_below_iff [OF cpo_lubI] is_ub_def) | |
| 38 | ||
| 16626 | 39 | lemma lub_range_mono: | 
| 31071 | 40 | "\<lbrakk>range X \<subseteq> range Y; chain Y; chain X\<rbrakk> | 
| 16626 | 41 | \<Longrightarrow> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)" | 
| 15563 | 42 | apply (erule is_lub_thelub) | 
| 43 | apply (rule ub_rangeI) | |
| 16626 | 44 | apply (subgoal_tac "\<exists>j. X i = Y j") | 
| 15563 | 45 | apply clarsimp | 
| 46 | apply (erule is_ub_thelub) | |
| 47 | apply auto | |
| 48 | done | |
| 49 | ||
| 16626 | 50 | lemma lub_range_shift: | 
| 31071 | 51 | "chain Y \<Longrightarrow> (\<Squnion>i. Y (i + j)) = (\<Squnion>i. Y i)" | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
31071diff
changeset | 52 | apply (rule below_antisym) | 
| 15563 | 53 | apply (rule lub_range_mono) | 
| 54 | apply fast | |
| 55 | apply assumption | |
| 56 | apply (erule chain_shift) | |
| 57 | apply (rule is_lub_thelub) | |
| 58 | apply assumption | |
| 59 | apply (rule ub_rangeI) | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
31071diff
changeset | 60 | apply (rule_tac y="Y (i + j)" in below_trans) | 
| 25922 
cb04d05e95fb
rename lemma chain_mono3 -> chain_mono, chain_mono -> chain_mono_less
 huffman parents: 
25921diff
changeset | 61 | apply (erule chain_mono) | 
| 15563 | 62 | apply (rule le_add1) | 
| 17813 | 63 | apply (rule is_ub_thelub) | 
| 64 | apply (erule chain_shift) | |
| 15563 | 65 | done | 
| 66 | ||
| 16626 | 67 | lemma maxinch_is_thelub: | 
| 31071 | 68 | "chain Y \<Longrightarrow> max_in_chain i Y = ((\<Squnion>i. Y i) = Y i)" | 
| 15563 | 69 | apply (rule iffI) | 
| 70 | apply (fast intro!: thelubI lub_finch1) | |
| 71 | apply (unfold max_in_chain_def) | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
31071diff
changeset | 72 | apply (safe intro!: below_antisym) | 
| 25922 
cb04d05e95fb
rename lemma chain_mono3 -> chain_mono, chain_mono -> chain_mono_less
 huffman parents: 
25921diff
changeset | 73 | apply (fast elim!: chain_mono) | 
| 15563 | 74 | apply (drule sym) | 
| 75 | apply (force elim!: is_ub_thelub) | |
| 76 | done | |
| 77 | ||
| 16626 | 78 | text {* the @{text "\<sqsubseteq>"} relation between two chains is preserved by their lubs *}
 | 
| 15563 | 79 | |
| 16626 | 80 | lemma lub_mono: | 
| 31071 | 81 | "\<lbrakk>chain X; chain Y; \<And>i. X i \<sqsubseteq> Y i\<rbrakk> | 
| 16626 | 82 | \<Longrightarrow> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)" | 
| 15563 | 83 | apply (erule is_lub_thelub) | 
| 84 | apply (rule ub_rangeI) | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
31071diff
changeset | 85 | apply (rule below_trans) | 
| 25923 | 86 | apply (erule meta_spec) | 
| 15563 | 87 | apply (erule is_ub_thelub) | 
| 88 | done | |
| 89 | ||
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 90 | text {* the = relation between two chains is preserved by their lubs *}
 | 
| 15563 | 91 | |
| 16626 | 92 | lemma lub_equal: | 
| 31071 | 93 | "\<lbrakk>chain X; chain Y; \<forall>k. X k = Y k\<rbrakk> | 
| 16626 | 94 | \<Longrightarrow> (\<Squnion>i. X i) = (\<Squnion>i. Y i)" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39199diff
changeset | 95 | by (simp only: fun_eq_iff [symmetric]) | 
| 15563 | 96 | |
| 35492 | 97 | lemma lub_eq: | 
| 98 | "(\<And>i. X i = Y i) \<Longrightarrow> (\<Squnion>i. X i) = (\<Squnion>i. Y i)" | |
| 99 | by simp | |
| 100 | ||
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 101 | text {* more results about mono and = of lubs of chains *}
 | 
| 3326 | 102 | |
| 16626 | 103 | lemma lub_mono2: | 
| 31071 | 104 | "\<lbrakk>\<exists>j. \<forall>i>j. X i = Y i; chain X; chain Y\<rbrakk> | 
| 16626 | 105 | \<Longrightarrow> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)" | 
| 15563 | 106 | apply (erule exE) | 
| 17813 | 107 | apply (subgoal_tac "(\<Squnion>i. X (i + Suc j)) \<sqsubseteq> (\<Squnion>i. Y (i + Suc j))") | 
| 108 | apply (thin_tac "\<forall>i>j. X i = Y i") | |
| 109 | apply (simp only: lub_range_shift) | |
| 16626 | 110 | apply simp | 
| 15563 | 111 | done | 
| 112 | ||
| 16626 | 113 | lemma lub_equal2: | 
| 31071 | 114 | "\<lbrakk>\<exists>j. \<forall>i>j. X i = Y i; chain X; chain Y\<rbrakk> | 
| 16626 | 115 | \<Longrightarrow> (\<Squnion>i. X i) = (\<Squnion>i. Y i)" | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
31071diff
changeset | 116 | by (blast intro: below_antisym lub_mono2 sym) | 
| 15563 | 117 | |
| 16626 | 118 | lemma lub_mono3: | 
| 31071 | 119 | "\<lbrakk>chain Y; chain X; \<forall>i. \<exists>j. Y i \<sqsubseteq> X j\<rbrakk> | 
| 16626 | 120 | \<Longrightarrow> (\<Squnion>i. Y i) \<sqsubseteq> (\<Squnion>i. X i)" | 
| 17813 | 121 | apply (erule is_lub_thelub) | 
| 15563 | 122 | apply (rule ub_rangeI) | 
| 123 | apply (erule allE) | |
| 124 | apply (erule exE) | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
31071diff
changeset | 125 | apply (erule below_trans) | 
| 16626 | 126 | apply (erule is_ub_thelub) | 
| 15563 | 127 | done | 
| 128 | ||
| 16203 | 129 | lemma ch2ch_lub: | 
| 130 | assumes 1: "\<And>j. chain (\<lambda>i. Y i j)" | |
| 131 | assumes 2: "\<And>i. chain (\<lambda>j. Y i j)" | |
| 132 | shows "chain (\<lambda>i. \<Squnion>j. Y i j)" | |
| 133 | apply (rule chainI) | |
| 25923 | 134 | apply (rule lub_mono [OF 2 2]) | 
| 16203 | 135 | apply (rule chainE [OF 1]) | 
| 136 | done | |
| 137 | ||
| 16201 | 138 | lemma diag_lub: | 
| 139 | assumes 1: "\<And>j. chain (\<lambda>i. Y i j)" | |
| 140 | assumes 2: "\<And>i. chain (\<lambda>j. Y i j)" | |
| 141 | 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: 
31071diff
changeset | 142 | proof (rule below_antisym) | 
| 16201 | 143 | have 3: "chain (\<lambda>i. Y i i)" | 
| 144 | 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: 
31071diff
changeset | 145 | apply (rule below_trans) | 
| 16201 | 146 | apply (rule chainE [OF 1]) | 
| 147 | apply (rule chainE [OF 2]) | |
| 148 | done | |
| 149 | have 4: "chain (\<lambda>i. \<Squnion>j. Y i j)" | |
| 16203 | 150 | by (rule ch2ch_lub [OF 1 2]) | 
| 16201 | 151 | show "(\<Squnion>i. \<Squnion>j. Y i j) \<sqsubseteq> (\<Squnion>i. Y i i)" | 
| 152 | apply (rule is_lub_thelub [OF 4]) | |
| 153 | apply (rule ub_rangeI) | |
| 16203 | 154 | apply (rule lub_mono3 [rule_format, OF 2 3]) | 
| 16201 | 155 | apply (rule exI) | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
31071diff
changeset | 156 | apply (rule below_trans) | 
| 25922 
cb04d05e95fb
rename lemma chain_mono3 -> chain_mono, chain_mono -> chain_mono_less
 huffman parents: 
25921diff
changeset | 157 | apply (rule chain_mono [OF 1 le_maxI1]) | 
| 
cb04d05e95fb
rename lemma chain_mono3 -> chain_mono, chain_mono -> chain_mono_less
 huffman parents: 
25921diff
changeset | 158 | apply (rule chain_mono [OF 2 le_maxI2]) | 
| 16201 | 159 | done | 
| 160 | show "(\<Squnion>i. Y i i) \<sqsubseteq> (\<Squnion>i. \<Squnion>j. Y i j)" | |
| 25923 | 161 | apply (rule lub_mono [OF 3 4]) | 
| 16201 | 162 | apply (rule is_ub_thelub [OF 2]) | 
| 163 | done | |
| 164 | qed | |
| 165 | ||
| 166 | lemma ex_lub: | |
| 167 | assumes 1: "\<And>j. chain (\<lambda>i. Y i j)" | |
| 168 | assumes 2: "\<And>i. chain (\<lambda>j. Y i j)" | |
| 169 | shows "(\<Squnion>i. \<Squnion>j. Y i j) = (\<Squnion>j. \<Squnion>i. Y i j)" | |
| 31071 | 170 | by (simp add: diag_lub 1 2) | 
| 16201 | 171 | |
| 31071 | 172 | end | 
| 16201 | 173 | |
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 174 | subsection {* Pointed cpos *}
 | 
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 175 | |
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 176 | text {* The class pcpo of pointed cpos *}
 | 
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 177 | |
| 29614 
1f7b1b0df292
simplified handling of base sort, dropped axclass
 haftmann parents: 
29138diff
changeset | 178 | class pcpo = cpo + | 
| 
1f7b1b0df292
simplified handling of base sort, dropped axclass
 haftmann parents: 
29138diff
changeset | 179 | assumes least: "\<exists>x. \<forall>y. x \<sqsubseteq> y" | 
| 31071 | 180 | begin | 
| 25723 | 181 | |
| 31071 | 182 | definition UU :: 'a where | 
| 25723 | 183 | "UU = (THE x. \<forall>y. x \<sqsubseteq> y)" | 
| 184 | ||
| 185 | notation (xsymbols) | |
| 186 |   UU  ("\<bottom>")
 | |
| 187 | ||
| 188 | text {* derive the old rule minimal *}
 | |
| 189 | ||
| 190 | lemma UU_least: "\<forall>z. \<bottom> \<sqsubseteq> z" | |
| 191 | apply (unfold UU_def) | |
| 192 | apply (rule theI') | |
| 193 | apply (rule ex_ex1I) | |
| 194 | apply (rule least) | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
31071diff
changeset | 195 | apply (blast intro: below_antisym) | 
| 25723 | 196 | done | 
| 197 | ||
| 198 | lemma minimal [iff]: "\<bottom> \<sqsubseteq> x" | |
| 199 | by (rule UU_least [THEN spec]) | |
| 200 | ||
| 31071 | 201 | end | 
| 202 | ||
| 31024 
0fdf666e08bf
reimplement reorientation simproc using theory data
 huffman parents: 
29634diff
changeset | 203 | text {* Simproc to rewrite @{term "\<bottom> = x"} to @{term "x = \<bottom>"}. *}
 | 
| 16739 | 204 | |
| 31024 
0fdf666e08bf
reimplement reorientation simproc using theory data
 huffman parents: 
29634diff
changeset | 205 | setup {*
 | 
| 33523 | 206 | Reorient_Proc.add | 
| 31024 
0fdf666e08bf
reimplement reorientation simproc using theory data
 huffman parents: 
29634diff
changeset | 207 |     (fn Const(@{const_name UU}, _) => true | _ => false)
 | 
| 
0fdf666e08bf
reimplement reorientation simproc using theory data
 huffman parents: 
29634diff
changeset | 208 | *} | 
| 25723 | 209 | |
| 33523 | 210 | simproc_setup reorient_bottom ("\<bottom> = x") = Reorient_Proc.proc
 | 
| 25723 | 211 | |
| 31071 | 212 | context pcpo | 
| 213 | begin | |
| 214 | ||
| 25723 | 215 | text {* useful lemmas about @{term \<bottom>} *}
 | 
| 216 | ||
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
31071diff
changeset | 217 | lemma below_UU_iff [simp]: "(x \<sqsubseteq> \<bottom>) = (x = \<bottom>)" | 
| 25723 | 218 | by (simp add: po_eq_conv) | 
| 219 | ||
| 220 | lemma eq_UU_iff: "(x = \<bottom>) = (x \<sqsubseteq> \<bottom>)" | |
| 221 | by simp | |
| 222 | ||
| 223 | lemma UU_I: "x \<sqsubseteq> \<bottom> \<Longrightarrow> x = \<bottom>" | |
| 224 | by (subst eq_UU_iff) | |
| 225 | ||
| 226 | lemma chain_UU_I: "\<lbrakk>chain Y; (\<Squnion>i. Y i) = \<bottom>\<rbrakk> \<Longrightarrow> \<forall>i. Y i = \<bottom>" | |
| 15563 | 227 | apply (rule allI) | 
| 16626 | 228 | apply (rule UU_I) | 
| 15563 | 229 | apply (erule subst) | 
| 230 | apply (erule is_ub_thelub) | |
| 231 | done | |
| 232 | ||
| 16626 | 233 | lemma chain_UU_I_inverse: "\<forall>i::nat. Y i = \<bottom> \<Longrightarrow> (\<Squnion>i. Y i) = \<bottom>" | 
| 15563 | 234 | apply (rule lub_chain_maxelem) | 
| 235 | apply (erule spec) | |
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 236 | apply simp | 
| 15563 | 237 | done | 
| 238 | ||
| 16626 | 239 | lemma chain_UU_I_inverse2: "(\<Squnion>i. Y i) \<noteq> \<bottom> \<Longrightarrow> \<exists>i::nat. Y i \<noteq> \<bottom>" | 
| 31071 | 240 | by (blast intro: chain_UU_I_inverse) | 
| 15563 | 241 | |
| 16626 | 242 | lemma notUU_I: "\<lbrakk>x \<sqsubseteq> y; x \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> y \<noteq> \<bottom>" | 
| 31071 | 243 | by (blast intro: UU_I) | 
| 15563 | 244 | |
| 16627 | 245 | lemma chain_mono2: "\<lbrakk>\<exists>j. Y j \<noteq> \<bottom>; chain Y\<rbrakk> \<Longrightarrow> \<exists>j. \<forall>i>j. Y i \<noteq> \<bottom>" | 
| 31071 | 246 | by (blast dest: notUU_I chain_mono_less) | 
| 247 | ||
| 248 | end | |
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 249 | |
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 250 | subsection {* Chain-finite and flat cpos *}
 | 
| 15563 | 251 | |
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 252 | text {* further useful classes for HOLCF domains *}
 | 
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 253 | |
| 31071 | 254 | class chfin = po + | 
| 255 | assumes chfin: "chain Y \<Longrightarrow> \<exists>n. max_in_chain n Y" | |
| 256 | begin | |
| 25814 | 257 | |
| 31071 | 258 | subclass cpo | 
| 259 | apply default | |
| 260 | apply (frule chfin) | |
| 261 | apply (blast intro: lub_finch1) | |
| 262 | done | |
| 15563 | 263 | |
| 31071 | 264 | lemma chfin2finch: "chain Y \<Longrightarrow> finite_chain Y" | 
| 265 | by (simp add: chfin finite_chain_def) | |
| 266 | ||
| 267 | end | |
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 268 | |
| 31071 | 269 | class finite_po = finite + po | 
| 270 | begin | |
| 25814 | 271 | |
| 31071 | 272 | subclass chfin | 
| 273 | apply default | |
| 25814 | 274 | apply (drule finite_range_imp_finch) | 
| 275 | apply (rule finite) | |
| 276 | apply (simp add: finite_chain_def) | |
| 277 | done | |
| 278 | ||
| 31071 | 279 | end | 
| 15640 
2d1d6ea579a1
chfin now a subclass of po, proved instance chfin < cpo
 huffman parents: 
15588diff
changeset | 280 | |
| 31071 | 281 | class flat = pcpo + | 
| 282 | assumes ax_flat: "x \<sqsubseteq> y \<Longrightarrow> x = \<bottom> \<or> x = y" | |
| 283 | begin | |
| 15640 
2d1d6ea579a1
chfin now a subclass of po, proved instance chfin < cpo
 huffman parents: 
15588diff
changeset | 284 | |
| 31071 | 285 | subclass chfin | 
| 286 | apply default | |
| 15563 | 287 | apply (unfold max_in_chain_def) | 
| 16626 | 288 | apply (case_tac "\<forall>i. Y i = \<bottom>") | 
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 289 | apply simp | 
| 15563 | 290 | apply simp | 
| 291 | apply (erule exE) | |
| 16626 | 292 | apply (rule_tac x="i" in exI) | 
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 293 | apply clarify | 
| 25922 
cb04d05e95fb
rename lemma chain_mono3 -> chain_mono, chain_mono -> chain_mono_less
 huffman parents: 
25921diff
changeset | 294 | apply (blast dest: chain_mono ax_flat) | 
| 15563 | 295 | done | 
| 296 | ||
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
31071diff
changeset | 297 | lemma flat_below_iff: | 
| 25826 | 298 | shows "(x \<sqsubseteq> y) = (x = \<bottom> \<or> x = y)" | 
| 31071 | 299 | by (safe dest!: ax_flat) | 
| 25826 | 300 | |
| 31071 | 301 | lemma flat_eq: "a \<noteq> \<bottom> \<Longrightarrow> a \<sqsubseteq> b = (a = b)" | 
| 302 | by (safe dest!: ax_flat) | |
| 15563 | 303 | |
| 31071 | 304 | end | 
| 15563 | 305 | |
| 26023 | 306 | text {* Discrete cpos *}
 | 
| 307 | ||
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
31071diff
changeset | 308 | class discrete_cpo = below + | 
| 29614 
1f7b1b0df292
simplified handling of base sort, dropped axclass
 haftmann parents: 
29138diff
changeset | 309 | assumes discrete_cpo [simp]: "x \<sqsubseteq> y \<longleftrightarrow> x = y" | 
| 31071 | 310 | begin | 
| 26023 | 311 | |
| 31071 | 312 | subclass po | 
| 29614 
1f7b1b0df292
simplified handling of base sort, dropped axclass
 haftmann parents: 
29138diff
changeset | 313 | proof qed simp_all | 
| 26023 | 314 | |
| 315 | text {* In a discrete cpo, every chain is constant *}
 | |
| 316 | ||
| 317 | lemma discrete_chain_const: | |
| 31071 | 318 | assumes S: "chain S" | 
| 26023 | 319 | shows "\<exists>x. S = (\<lambda>i. x)" | 
| 320 | proof (intro exI ext) | |
| 321 | fix i :: nat | |
| 322 | have "S 0 \<sqsubseteq> S i" using S le0 by (rule chain_mono) | |
| 323 | hence "S 0 = S i" by simp | |
| 324 | thus "S i = S 0" by (rule sym) | |
| 325 | qed | |
| 326 | ||
| 31071 | 327 | subclass cpo | 
| 26023 | 328 | proof | 
| 329 | fix S :: "nat \<Rightarrow> 'a" | |
| 330 | assume S: "chain S" | |
| 331 | hence "\<exists>x. S = (\<lambda>i. x)" | |
| 332 | by (rule discrete_chain_const) | |
| 333 | thus "\<exists>x. range S <<| x" | |
| 334 | by (fast intro: lub_const) | |
| 335 | qed | |
| 336 | ||
| 31071 | 337 | end | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: 
15563diff
changeset | 338 | |
| 16626 | 339 | end |