| author | wenzelm | 
| Tue, 10 Jan 2006 19:33:37 +0100 | |
| changeset 18638 | e135f6a1b76c | 
| parent 17813 | 03133f6606a1 | 
| child 18924 | 83acd39b1bab | 
| 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: 
15563diff
changeset | 5 | |
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: 
15563diff
changeset | 6 | header {* Classes cpo and pcpo *}
 | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: 
15563diff
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: 
15577diff
changeset | 12 | subsection {* Complete partial orders *}
 | 
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 13 | |
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 14 | text {* The class cpo of chain complete partial orders *}
 | 
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 15 | |
| 2640 | 16 | axclass cpo < po | 
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 17 |         -- {* class axiom: *}
 | 
| 16626 | 18 | cpo: "chain S \<Longrightarrow> \<exists>x. range S <<| x" | 
| 2394 | 19 | |
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 20 | text {* in cpo's everthing equal to THE lub has lub properties for every chain *}
 | 
| 15563 | 21 | |
| 16626 | 22 | 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: 
15577diff
changeset | 23 | by (blast dest: cpo intro: lubI) | 
| 15563 | 24 | |
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 25 | text {* Properties of the lub *}
 | 
| 15563 | 26 | |
| 16626 | 27 | 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: 
15577diff
changeset | 28 | by (blast dest: cpo intro: lubI [THEN is_ub_lub]) | 
| 15563 | 29 | |
| 16626 | 30 | lemma is_lub_thelub: | 
| 31 | "\<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: 
15577diff
changeset | 32 | by (blast dest: cpo intro: lubI [THEN is_lub_lub]) | 
| 15563 | 33 | |
| 16626 | 34 | lemma lub_range_mono: | 
| 35 | "\<lbrakk>range X \<subseteq> range Y; chain Y; chain (X::nat \<Rightarrow> 'a::cpo)\<rbrakk> | |
| 36 | \<Longrightarrow> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)" | |
| 15563 | 37 | apply (erule is_lub_thelub) | 
| 38 | apply (rule ub_rangeI) | |
| 16626 | 39 | apply (subgoal_tac "\<exists>j. X i = Y j") | 
| 15563 | 40 | apply clarsimp | 
| 41 | apply (erule is_ub_thelub) | |
| 42 | apply auto | |
| 43 | done | |
| 44 | ||
| 16626 | 45 | lemma lub_range_shift: | 
| 46 | "chain (Y::nat \<Rightarrow> 'a::cpo) \<Longrightarrow> (\<Squnion>i. Y (i + j)) = (\<Squnion>i. Y i)" | |
| 15563 | 47 | apply (rule antisym_less) | 
| 48 | apply (rule lub_range_mono) | |
| 49 | apply fast | |
| 50 | apply assumption | |
| 51 | apply (erule chain_shift) | |
| 52 | apply (rule is_lub_thelub) | |
| 53 | apply assumption | |
| 54 | apply (rule ub_rangeI) | |
| 17813 | 55 | apply (rule_tac y="Y (i + j)" in trans_less) | 
| 15563 | 56 | apply (erule chain_mono3) | 
| 57 | apply (rule le_add1) | |
| 17813 | 58 | apply (rule is_ub_thelub) | 
| 59 | apply (erule chain_shift) | |
| 15563 | 60 | done | 
| 61 | ||
| 16626 | 62 | lemma maxinch_is_thelub: | 
| 63 | "chain Y \<Longrightarrow> max_in_chain i Y = ((\<Squnion>i. Y i) = ((Y i)::'a::cpo))" | |
| 15563 | 64 | apply (rule iffI) | 
| 65 | apply (fast intro!: thelubI lub_finch1) | |
| 66 | apply (unfold max_in_chain_def) | |
| 67 | apply (safe intro!: antisym_less) | |
| 68 | apply (fast elim!: chain_mono3) | |
| 69 | apply (drule sym) | |
| 70 | apply (force elim!: is_ub_thelub) | |
| 71 | done | |
| 72 | ||
| 16626 | 73 | text {* the @{text "\<sqsubseteq>"} relation between two chains is preserved by their lubs *}
 | 
| 15563 | 74 | |
| 16626 | 75 | lemma lub_mono: | 
| 76 | "\<lbrakk>chain (X::nat \<Rightarrow> 'a::cpo); chain Y; \<forall>k. X k \<sqsubseteq> Y k\<rbrakk> | |
| 77 | \<Longrightarrow> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)" | |
| 15563 | 78 | apply (erule is_lub_thelub) | 
| 79 | apply (rule ub_rangeI) | |
| 80 | apply (rule trans_less) | |
| 81 | apply (erule spec) | |
| 82 | apply (erule is_ub_thelub) | |
| 83 | done | |
| 84 | ||
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 85 | text {* the = relation between two chains is preserved by their lubs *}
 | 
| 15563 | 86 | |
| 16626 | 87 | lemma lub_equal: | 
| 88 | "\<lbrakk>chain (X::nat \<Rightarrow> 'a::cpo); chain Y; \<forall>k. X k = Y k\<rbrakk> | |
| 89 | \<Longrightarrow> (\<Squnion>i. X i) = (\<Squnion>i. Y i)" | |
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 90 | by (simp only: expand_fun_eq [symmetric]) | 
| 15563 | 91 | |
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 92 | text {* more results about mono and = of lubs of chains *}
 | 
| 3326 | 93 | |
| 16626 | 94 | lemma lub_mono2: | 
| 17813 | 95 | "\<lbrakk>\<exists>j. \<forall>i>j. X i = Y i; chain (X::nat \<Rightarrow> 'a::cpo); chain Y\<rbrakk> | 
| 16626 | 96 | \<Longrightarrow> (\<Squnion>i. X i) \<sqsubseteq> (\<Squnion>i. Y i)" | 
| 15563 | 97 | apply (erule exE) | 
| 17813 | 98 | apply (subgoal_tac "(\<Squnion>i. X (i + Suc j)) \<sqsubseteq> (\<Squnion>i. Y (i + Suc j))") | 
| 99 | apply (thin_tac "\<forall>i>j. X i = Y i") | |
| 100 | apply (simp only: lub_range_shift) | |
| 16626 | 101 | apply simp | 
| 15563 | 102 | done | 
| 103 | ||
| 16626 | 104 | lemma lub_equal2: | 
| 105 | "\<lbrakk>\<exists>j. \<forall>i>j. X i = Y i; chain (X::nat \<Rightarrow> 'a::cpo); chain Y\<rbrakk> | |
| 106 | \<Longrightarrow> (\<Squnion>i. X i) = (\<Squnion>i. Y i)" | |
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 107 | by (blast intro: antisym_less lub_mono2 sym) | 
| 15563 | 108 | |
| 16626 | 109 | lemma lub_mono3: | 
| 110 | "\<lbrakk>chain (Y::nat \<Rightarrow> 'a::cpo); chain X; \<forall>i. \<exists>j. Y i \<sqsubseteq> X j\<rbrakk> | |
| 111 | \<Longrightarrow> (\<Squnion>i. Y i) \<sqsubseteq> (\<Squnion>i. X i)" | |
| 17813 | 112 | apply (erule is_lub_thelub) | 
| 15563 | 113 | apply (rule ub_rangeI) | 
| 114 | apply (erule allE) | |
| 115 | apply (erule exE) | |
| 16626 | 116 | apply (erule trans_less) | 
| 117 | apply (erule is_ub_thelub) | |
| 15563 | 118 | done | 
| 119 | ||
| 16203 | 120 | lemma ch2ch_lub: | 
| 121 | fixes Y :: "nat \<Rightarrow> nat \<Rightarrow> 'a::cpo" | |
| 122 | assumes 1: "\<And>j. chain (\<lambda>i. Y i j)" | |
| 123 | assumes 2: "\<And>i. chain (\<lambda>j. Y i j)" | |
| 124 | shows "chain (\<lambda>i. \<Squnion>j. Y i j)" | |
| 125 | apply (rule chainI) | |
| 126 | apply (rule lub_mono [rule_format, OF 2 2]) | |
| 127 | apply (rule chainE [OF 1]) | |
| 128 | done | |
| 129 | ||
| 16201 | 130 | lemma diag_lub: | 
| 131 | fixes Y :: "nat \<Rightarrow> nat \<Rightarrow> 'a::cpo" | |
| 132 | assumes 1: "\<And>j. chain (\<lambda>i. Y i j)" | |
| 133 | assumes 2: "\<And>i. chain (\<lambda>j. Y i j)" | |
| 134 | shows "(\<Squnion>i. \<Squnion>j. Y i j) = (\<Squnion>i. Y i i)" | |
| 135 | proof (rule antisym_less) | |
| 136 | have 3: "chain (\<lambda>i. Y i i)" | |
| 137 | apply (rule chainI) | |
| 138 | apply (rule trans_less) | |
| 139 | apply (rule chainE [OF 1]) | |
| 140 | apply (rule chainE [OF 2]) | |
| 141 | done | |
| 142 | have 4: "chain (\<lambda>i. \<Squnion>j. Y i j)" | |
| 16203 | 143 | by (rule ch2ch_lub [OF 1 2]) | 
| 16201 | 144 | show "(\<Squnion>i. \<Squnion>j. Y i j) \<sqsubseteq> (\<Squnion>i. Y i i)" | 
| 145 | apply (rule is_lub_thelub [OF 4]) | |
| 146 | apply (rule ub_rangeI) | |
| 16203 | 147 | apply (rule lub_mono3 [rule_format, OF 2 3]) | 
| 16201 | 148 | apply (rule exI) | 
| 149 | apply (rule trans_less) | |
| 150 | apply (rule chain_mono3 [OF 1 le_maxI1]) | |
| 151 | apply (rule chain_mono3 [OF 2 le_maxI2]) | |
| 152 | done | |
| 153 | show "(\<Squnion>i. Y i i) \<sqsubseteq> (\<Squnion>i. \<Squnion>j. Y i j)" | |
| 16203 | 154 | apply (rule lub_mono [rule_format, OF 3 4]) | 
| 16201 | 155 | apply (rule is_ub_thelub [OF 2]) | 
| 156 | done | |
| 157 | qed | |
| 158 | ||
| 159 | lemma ex_lub: | |
| 160 | fixes Y :: "nat \<Rightarrow> nat \<Rightarrow> 'a::cpo" | |
| 161 | assumes 1: "\<And>j. chain (\<lambda>i. Y i j)" | |
| 162 | assumes 2: "\<And>i. chain (\<lambda>j. Y i j)" | |
| 163 | shows "(\<Squnion>i. \<Squnion>j. Y i j) = (\<Squnion>j. \<Squnion>i. Y i j)" | |
| 164 | by (simp add: diag_lub 1 2) | |
| 165 | ||
| 166 | ||
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 167 | subsection {* Pointed cpos *}
 | 
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 168 | |
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 169 | text {* The class pcpo of pointed cpos *}
 | 
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 170 | |
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 171 | axclass pcpo < cpo | 
| 16626 | 172 | least: "\<exists>x. \<forall>y. x \<sqsubseteq> y" | 
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 173 | |
| 16626 | 174 | constdefs | 
| 175 | UU :: "'a::pcpo" | |
| 17813 | 176 | "UU \<equiv> THE x. \<forall>y. x \<sqsubseteq> y" | 
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 177 | |
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 178 | syntax (xsymbols) | 
| 16626 | 179 |   UU :: "'a::pcpo" ("\<bottom>")
 | 
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 180 | |
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 181 | text {* derive the old rule minimal *}
 | 
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 182 | |
| 16626 | 183 | lemma UU_least: "\<forall>z. \<bottom> \<sqsubseteq> z" | 
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 184 | apply (unfold UU_def) | 
| 15930 
145651bc64a8
Replaced all unnecessary uses of SOME with THE or LEAST
 huffman parents: 
15640diff
changeset | 185 | apply (rule theI') | 
| 
145651bc64a8
Replaced all unnecessary uses of SOME with THE or LEAST
 huffman parents: 
15640diff
changeset | 186 | apply (rule ex_ex1I) | 
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 187 | apply (rule least) | 
| 15930 
145651bc64a8
Replaced all unnecessary uses of SOME with THE or LEAST
 huffman parents: 
15640diff
changeset | 188 | apply (blast intro: antisym_less) | 
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 189 | done | 
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 190 | |
| 16626 | 191 | lemma minimal [iff]: "\<bottom> \<sqsubseteq> x" | 
| 192 | by (rule UU_least [THEN spec]) | |
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 193 | |
| 16739 | 194 | lemma UU_reorient: "(\<bottom> = x) = (x = \<bottom>)" | 
| 195 | by auto | |
| 196 | ||
| 197 | ML_setup {*
 | |
| 198 | local | |
| 199 | val meta_UU_reorient = thm "UU_reorient" RS eq_reflection; | |
| 200 |   fun is_UU (Const ("Pcpo.UU",_)) = true
 | |
| 201 | | is_UU _ = false; | |
| 202 | fun reorient_proc sg _ (_ $ t $ u) = | |
| 203 | if is_UU u then NONE else SOME meta_UU_reorient; | |
| 204 | in | |
| 205 | val UU_reorient_simproc = | |
| 206 | Simplifier.simproc (the_context ()) | |
| 207 | "UU_reorient_simproc" ["UU=x"] reorient_proc | |
| 208 | end; | |
| 209 | ||
| 210 | Addsimprocs [UU_reorient_simproc]; | |
| 211 | *} | |
| 212 | ||
| 16626 | 213 | text {* useful lemmas about @{term \<bottom>} *}
 | 
| 15563 | 214 | |
| 16626 | 215 | lemma eq_UU_iff: "(x = \<bottom>) = (x \<sqsubseteq> \<bottom>)" | 
| 17813 | 216 | by (simp add: po_eq_conv) | 
| 15563 | 217 | |
| 16626 | 218 | lemma UU_I: "x \<sqsubseteq> \<bottom> \<Longrightarrow> x = \<bottom>" | 
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 219 | by (subst eq_UU_iff) | 
| 15563 | 220 | |
| 16626 | 221 | lemma not_less2not_eq: "\<not> (x::'a::po) \<sqsubseteq> y \<Longrightarrow> x \<noteq> y" | 
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 222 | by auto | 
| 15563 | 223 | |
| 16626 | 224 | lemma chain_UU_I: "\<lbrakk>chain Y; (\<Squnion>i. Y i) = \<bottom>\<rbrakk> \<Longrightarrow> \<forall>i. Y i = \<bottom>" | 
| 15563 | 225 | apply (rule allI) | 
| 16626 | 226 | apply (rule UU_I) | 
| 15563 | 227 | apply (erule subst) | 
| 228 | apply (erule is_ub_thelub) | |
| 229 | done | |
| 230 | ||
| 16626 | 231 | lemma chain_UU_I_inverse: "\<forall>i::nat. Y i = \<bottom> \<Longrightarrow> (\<Squnion>i. Y i) = \<bottom>" | 
| 15563 | 232 | apply (rule lub_chain_maxelem) | 
| 233 | apply (erule spec) | |
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 234 | apply simp | 
| 15563 | 235 | done | 
| 236 | ||
| 16626 | 237 | 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: 
15577diff
changeset | 238 | by (blast intro: chain_UU_I_inverse) | 
| 15563 | 239 | |
| 16626 | 240 | 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: 
15577diff
changeset | 241 | by (blast intro: UU_I) | 
| 15563 | 242 | |
| 16627 | 243 | lemma chain_mono2: "\<lbrakk>\<exists>j. Y j \<noteq> \<bottom>; chain Y\<rbrakk> \<Longrightarrow> \<exists>j. \<forall>i>j. Y i \<noteq> \<bottom>" | 
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 244 | by (blast dest: notUU_I chain_mono) | 
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 245 | |
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 246 | subsection {* Chain-finite and flat cpos *}
 | 
| 15563 | 247 | |
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 248 | text {* further useful classes for HOLCF domains *}
 | 
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 249 | |
| 15640 
2d1d6ea579a1
chfin now a subclass of po, proved instance chfin < cpo
 huffman parents: 
15588diff
changeset | 250 | axclass chfin < po | 
| 16626 | 251 | chfin: "\<forall>Y. chain Y \<longrightarrow> (\<exists>n. max_in_chain n Y)" | 
| 15563 | 252 | |
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 253 | axclass flat < pcpo | 
| 16626 | 254 | ax_flat: "\<forall>x y. x \<sqsubseteq> y \<longrightarrow> (x = \<bottom>) \<or> (x = y)" | 
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 255 | |
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 256 | text {* some properties for chfin and flat *}
 | 
| 15563 | 257 | |
| 15640 
2d1d6ea579a1
chfin now a subclass of po, proved instance chfin < cpo
 huffman parents: 
15588diff
changeset | 258 | text {* chfin types are cpo *}
 | 
| 
2d1d6ea579a1
chfin now a subclass of po, proved instance chfin < cpo
 huffman parents: 
15588diff
changeset | 259 | |
| 
2d1d6ea579a1
chfin now a subclass of po, proved instance chfin < cpo
 huffman parents: 
15588diff
changeset | 260 | lemma chfin_imp_cpo: | 
| 16626 | 261 | "chain (S::nat \<Rightarrow> 'a::chfin) \<Longrightarrow> \<exists>x. range S <<| x" | 
| 15640 
2d1d6ea579a1
chfin now a subclass of po, proved instance chfin < cpo
 huffman parents: 
15588diff
changeset | 262 | apply (frule chfin [rule_format]) | 
| 
2d1d6ea579a1
chfin now a subclass of po, proved instance chfin < cpo
 huffman parents: 
15588diff
changeset | 263 | apply (blast intro: lub_finch1) | 
| 
2d1d6ea579a1
chfin now a subclass of po, proved instance chfin < cpo
 huffman parents: 
15588diff
changeset | 264 | done | 
| 
2d1d6ea579a1
chfin now a subclass of po, proved instance chfin < cpo
 huffman parents: 
15588diff
changeset | 265 | |
| 
2d1d6ea579a1
chfin now a subclass of po, proved instance chfin < cpo
 huffman parents: 
15588diff
changeset | 266 | instance chfin < cpo | 
| 
2d1d6ea579a1
chfin now a subclass of po, proved instance chfin < cpo
 huffman parents: 
15588diff
changeset | 267 | by intro_classes (rule chfin_imp_cpo) | 
| 
2d1d6ea579a1
chfin now a subclass of po, proved instance chfin < cpo
 huffman parents: 
15588diff
changeset | 268 | |
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 269 | text {* flat types are chfin *}
 | 
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 270 | |
| 15563 | 271 | lemma flat_imp_chfin: | 
| 16626 | 272 | "\<forall>Y::nat \<Rightarrow> 'a::flat. chain Y \<longrightarrow> (\<exists>n. max_in_chain n Y)" | 
| 15563 | 273 | apply (unfold max_in_chain_def) | 
| 274 | apply clarify | |
| 16626 | 275 | apply (case_tac "\<forall>i. Y i = \<bottom>") | 
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 276 | apply simp | 
| 15563 | 277 | apply simp | 
| 278 | apply (erule exE) | |
| 16626 | 279 | apply (rule_tac x="i" in exI) | 
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 280 | apply clarify | 
| 17813 | 281 | apply (blast dest: chain_mono3 ax_flat [rule_format]) | 
| 15563 | 282 | done | 
| 283 | ||
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 284 | instance flat < chfin | 
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 285 | by intro_classes (rule flat_imp_chfin) | 
| 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 286 | |
| 16627 | 287 | text {* flat subclass of chfin; @{text adm_flat} not needed *}
 | 
| 15563 | 288 | |
| 16626 | 289 | lemma flat_eq: "(a::'a::flat) \<noteq> \<bottom> \<Longrightarrow> a \<sqsubseteq> b = (a = b)" | 
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 290 | by (safe dest!: ax_flat [rule_format]) | 
| 15563 | 291 | |
| 16626 | 292 | lemma chfin2finch: "chain (Y::nat \<Rightarrow> 'a::chfin) \<Longrightarrow> finite_chain Y" | 
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 293 | by (simp add: chfin finite_chain_def) | 
| 15563 | 294 | |
| 15588 
14e3228f18cc
arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 295 | text {* lemmata for improved admissibility introdution rule *}
 | 
| 15563 | 296 | |
| 297 | lemma infinite_chain_adm_lemma: | |
| 16626 | 298 | "\<lbrakk>chain Y; \<forall>i. P (Y i); | 
| 299 | \<And>Y. \<lbrakk>chain Y; \<forall>i. P (Y i); \<not> finite_chain Y\<rbrakk> \<Longrightarrow> P (\<Squnion>i. Y i)\<rbrakk> | |
| 300 | \<Longrightarrow> P (\<Squnion>i. Y i)" | |
| 15563 | 301 | apply (case_tac "finite_chain Y") | 
| 302 | prefer 2 apply fast | |
| 303 | apply (unfold finite_chain_def) | |
| 304 | apply safe | |
| 305 | apply (erule lub_finch1 [THEN thelubI, THEN ssubst]) | |
| 306 | apply assumption | |
| 307 | apply (erule spec) | |
| 308 | done | |
| 309 | ||
| 310 | lemma increasing_chain_adm_lemma: | |
| 16626 | 311 | "\<lbrakk>chain Y; \<forall>i. P (Y i); \<And>Y. \<lbrakk>chain Y; \<forall>i. P (Y i); | 
| 312 | \<forall>i. \<exists>j>i. Y i \<noteq> Y j \<and> Y i \<sqsubseteq> Y j\<rbrakk> \<Longrightarrow> P (\<Squnion>i. Y i)\<rbrakk> | |
| 313 | \<Longrightarrow> P (\<Squnion>i. Y i)" | |
| 15563 | 314 | apply (erule infinite_chain_adm_lemma) | 
| 315 | apply assumption | |
| 316 | apply (erule thin_rl) | |
| 317 | apply (unfold finite_chain_def) | |
| 318 | apply (unfold max_in_chain_def) | |
| 319 | apply (fast dest: le_imp_less_or_eq elim: chain_mono) | |
| 320 | done | |
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: 
15563diff
changeset | 321 | |
| 16626 | 322 | end |