| author | wenzelm | 
| Fri, 17 Jan 2025 11:49:31 +0100 | |
| changeset 81845 | acd9849d4e9e | 
| parent 81583 | b6df83045178 | 
| child 82621 | b444e7150e1f | 
| permissions | -rw-r--r-- | 
| 81575 | 1 | (* Title: HOL/HOLCF/Cpo.thy | 
| 2 | Author: Franz Regensburger | |
| 3 | Author: Tobias Nipkow | |
| 4 | Author: Brian Huffman | |
| 5 | ||
| 6 | Foundations of HOLCF: complete partial orders etc. | |
| 7 | *) | |
| 8 | ||
| 9 | theory Cpo | |
| 10 | imports Main | |
| 11 | begin | |
| 12 | ||
| 13 | section \<open>Partial orders\<close> | |
| 14 | ||
| 15 | declare [[typedef_overloaded]] | |
| 16 | ||
| 17 | ||
| 18 | subsection \<open>Type class for partial orders\<close> | |
| 19 | ||
| 20 | class below = | |
| 21 | fixes below :: "'a \<Rightarrow> 'a \<Rightarrow> bool" | |
| 22 | begin | |
| 23 | ||
| 24 | notation (ASCII) | |
| 25 | below (infix \<open><<\<close> 50) | |
| 26 | ||
| 27 | notation | |
| 28 | below (infix \<open>\<sqsubseteq>\<close> 50) | |
| 29 | ||
| 30 | abbreviation not_below :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix \<open>\<notsqsubseteq>\<close> 50) | |
| 31 | where "not_below x y \<equiv> \<not> below x y" | |
| 32 | ||
| 33 | notation (ASCII) | |
| 34 | not_below (infix \<open>~<<\<close> 50) | |
| 35 | ||
| 36 | lemma below_eq_trans: "a \<sqsubseteq> b \<Longrightarrow> b = c \<Longrightarrow> a \<sqsubseteq> c" | |
| 37 | by (rule subst) | |
| 38 | ||
| 39 | lemma eq_below_trans: "a = b \<Longrightarrow> b \<sqsubseteq> c \<Longrightarrow> a \<sqsubseteq> c" | |
| 40 | by (rule ssubst) | |
| 41 | ||
| 42 | end | |
| 43 | ||
| 44 | class po = below + | |
| 45 | assumes below_refl [iff]: "x \<sqsubseteq> x" | |
| 46 | assumes below_trans: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> z" | |
| 47 | assumes below_antisym: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> x \<Longrightarrow> x = y" | |
| 48 | begin | |
| 49 | ||
| 50 | lemma eq_imp_below: "x = y \<Longrightarrow> x \<sqsubseteq> y" | |
| 51 | by simp | |
| 52 | ||
| 53 | lemma box_below: "a \<sqsubseteq> b \<Longrightarrow> c \<sqsubseteq> a \<Longrightarrow> b \<sqsubseteq> d \<Longrightarrow> c \<sqsubseteq> d" | |
| 54 | by (rule below_trans [OF below_trans]) | |
| 55 | ||
| 56 | lemma po_eq_conv: "x = y \<longleftrightarrow> x \<sqsubseteq> y \<and> y \<sqsubseteq> x" | |
| 57 | by (fast intro!: below_antisym) | |
| 58 | ||
| 59 | lemma rev_below_trans: "y \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> y \<Longrightarrow> x \<sqsubseteq> z" | |
| 60 | by (rule below_trans) | |
| 61 | ||
| 62 | lemma not_below2not_eq: "x \<notsqsubseteq> y \<Longrightarrow> x \<noteq> y" | |
| 63 | by auto | |
| 64 | ||
| 65 | end | |
| 66 | ||
| 67 | lemmas HOLCF_trans_rules [trans] = | |
| 68 | below_trans | |
| 69 | below_antisym | |
| 70 | below_eq_trans | |
| 71 | eq_below_trans | |
| 72 | ||
| 73 | context po | |
| 74 | begin | |
| 75 | ||
| 76 | subsection \<open>Upper bounds\<close> | |
| 77 | ||
| 78 | definition is_ub :: "'a set \<Rightarrow> 'a \<Rightarrow> bool" (infix \<open><|\<close> 55) | |
| 79 | where "S <| x \<longleftrightarrow> (\<forall>y\<in>S. y \<sqsubseteq> x)" | |
| 80 | ||
| 81 | lemma is_ubI: "(\<And>x. x \<in> S \<Longrightarrow> x \<sqsubseteq> u) \<Longrightarrow> S <| u" | |
| 82 | by (simp add: is_ub_def) | |
| 83 | ||
| 84 | lemma is_ubD: "\<lbrakk>S <| u; x \<in> S\<rbrakk> \<Longrightarrow> x \<sqsubseteq> u" | |
| 85 | by (simp add: is_ub_def) | |
| 86 | ||
| 87 | lemma ub_imageI: "(\<And>x. x \<in> S \<Longrightarrow> f x \<sqsubseteq> u) \<Longrightarrow> (\<lambda>x. f x) ` S <| u" | |
| 88 | unfolding is_ub_def by fast | |
| 89 | ||
| 90 | lemma ub_imageD: "\<lbrakk>f ` S <| u; x \<in> S\<rbrakk> \<Longrightarrow> f x \<sqsubseteq> u" | |
| 91 | unfolding is_ub_def by fast | |
| 92 | ||
| 93 | lemma ub_rangeI: "(\<And>i. S i \<sqsubseteq> x) \<Longrightarrow> range S <| x" | |
| 94 | unfolding is_ub_def by fast | |
| 95 | ||
| 96 | lemma ub_rangeD: "range S <| x \<Longrightarrow> S i \<sqsubseteq> x" | |
| 97 | unfolding is_ub_def by fast | |
| 98 | ||
| 99 | lemma is_ub_empty [simp]: "{} <| u"
 | |
| 100 | unfolding is_ub_def by fast | |
| 101 | ||
| 102 | lemma is_ub_insert [simp]: "(insert x A) <| y = (x \<sqsubseteq> y \<and> A <| y)" | |
| 103 | unfolding is_ub_def by fast | |
| 104 | ||
| 105 | lemma is_ub_upward: "\<lbrakk>S <| x; x \<sqsubseteq> y\<rbrakk> \<Longrightarrow> S <| y" | |
| 106 | unfolding is_ub_def by (fast intro: below_trans) | |
| 107 | ||
| 108 | ||
| 109 | subsection \<open>Least upper bounds\<close> | |
| 110 | ||
| 111 | definition is_lub :: "'a set \<Rightarrow> 'a \<Rightarrow> bool" (infix \<open><<|\<close> 55) | |
| 112 | where "S <<| x \<longleftrightarrow> S <| x \<and> (\<forall>u. S <| u \<longrightarrow> x \<sqsubseteq> u)" | |
| 113 | ||
| 114 | definition lub :: "'a set \<Rightarrow> 'a" | |
| 115 | where "lub S = (THE x. S <<| x)" | |
| 116 | ||
| 117 | end | |
| 118 | ||
| 119 | syntax (ASCII) | |
| 120 | "_BLub" :: "[pttrn, 'a set, 'b] \<Rightarrow> 'b" (\<open>(\<open>indent=3 notation=\<open>binder LUB\<close>\<close>LUB _:_./ _)\<close> [0,0, 10] 10) | |
| 121 | ||
| 122 | syntax | |
| 123 | "_BLub" :: "[pttrn, 'a set, 'b] \<Rightarrow> 'b" (\<open>(\<open>indent=3 notation=\<open>binder \<Squnion>\<close>\<close>\<Squnion>_\<in>_./ _)\<close> [0,0, 10] 10) | |
| 124 | ||
| 125 | syntax_consts | |
| 126 | "_BLub" \<rightleftharpoons> lub | |
| 127 | ||
| 128 | translations | |
| 129 | "LUB x:A. t" \<rightleftharpoons> "CONST lub ((\<lambda>x. t) ` A)" | |
| 130 | ||
| 131 | context po | |
| 132 | begin | |
| 133 | ||
| 134 | abbreviation Lub (binder \<open>\<Squnion>\<close> 10) | |
| 135 | where "\<Squnion>n. t n \<equiv> lub (range t)" | |
| 136 | ||
| 137 | notation (ASCII) | |
| 138 | Lub (binder \<open>LUB \<close> 10) | |
| 139 | ||
| 140 | text \<open>access to some definition as inference rule\<close> | |
| 141 | ||
| 142 | lemma is_lubD1: "S <<| x \<Longrightarrow> S <| x" | |
| 143 | unfolding is_lub_def by fast | |
| 144 | ||
| 145 | lemma is_lubD2: "\<lbrakk>S <<| x; S <| u\<rbrakk> \<Longrightarrow> x \<sqsubseteq> u" | |
| 146 | unfolding is_lub_def by fast | |
| 147 | ||
| 148 | lemma is_lubI: "\<lbrakk>S <| x; \<And>u. S <| u \<Longrightarrow> x \<sqsubseteq> u\<rbrakk> \<Longrightarrow> S <<| x" | |
| 149 | unfolding is_lub_def by fast | |
| 150 | ||
| 151 | lemma is_lub_below_iff: "S <<| x \<Longrightarrow> x \<sqsubseteq> u \<longleftrightarrow> S <| u" | |
| 152 | unfolding is_lub_def is_ub_def by (metis below_trans) | |
| 153 | ||
| 154 | text \<open>lubs are unique\<close> | |
| 155 | ||
| 156 | lemma is_lub_unique: "S <<| x \<Longrightarrow> S <<| y \<Longrightarrow> x = y" | |
| 157 | unfolding is_lub_def is_ub_def by (blast intro: below_antisym) | |
| 158 | ||
| 159 | text \<open>technical lemmas about \<^term>\<open>lub\<close> and \<^term>\<open>is_lub\<close>\<close> | |
| 160 | ||
| 161 | lemma is_lub_lub: "M <<| x \<Longrightarrow> M <<| lub M" | |
| 162 | unfolding lub_def by (rule theI [OF _ is_lub_unique]) | |
| 163 | ||
| 164 | lemma lub_eqI: "M <<| l \<Longrightarrow> lub M = l" | |
| 165 | by (rule is_lub_unique [OF is_lub_lub]) | |
| 166 | ||
| 167 | lemma is_lub_singleton [simp]: "{x} <<| x"
 | |
| 168 | by (simp add: is_lub_def) | |
| 169 | ||
| 170 | lemma lub_singleton [simp]: "lub {x} = x"
 | |
| 171 | by (rule is_lub_singleton [THEN lub_eqI]) | |
| 172 | ||
| 173 | lemma is_lub_bin: "x \<sqsubseteq> y \<Longrightarrow> {x, y} <<| y"
 | |
| 174 | by (simp add: is_lub_def) | |
| 175 | ||
| 176 | lemma lub_bin: "x \<sqsubseteq> y \<Longrightarrow> lub {x, y} = y"
 | |
| 177 | by (rule is_lub_bin [THEN lub_eqI]) | |
| 178 | ||
| 179 | lemma is_lub_maximal: "S <| x \<Longrightarrow> x \<in> S \<Longrightarrow> S <<| x" | |
| 180 | by (erule is_lubI, erule (1) is_ubD) | |
| 181 | ||
| 182 | lemma lub_maximal: "S <| x \<Longrightarrow> x \<in> S \<Longrightarrow> lub S = x" | |
| 183 | by (rule is_lub_maximal [THEN lub_eqI]) | |
| 184 | ||
| 185 | ||
| 186 | subsection \<open>Countable chains\<close> | |
| 187 | ||
| 188 | definition chain :: "(nat \<Rightarrow> 'a) \<Rightarrow> bool" | |
| 189 | where \<comment> \<open>Here we use countable chains and I prefer to code them as functions!\<close> | |
| 190 | "chain Y = (\<forall>i. Y i \<sqsubseteq> Y (Suc i))" | |
| 191 | ||
| 192 | lemma chainI: "(\<And>i. Y i \<sqsubseteq> Y (Suc i)) \<Longrightarrow> chain Y" | |
| 193 | unfolding chain_def by fast | |
| 194 | ||
| 195 | lemma chainE: "chain Y \<Longrightarrow> Y i \<sqsubseteq> Y (Suc i)" | |
| 196 | unfolding chain_def by fast | |
| 197 | ||
| 198 | text \<open>chains are monotone functions\<close> | |
| 199 | ||
| 200 | lemma chain_mono_less: "chain Y \<Longrightarrow> i < j \<Longrightarrow> Y i \<sqsubseteq> Y j" | |
| 201 | by (erule less_Suc_induct, erule chainE, erule below_trans) | |
| 202 | ||
| 203 | lemma chain_mono: "chain Y \<Longrightarrow> i \<le> j \<Longrightarrow> Y i \<sqsubseteq> Y j" | |
| 204 | by (cases "i = j") (simp_all add: chain_mono_less) | |
| 205 | ||
| 206 | lemma chain_shift: "chain Y \<Longrightarrow> chain (\<lambda>i. Y (i + j))" | |
| 207 | by (rule chainI, simp, erule chainE) | |
| 208 | ||
| 209 | text \<open>technical lemmas about (least) upper bounds of chains\<close> | |
| 210 | ||
| 211 | lemma is_lub_rangeD1: "range S <<| x \<Longrightarrow> S i \<sqsubseteq> x" | |
| 212 | by (rule is_lubD1 [THEN ub_rangeD]) | |
| 213 | ||
| 214 | lemma is_ub_range_shift: "chain S \<Longrightarrow> range (\<lambda>i. S (i + j)) <| x = range S <| x" | |
| 215 | apply (rule iffI) | |
| 216 | apply (rule ub_rangeI) | |
| 217 | apply (rule_tac y="S (i + j)" in below_trans) | |
| 218 | apply (erule chain_mono) | |
| 219 | apply (rule le_add1) | |
| 220 | apply (erule ub_rangeD) | |
| 221 | apply (rule ub_rangeI) | |
| 222 | apply (erule ub_rangeD) | |
| 223 | done | |
| 224 | ||
| 225 | lemma is_lub_range_shift: "chain S \<Longrightarrow> range (\<lambda>i. S (i + j)) <<| x = range S <<| x" | |
| 226 | by (simp add: is_lub_def is_ub_range_shift) | |
| 227 | ||
| 228 | text \<open>the lub of a constant chain is the constant\<close> | |
| 229 | ||
| 230 | lemma chain_const [simp]: "chain (\<lambda>i. c)" | |
| 231 | by (simp add: chainI) | |
| 232 | ||
| 233 | lemma is_lub_const: "range (\<lambda>x. c) <<| c" | |
| 234 | by (blast dest: ub_rangeD intro: is_lubI ub_rangeI) | |
| 235 | ||
| 236 | lemma lub_const [simp]: "(\<Squnion>i. c) = c" | |
| 237 | by (rule is_lub_const [THEN lub_eqI]) | |
| 238 | ||
| 239 | ||
| 240 | subsection \<open>Finite chains\<close> | |
| 241 | ||
| 242 | definition max_in_chain :: "nat \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> bool" | |
| 243 | where \<comment> \<open>finite chains, needed for monotony of continuous functions\<close> | |
| 244 | "max_in_chain i C \<longleftrightarrow> (\<forall>j. i \<le> j \<longrightarrow> C i = C j)" | |
| 245 | ||
| 246 | definition finite_chain :: "(nat \<Rightarrow> 'a) \<Rightarrow> bool" | |
| 247 | where "finite_chain C = (chain C \<and> (\<exists>i. max_in_chain i C))" | |
| 248 | ||
| 249 | text \<open>results about finite chains\<close> | |
| 250 | ||
| 251 | lemma max_in_chainI: "(\<And>j. i \<le> j \<Longrightarrow> Y i = Y j) \<Longrightarrow> max_in_chain i Y" | |
| 252 | unfolding max_in_chain_def by fast | |
| 253 | ||
| 254 | lemma max_in_chainD: "max_in_chain i Y \<Longrightarrow> i \<le> j \<Longrightarrow> Y i = Y j" | |
| 255 | unfolding max_in_chain_def by fast | |
| 256 | ||
| 257 | lemma finite_chainI: "chain C \<Longrightarrow> max_in_chain i C \<Longrightarrow> finite_chain C" | |
| 258 | unfolding finite_chain_def by fast | |
| 259 | ||
| 260 | lemma finite_chainE: "\<lbrakk>finite_chain C; \<And>i. \<lbrakk>chain C; max_in_chain i C\<rbrakk> \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R" | |
| 261 | unfolding finite_chain_def by fast | |
| 262 | ||
| 263 | lemma lub_finch1: "chain C \<Longrightarrow> max_in_chain i C \<Longrightarrow> range C <<| C i" | |
| 264 | apply (rule is_lubI) | |
| 265 | apply (rule ub_rangeI, rename_tac j) | |
| 266 | apply (rule_tac x=i and y=j in linorder_le_cases) | |
| 267 | apply (drule (1) max_in_chainD, simp) | |
| 268 | apply (erule (1) chain_mono) | |
| 269 | apply (erule ub_rangeD) | |
| 270 | done | |
| 271 | ||
| 272 | lemma lub_finch2: "finite_chain C \<Longrightarrow> range C <<| C (LEAST i. max_in_chain i C)" | |
| 273 | apply (erule finite_chainE) | |
| 274 | apply (erule LeastI2 [where Q="\<lambda>i. range C <<| C i"]) | |
| 275 | apply (erule (1) lub_finch1) | |
| 276 | done | |
| 277 | ||
| 278 | lemma finch_imp_finite_range: "finite_chain Y \<Longrightarrow> finite (range Y)" | |
| 279 | apply (erule finite_chainE) | |
| 280 |   apply (rule_tac B="Y ` {..i}" in finite_subset)
 | |
| 281 | apply (rule subsetI) | |
| 282 | apply (erule rangeE, rename_tac j) | |
| 283 | apply (rule_tac x=i and y=j in linorder_le_cases) | |
| 284 | apply (subgoal_tac "Y j = Y i", simp) | |
| 285 | apply (simp add: max_in_chain_def) | |
| 286 | apply simp | |
| 287 | apply simp | |
| 288 | done | |
| 289 | ||
| 290 | lemma finite_range_has_max: | |
| 291 | fixes f :: "nat \<Rightarrow> 'a" | |
| 292 | and r :: "'a \<Rightarrow> 'a \<Rightarrow> bool" | |
| 293 | assumes mono: "\<And>i j. i \<le> j \<Longrightarrow> r (f i) (f j)" | |
| 294 | assumes finite_range: "finite (range f)" | |
| 295 | shows "\<exists>k. \<forall>i. r (f i) (f k)" | |
| 296 | proof (intro exI allI) | |
| 297 | fix i :: nat | |
| 298 | let ?j = "LEAST k. f k = f i" | |
| 299 | let ?k = "Max ((\<lambda>x. LEAST k. f k = x) ` range f)" | |
| 300 | have "?j \<le> ?k" | |
| 301 | proof (rule Max_ge) | |
| 302 | show "finite ((\<lambda>x. LEAST k. f k = x) ` range f)" | |
| 303 | using finite_range by (rule finite_imageI) | |
| 304 | show "?j \<in> (\<lambda>x. LEAST k. f k = x) ` range f" | |
| 305 | by (intro imageI rangeI) | |
| 306 | qed | |
| 307 | hence "r (f ?j) (f ?k)" | |
| 308 | by (rule mono) | |
| 309 | also have "f ?j = f i" | |
| 310 | by (rule LeastI, rule refl) | |
| 311 | finally show "r (f i) (f ?k)" . | |
| 312 | qed | |
| 313 | ||
| 314 | lemma finite_range_imp_finch: "chain Y \<Longrightarrow> finite (range Y) \<Longrightarrow> finite_chain Y" | |
| 315 | apply (subgoal_tac "\<exists>k. \<forall>i. Y i \<sqsubseteq> Y k") | |
| 316 | apply (erule exE) | |
| 317 | apply (rule finite_chainI, assumption) | |
| 318 | apply (rule max_in_chainI) | |
| 319 | apply (rule below_antisym) | |
| 320 | apply (erule (1) chain_mono) | |
| 321 | apply (erule spec) | |
| 322 | apply (rule finite_range_has_max) | |
| 323 | apply (erule (1) chain_mono) | |
| 324 | apply assumption | |
| 325 | done | |
| 326 | ||
| 327 | lemma bin_chain: "x \<sqsubseteq> y \<Longrightarrow> chain (\<lambda>i. if i=0 then x else y)" | |
| 328 | by (rule chainI) simp | |
| 329 | ||
| 330 | lemma bin_chainmax: "x \<sqsubseteq> y \<Longrightarrow> max_in_chain (Suc 0) (\<lambda>i. if i=0 then x else y)" | |
| 331 | by (simp add: max_in_chain_def) | |
| 332 | ||
| 333 | lemma is_lub_bin_chain: "x \<sqsubseteq> y \<Longrightarrow> range (\<lambda>i::nat. if i=0 then x else y) <<| y" | |
| 334 | apply (frule bin_chain) | |
| 335 | apply (drule bin_chainmax) | |
| 336 | apply (drule (1) lub_finch1) | |
| 337 | apply simp | |
| 338 | done | |
| 339 | ||
| 340 | text \<open>the maximal element in a chain is its lub\<close> | |
| 341 | ||
| 342 | lemma lub_chain_maxelem: "Y i = c \<Longrightarrow> \<forall>i. Y i \<sqsubseteq> c \<Longrightarrow> lub (range Y) = c" | |
| 343 | by (blast dest: ub_rangeD intro: lub_eqI is_lubI ub_rangeI) | |
| 344 | ||
| 345 | end | |
| 346 | ||
| 347 | ||
| 348 | section \<open>Classes cpo and pcpo\<close> | |
| 349 | ||
| 350 | subsection \<open>Complete partial orders\<close> | |
| 351 | ||
| 352 | text \<open>The class cpo of chain complete partial orders\<close> | |
| 353 | ||
| 354 | class cpo = po + | |
| 355 | assumes cpo: "chain S \<Longrightarrow> \<exists>x. range S <<| x" | |
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 356 | |
| 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 357 | default_sort cpo | 
| 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 358 | |
| 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 359 | context cpo | 
| 81575 | 360 | begin | 
| 361 | ||
| 362 | text \<open>in cpo's everthing equal to THE lub has lub properties for every chain\<close> | |
| 363 | ||
| 364 | lemma cpo_lubI: "chain S \<Longrightarrow> range S <<| (\<Squnion>i. S i)" | |
| 365 | by (fast dest: cpo elim: is_lub_lub) | |
| 366 | ||
| 367 | lemma thelubE: "\<lbrakk>chain S; (\<Squnion>i. S i) = l\<rbrakk> \<Longrightarrow> range S <<| l" | |
| 368 | by (blast dest: cpo intro: is_lub_lub) | |
| 369 | ||
| 370 | text \<open>Properties of the lub\<close> | |
| 371 | ||
| 372 | lemma is_ub_thelub: "chain S \<Longrightarrow> S x \<sqsubseteq> (\<Squnion>i. S i)" | |
| 373 | by (blast dest: cpo intro: is_lub_lub [THEN is_lub_rangeD1]) | |
| 374 | ||
| 375 | lemma is_lub_thelub: "\<lbrakk>chain S; range S <| x\<rbrakk> \<Longrightarrow> (\<Squnion>i. S i) \<sqsubseteq> x" | |
| 376 | by (blast dest: cpo intro: is_lub_lub [THEN is_lubD2]) | |
| 377 | ||
| 378 | lemma lub_below_iff: "chain S \<Longrightarrow> (\<Squnion>i. S i) \<sqsubseteq> x \<longleftrightarrow> (\<forall>i. S i \<sqsubseteq> x)" | |
| 379 | by (simp add: is_lub_below_iff [OF cpo_lubI] is_ub_def) | |
| 380 | ||
| 381 | lemma lub_below: "\<lbrakk>chain S; \<And>i. S i \<sqsubseteq> x\<rbrakk> \<Longrightarrow> (\<Squnion>i. S i) \<sqsubseteq> x" | |
| 382 | by (simp add: lub_below_iff) | |
| 383 | ||
| 384 | lemma below_lub: "\<lbrakk>chain S; x \<sqsubseteq> S i\<rbrakk> \<Longrightarrow> x \<sqsubseteq> (\<Squnion>i. S i)" | |
| 385 | by (erule below_trans, erule is_ub_thelub) | |
| 386 | ||
| 387 | 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)" | |
| 388 | apply (erule lub_below) | |
| 389 | apply (subgoal_tac "\<exists>j. X i = Y j") | |
| 390 | apply clarsimp | |
| 391 | apply (erule is_ub_thelub) | |
| 392 | apply auto | |
| 393 | done | |
| 394 | ||
| 395 | lemma lub_range_shift: "chain Y \<Longrightarrow> (\<Squnion>i. Y (i + j)) = (\<Squnion>i. Y i)" | |
| 396 | apply (rule below_antisym) | |
| 397 | apply (rule lub_range_mono) | |
| 398 | apply fast | |
| 399 | apply assumption | |
| 400 | apply (erule chain_shift) | |
| 401 | apply (rule lub_below) | |
| 402 | apply assumption | |
| 403 | apply (rule_tac i="i" in below_lub) | |
| 404 | apply (erule chain_shift) | |
| 405 | apply (erule chain_mono) | |
| 406 | apply (rule le_add1) | |
| 407 | done | |
| 408 | ||
| 409 | lemma maxinch_is_thelub: "chain Y \<Longrightarrow> max_in_chain i Y = ((\<Squnion>i. Y i) = Y i)" | |
| 410 | apply (rule iffI) | |
| 411 | apply (fast intro!: lub_eqI lub_finch1) | |
| 412 | apply (unfold max_in_chain_def) | |
| 413 | apply (safe intro!: below_antisym) | |
| 414 | apply (fast elim!: chain_mono) | |
| 415 | apply (drule sym) | |
| 416 | apply (force elim!: is_ub_thelub) | |
| 417 | done | |
| 418 | ||
| 419 | text \<open>the \<open>\<sqsubseteq>\<close> relation between two chains is preserved by their lubs\<close> | |
| 420 | ||
| 421 | 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)" | |
| 422 | by (fast elim: lub_below below_lub) | |
| 423 | ||
| 424 | text \<open>the = relation between two chains is preserved by their lubs\<close> | |
| 425 | ||
| 426 | lemma lub_eq: "(\<And>i. X i = Y i) \<Longrightarrow> (\<Squnion>i. X i) = (\<Squnion>i. Y i)" | |
| 427 | by simp | |
| 428 | ||
| 429 | lemma ch2ch_lub: | |
| 430 | assumes 1: "\<And>j. chain (\<lambda>i. Y i j)" | |
| 431 | assumes 2: "\<And>i. chain (\<lambda>j. Y i j)" | |
| 432 | shows "chain (\<lambda>i. \<Squnion>j. Y i j)" | |
| 433 | apply (rule chainI) | |
| 434 | apply (rule lub_mono [OF 2 2]) | |
| 435 | apply (rule chainE [OF 1]) | |
| 436 | done | |
| 437 | ||
| 438 | lemma diag_lub: | |
| 439 | assumes 1: "\<And>j. chain (\<lambda>i. Y i j)" | |
| 440 | assumes 2: "\<And>i. chain (\<lambda>j. Y i j)" | |
| 441 | shows "(\<Squnion>i. \<Squnion>j. Y i j) = (\<Squnion>i. Y i i)" | |
| 442 | proof (rule below_antisym) | |
| 443 | have 3: "chain (\<lambda>i. Y i i)" | |
| 444 | apply (rule chainI) | |
| 445 | apply (rule below_trans) | |
| 446 | apply (rule chainE [OF 1]) | |
| 447 | apply (rule chainE [OF 2]) | |
| 448 | done | |
| 449 | have 4: "chain (\<lambda>i. \<Squnion>j. Y i j)" | |
| 450 | by (rule ch2ch_lub [OF 1 2]) | |
| 451 | show "(\<Squnion>i. \<Squnion>j. Y i j) \<sqsubseteq> (\<Squnion>i. Y i i)" | |
| 452 | apply (rule lub_below [OF 4]) | |
| 453 | apply (rule lub_below [OF 2]) | |
| 454 | apply (rule below_lub [OF 3]) | |
| 455 | apply (rule below_trans) | |
| 456 | apply (rule chain_mono [OF 1 max.cobounded1]) | |
| 457 | apply (rule chain_mono [OF 2 max.cobounded2]) | |
| 458 | done | |
| 459 | show "(\<Squnion>i. Y i i) \<sqsubseteq> (\<Squnion>i. \<Squnion>j. Y i j)" | |
| 460 | apply (rule lub_mono [OF 3 4]) | |
| 461 | apply (rule is_ub_thelub [OF 2]) | |
| 462 | done | |
| 463 | qed | |
| 464 | ||
| 465 | lemma ex_lub: | |
| 466 | assumes 1: "\<And>j. chain (\<lambda>i. Y i j)" | |
| 467 | assumes 2: "\<And>i. chain (\<lambda>j. Y i j)" | |
| 468 | shows "(\<Squnion>i. \<Squnion>j. Y i j) = (\<Squnion>j. \<Squnion>i. Y i j)" | |
| 469 | by (simp add: diag_lub 1 2) | |
| 470 | ||
| 471 | end | |
| 472 | ||
| 473 | ||
| 474 | subsection \<open>Pointed cpos\<close> | |
| 475 | ||
| 476 | text \<open>The class pcpo of pointed cpos\<close> | |
| 477 | ||
| 478 | class pcpo = cpo + | |
| 479 | assumes least: "\<exists>x. \<forall>y. x \<sqsubseteq> y" | |
| 480 | begin | |
| 481 | ||
| 482 | definition bottom :: "'a" (\<open>\<bottom>\<close>) | |
| 483 | where "bottom = (THE x. \<forall>y. x \<sqsubseteq> y)" | |
| 484 | ||
| 485 | lemma minimal [iff]: "\<bottom> \<sqsubseteq> x" | |
| 486 | unfolding bottom_def | |
| 487 | apply (rule the1I2) | |
| 488 | apply (rule ex_ex1I) | |
| 489 | apply (rule least) | |
| 490 | apply (blast intro: below_antisym) | |
| 491 | apply simp | |
| 492 | done | |
| 493 | ||
| 494 | end | |
| 495 | ||
| 496 | text \<open>Old "UU" syntax:\<close> | |
| 497 | abbreviation (input) "UU \<equiv> bottom" | |
| 498 | ||
| 499 | text \<open>Simproc to rewrite \<^term>\<open>\<bottom> = x\<close> to \<^term>\<open>x = \<bottom>\<close>.\<close> | |
| 500 | setup \<open>Reorient_Proc.add (fn \<^Const_>\<open>bottom _\<close> => true | _ => false)\<close> | |
| 501 | simproc_setup reorient_bottom ("\<bottom> = x") = \<open>K Reorient_Proc.proc\<close>
 | |
| 502 | ||
| 503 | text \<open>useful lemmas about \<^term>\<open>\<bottom>\<close>\<close> | |
| 504 | ||
| 505 | lemma below_bottom_iff [simp]: "x \<sqsubseteq> \<bottom> \<longleftrightarrow> x = \<bottom>" | |
| 506 | by (simp add: po_eq_conv) | |
| 507 | ||
| 508 | lemma eq_bottom_iff: "x = \<bottom> \<longleftrightarrow> x \<sqsubseteq> \<bottom>" | |
| 509 | by simp | |
| 510 | ||
| 511 | lemma bottomI: "x \<sqsubseteq> \<bottom> \<Longrightarrow> x = \<bottom>" | |
| 512 | by (subst eq_bottom_iff) | |
| 513 | ||
| 514 | lemma lub_eq_bottom_iff: "chain Y \<Longrightarrow> (\<Squnion>i. Y i) = \<bottom> \<longleftrightarrow> (\<forall>i. Y i = \<bottom>)" | |
| 515 | by (simp only: eq_bottom_iff lub_below_iff) | |
| 516 | ||
| 517 | ||
| 518 | subsection \<open>Chain-finite and flat cpos\<close> | |
| 519 | ||
| 520 | text \<open>further useful classes for HOLCF domains\<close> | |
| 521 | ||
| 522 | class chfin = po + | |
| 523 | assumes chfin: "chain Y \<Longrightarrow> \<exists>n. max_in_chain n Y" | |
| 524 | begin | |
| 525 | ||
| 526 | subclass cpo | |
| 527 | apply standard | |
| 528 | apply (frule chfin) | |
| 529 | apply (blast intro: lub_finch1) | |
| 530 | done | |
| 531 | ||
| 532 | lemma chfin2finch: "chain Y \<Longrightarrow> finite_chain Y" | |
| 533 | by (simp add: chfin finite_chain_def) | |
| 534 | ||
| 535 | end | |
| 536 | ||
| 537 | class flat = pcpo + | |
| 538 | assumes ax_flat: "x \<sqsubseteq> y \<Longrightarrow> x = \<bottom> \<or> x = y" | |
| 539 | begin | |
| 540 | ||
| 541 | subclass chfin | |
| 542 | proof | |
| 543 | fix Y | |
| 544 | assume *: "chain Y" | |
| 545 | show "\<exists>n. max_in_chain n Y" | |
| 546 | apply (unfold max_in_chain_def) | |
| 547 | apply (cases "\<forall>i. Y i = \<bottom>") | |
| 548 | apply simp | |
| 549 | apply simp | |
| 550 | apply (erule exE) | |
| 551 | apply (rule_tac x="i" in exI) | |
| 552 | apply clarify | |
| 553 | using * apply (blast dest: chain_mono ax_flat) | |
| 554 | done | |
| 555 | qed | |
| 556 | ||
| 557 | lemma flat_below_iff: "x \<sqsubseteq> y \<longleftrightarrow> x = \<bottom> \<or> x = y" | |
| 558 | by (safe dest!: ax_flat) | |
| 559 | ||
| 560 | lemma flat_eq: "a \<noteq> \<bottom> \<Longrightarrow> a \<sqsubseteq> b = (a = b)" | |
| 561 | by (safe dest!: ax_flat) | |
| 562 | ||
| 563 | end | |
| 564 | ||
| 81577 | 565 | |
| 81575 | 566 | subsection \<open>Discrete cpos\<close> | 
| 567 | ||
| 568 | class discrete_cpo = below + | |
| 569 | assumes discrete_cpo [simp]: "x \<sqsubseteq> y \<longleftrightarrow> x = y" | |
| 570 | begin | |
| 571 | ||
| 572 | subclass po | |
| 573 | by standard simp_all | |
| 574 | ||
| 575 | text \<open>In a discrete cpo, every chain is constant\<close> | |
| 576 | ||
| 577 | lemma discrete_chain_const: | |
| 578 | assumes S: "chain S" | |
| 579 | shows "\<exists>x. S = (\<lambda>i. x)" | |
| 580 | proof (intro exI ext) | |
| 581 | fix i :: nat | |
| 582 | from S le0 have "S 0 \<sqsubseteq> S i" by (rule chain_mono) | |
| 583 | then have "S 0 = S i" by simp | |
| 584 | then show "S i = S 0" by (rule sym) | |
| 585 | qed | |
| 586 | ||
| 587 | subclass chfin | |
| 588 | proof | |
| 589 | fix S :: "nat \<Rightarrow> 'a" | |
| 590 | assume S: "chain S" | |
| 591 | then have "\<exists>x. S = (\<lambda>i. x)" | |
| 592 | by (rule discrete_chain_const) | |
| 593 | then have "max_in_chain 0 S" | |
| 594 | by (auto simp: max_in_chain_def) | |
| 595 | then show "\<exists>i. max_in_chain i S" .. | |
| 596 | qed | |
| 597 | ||
| 598 | end | |
| 599 | ||
| 600 | ||
| 601 | section \<open>Continuity and monotonicity\<close> | |
| 602 | ||
| 603 | subsection \<open>Definitions\<close> | |
| 604 | ||
| 81576 | 605 | definition monofun :: "('a::po \<Rightarrow> 'b::po) \<Rightarrow> bool"  \<comment> \<open>monotonicity\<close>
 | 
| 81575 | 606 | where "monofun f \<longleftrightarrow> (\<forall>x y. x \<sqsubseteq> y \<longrightarrow> f x \<sqsubseteq> f y)" | 
| 607 | ||
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 608 | definition cont :: "('a \<Rightarrow> 'b) \<Rightarrow> bool"
 | 
| 81575 | 609 | where "cont f = (\<forall>Y. chain Y \<longrightarrow> range (\<lambda>i. f (Y i)) <<| f (\<Squnion>i. Y i))" | 
| 610 | ||
| 611 | lemma contI: "(\<And>Y. chain Y \<Longrightarrow> range (\<lambda>i. f (Y i)) <<| f (\<Squnion>i. Y i)) \<Longrightarrow> cont f" | |
| 612 | by (simp add: cont_def) | |
| 613 | ||
| 614 | lemma contE: "cont f \<Longrightarrow> chain Y \<Longrightarrow> range (\<lambda>i. f (Y i)) <<| f (\<Squnion>i. Y i)" | |
| 615 | by (simp add: cont_def) | |
| 616 | ||
| 617 | lemma monofunI: "(\<And>x y. x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y) \<Longrightarrow> monofun f" | |
| 618 | by (simp add: monofun_def) | |
| 619 | ||
| 620 | lemma monofunE: "monofun f \<Longrightarrow> x \<sqsubseteq> y \<Longrightarrow> f x \<sqsubseteq> f y" | |
| 621 | by (simp add: monofun_def) | |
| 622 | ||
| 623 | ||
| 624 | subsection \<open>Equivalence of alternate definition\<close> | |
| 625 | ||
| 626 | text \<open>monotone functions map chains to chains\<close> | |
| 627 | ||
| 628 | lemma ch2ch_monofun: "monofun f \<Longrightarrow> chain Y \<Longrightarrow> chain (\<lambda>i. f (Y i))" | |
| 629 | apply (rule chainI) | |
| 630 | apply (erule monofunE) | |
| 631 | apply (erule chainE) | |
| 632 | done | |
| 633 | ||
| 634 | text \<open>monotone functions map upper bound to upper bounds\<close> | |
| 635 | ||
| 636 | lemma ub2ub_monofun: "monofun f \<Longrightarrow> range Y <| u \<Longrightarrow> range (\<lambda>i. f (Y i)) <| f u" | |
| 637 | apply (rule ub_rangeI) | |
| 638 | apply (erule monofunE) | |
| 639 | apply (erule ub_rangeD) | |
| 640 | done | |
| 641 | ||
| 642 | text \<open>a lemma about binary chains\<close> | |
| 643 | ||
| 644 | lemma binchain_cont: "cont f \<Longrightarrow> x \<sqsubseteq> y \<Longrightarrow> range (\<lambda>i::nat. f (if i = 0 then x else y)) <<| f y" | |
| 645 | apply (subgoal_tac "f (\<Squnion>i::nat. if i = 0 then x else y) = f y") | |
| 646 | apply (erule subst) | |
| 647 | apply (erule contE) | |
| 648 | apply (erule bin_chain) | |
| 649 | apply (rule_tac f=f in arg_cong) | |
| 650 | apply (erule is_lub_bin_chain [THEN lub_eqI]) | |
| 651 | done | |
| 652 | ||
| 653 | text \<open>continuity implies monotonicity\<close> | |
| 654 | ||
| 655 | lemma cont2mono: "cont f \<Longrightarrow> monofun f" | |
| 656 | apply (rule monofunI) | |
| 657 | apply (drule (1) binchain_cont) | |
| 658 | apply (drule_tac i=0 in is_lub_rangeD1) | |
| 659 | apply simp | |
| 660 | done | |
| 661 | ||
| 662 | lemmas cont2monofunE = cont2mono [THEN monofunE] | |
| 663 | ||
| 664 | lemmas ch2ch_cont = cont2mono [THEN ch2ch_monofun] | |
| 665 | ||
| 666 | text \<open>continuity implies preservation of lubs\<close> | |
| 667 | ||
| 668 | lemma cont2contlubE: "cont f \<Longrightarrow> chain Y \<Longrightarrow> f (\<Squnion>i. Y i) = (\<Squnion>i. f (Y i))" | |
| 669 | apply (rule lub_eqI [symmetric]) | |
| 670 | apply (erule (1) contE) | |
| 671 | done | |
| 672 | ||
| 673 | lemma contI2: | |
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 674 | fixes f :: "'a \<Rightarrow> 'b" | 
| 81575 | 675 | assumes mono: "monofun f" | 
| 676 | assumes below: "\<And>Y. \<lbrakk>chain Y; chain (\<lambda>i. f (Y i))\<rbrakk> \<Longrightarrow> f (\<Squnion>i. Y i) \<sqsubseteq> (\<Squnion>i. f (Y i))" | |
| 677 | shows "cont f" | |
| 678 | proof (rule contI) | |
| 679 | fix Y :: "nat \<Rightarrow> 'a" | |
| 680 | assume Y: "chain Y" | |
| 681 | with mono have fY: "chain (\<lambda>i. f (Y i))" | |
| 682 | by (rule ch2ch_monofun) | |
| 683 | have "(\<Squnion>i. f (Y i)) = f (\<Squnion>i. Y i)" | |
| 684 | apply (rule below_antisym) | |
| 685 | apply (rule lub_below [OF fY]) | |
| 686 | apply (rule monofunE [OF mono]) | |
| 687 | apply (rule is_ub_thelub [OF Y]) | |
| 688 | apply (rule below [OF Y fY]) | |
| 689 | done | |
| 690 | with fY show "range (\<lambda>i. f (Y i)) <<| f (\<Squnion>i. Y i)" | |
| 691 | by (rule thelubE) | |
| 692 | qed | |
| 693 | ||
| 694 | ||
| 695 | subsection \<open>Collection of continuity rules\<close> | |
| 696 | ||
| 697 | named_theorems cont2cont "continuity intro rule" | |
| 698 | ||
| 699 | ||
| 700 | subsection \<open>Continuity of basic functions\<close> | |
| 701 | ||
| 702 | text \<open>The identity function is continuous\<close> | |
| 703 | ||
| 704 | lemma cont_id [simp, cont2cont]: "cont (\<lambda>x. x)" | |
| 705 | apply (rule contI) | |
| 706 | apply (erule cpo_lubI) | |
| 707 | done | |
| 708 | ||
| 709 | text \<open>constant functions are continuous\<close> | |
| 710 | ||
| 711 | lemma cont_const [simp, cont2cont]: "cont (\<lambda>x. c)" | |
| 712 | using is_lub_const by (rule contI) | |
| 713 | ||
| 714 | text \<open>application of functions is continuous\<close> | |
| 715 | ||
| 716 | lemma cont_apply: | |
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 717 | fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c" and t :: "'a \<Rightarrow> 'b" | 
| 81575 | 718 | assumes 1: "cont (\<lambda>x. t x)" | 
| 719 | assumes 2: "\<And>x. cont (\<lambda>y. f x y)" | |
| 720 | assumes 3: "\<And>y. cont (\<lambda>x. f x y)" | |
| 721 | shows "cont (\<lambda>x. (f x) (t x))" | |
| 722 | proof (rule contI2 [OF monofunI]) | |
| 723 | fix x y :: "'a" | |
| 724 | assume "x \<sqsubseteq> y" | |
| 725 | then show "f x (t x) \<sqsubseteq> f y (t y)" | |
| 726 | by (auto intro: cont2monofunE [OF 1] | |
| 727 | cont2monofunE [OF 2] | |
| 728 | cont2monofunE [OF 3] | |
| 729 | below_trans) | |
| 730 | next | |
| 731 | fix Y :: "nat \<Rightarrow> 'a" | |
| 732 | assume "chain Y" | |
| 733 | then show "f (\<Squnion>i. Y i) (t (\<Squnion>i. Y i)) \<sqsubseteq> (\<Squnion>i. f (Y i) (t (Y i)))" | |
| 734 | by (simp only: cont2contlubE [OF 1] ch2ch_cont [OF 1] | |
| 735 | cont2contlubE [OF 2] ch2ch_cont [OF 2] | |
| 736 | cont2contlubE [OF 3] ch2ch_cont [OF 3] | |
| 737 | diag_lub below_refl) | |
| 738 | qed | |
| 739 | ||
| 740 | lemma cont_compose: "cont c \<Longrightarrow> cont (\<lambda>x. f x) \<Longrightarrow> cont (\<lambda>x. c (f x))" | |
| 741 | by (rule cont_apply [OF _ _ cont_const]) | |
| 742 | ||
| 743 | text \<open>Least upper bounds preserve continuity\<close> | |
| 744 | ||
| 745 | lemma cont2cont_lub [simp]: | |
| 746 | assumes chain: "\<And>x. chain (\<lambda>i. F i x)" | |
| 747 | and cont: "\<And>i. cont (\<lambda>x. F i x)" | |
| 748 | shows "cont (\<lambda>x. \<Squnion>i. F i x)" | |
| 749 | apply (rule contI2) | |
| 750 | apply (simp add: monofunI cont2monofunE [OF cont] lub_mono chain) | |
| 751 | apply (simp add: cont2contlubE [OF cont]) | |
| 752 | apply (simp add: diag_lub ch2ch_cont [OF cont] chain) | |
| 753 | done | |
| 754 | ||
| 755 | text \<open>if-then-else is continuous\<close> | |
| 756 | ||
| 757 | lemma cont_if [simp, cont2cont]: "cont f \<Longrightarrow> cont g \<Longrightarrow> cont (\<lambda>x. if b then f x else g x)" | |
| 758 | by (induct b) simp_all | |
| 759 | ||
| 760 | ||
| 761 | subsection \<open>Finite chains and flat pcpos\<close> | |
| 762 | ||
| 763 | text \<open>Monotone functions map finite chains to finite chains.\<close> | |
| 764 | ||
| 765 | lemma monofun_finch2finch: "monofun f \<Longrightarrow> finite_chain Y \<Longrightarrow> finite_chain (\<lambda>n. f (Y n))" | |
| 766 | by (force simp add: finite_chain_def ch2ch_monofun max_in_chain_def) | |
| 767 | ||
| 768 | text \<open>The same holds for continuous functions.\<close> | |
| 769 | ||
| 770 | lemma cont_finch2finch: "cont f \<Longrightarrow> finite_chain Y \<Longrightarrow> finite_chain (\<lambda>n. f (Y n))" | |
| 771 | by (rule cont2mono [THEN monofun_finch2finch]) | |
| 772 | ||
| 773 | text \<open>All monotone functions with chain-finite domain are continuous.\<close> | |
| 774 | ||
| 775 | lemma chfindom_monofun2cont: "monofun f \<Longrightarrow> cont f" | |
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 776 | for f :: "'a::chfin \<Rightarrow> 'b" | 
| 81575 | 777 | apply (erule contI2) | 
| 778 | apply (frule chfin2finch) | |
| 779 | apply (clarsimp simp add: finite_chain_def) | |
| 780 | apply (subgoal_tac "max_in_chain i (\<lambda>i. f (Y i))") | |
| 781 | apply (simp add: maxinch_is_thelub ch2ch_monofun) | |
| 782 | apply (force simp add: max_in_chain_def) | |
| 783 | done | |
| 784 | ||
| 785 | text \<open>All strict functions with flat domain are continuous.\<close> | |
| 786 | ||
| 787 | lemma flatdom_strict2mono: "f \<bottom> = \<bottom> \<Longrightarrow> monofun f" | |
| 788 | for f :: "'a::flat \<Rightarrow> 'b::pcpo" | |
| 789 | apply (rule monofunI) | |
| 790 | apply (drule ax_flat) | |
| 791 | apply auto | |
| 792 | done | |
| 793 | ||
| 794 | lemma flatdom_strict2cont: "f \<bottom> = \<bottom> \<Longrightarrow> cont f" | |
| 795 | for f :: "'a::flat \<Rightarrow> 'b::pcpo" | |
| 796 | by (rule flatdom_strict2mono [THEN chfindom_monofun2cont]) | |
| 797 | ||
| 798 | text \<open>All functions with discrete domain are continuous.\<close> | |
| 799 | ||
| 800 | lemma cont_discrete_cpo [simp, cont2cont]: "cont f" | |
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 801 | for f :: "'a::discrete_cpo \<Rightarrow> 'b" | 
| 81575 | 802 | apply (rule contI) | 
| 803 | apply (drule discrete_chain_const, clarify) | |
| 804 | apply simp | |
| 805 | done | |
| 806 | ||
| 81577 | 807 | |
| 81575 | 808 | section \<open>Admissibility and compactness\<close> | 
| 809 | ||
| 810 | subsection \<open>Definitions\<close> | |
| 811 | ||
| 81582 | 812 | context cpo | 
| 813 | begin | |
| 814 | ||
| 815 | definition adm :: "('a \<Rightarrow> bool) \<Rightarrow> bool"
 | |
| 81575 | 816 | where "adm P \<longleftrightarrow> (\<forall>Y. chain Y \<longrightarrow> (\<forall>i. P (Y i)) \<longrightarrow> P (\<Squnion>i. Y i))" | 
| 817 | ||
| 818 | lemma admI: "(\<And>Y. \<lbrakk>chain Y; \<forall>i. P (Y i)\<rbrakk> \<Longrightarrow> P (\<Squnion>i. Y i)) \<Longrightarrow> adm P" | |
| 819 | unfolding adm_def by fast | |
| 820 | ||
| 821 | lemma admD: "adm P \<Longrightarrow> chain Y \<Longrightarrow> (\<And>i. P (Y i)) \<Longrightarrow> P (\<Squnion>i. Y i)" | |
| 822 | unfolding adm_def by fast | |
| 823 | ||
| 824 | lemma admD2: "adm (\<lambda>x. \<not> P x) \<Longrightarrow> chain Y \<Longrightarrow> P (\<Squnion>i. Y i) \<Longrightarrow> \<exists>i. P (Y i)" | |
| 825 | unfolding adm_def by fast | |
| 826 | ||
| 827 | lemma triv_admI: "\<forall>x. P x \<Longrightarrow> adm P" | |
| 828 | by (rule admI) (erule spec) | |
| 829 | ||
| 81582 | 830 | end | 
| 831 | ||
| 81575 | 832 | |
| 833 | subsection \<open>Admissibility on chain-finite types\<close> | |
| 834 | ||
| 835 | text \<open>For chain-finite (easy) types every formula is admissible.\<close> | |
| 836 | ||
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 837 | lemma adm_chfin [simp]: "adm P" for P :: "'a::chfin \<Rightarrow> bool" | 
| 81575 | 838 | by (rule admI, frule chfin, auto simp add: maxinch_is_thelub) | 
| 839 | ||
| 840 | ||
| 841 | subsection \<open>Admissibility of special formulae and propagation\<close> | |
| 842 | ||
| 81582 | 843 | context cpo | 
| 844 | begin | |
| 845 | ||
| 81575 | 846 | lemma adm_const [simp]: "adm (\<lambda>x. t)" | 
| 847 | by (rule admI, simp) | |
| 848 | ||
| 849 | lemma adm_conj [simp]: "adm (\<lambda>x. P x) \<Longrightarrow> adm (\<lambda>x. Q x) \<Longrightarrow> adm (\<lambda>x. P x \<and> Q x)" | |
| 850 | by (fast intro: admI elim: admD) | |
| 851 | ||
| 852 | lemma adm_all [simp]: "(\<And>y. adm (\<lambda>x. P x y)) \<Longrightarrow> adm (\<lambda>x. \<forall>y. P x y)" | |
| 853 | by (fast intro: admI elim: admD) | |
| 854 | ||
| 855 | lemma adm_ball [simp]: "(\<And>y. y \<in> A \<Longrightarrow> adm (\<lambda>x. P x y)) \<Longrightarrow> adm (\<lambda>x. \<forall>y\<in>A. P x y)" | |
| 856 | by (fast intro: admI elim: admD) | |
| 857 | ||
| 858 | text \<open>Admissibility for disjunction is hard to prove. It requires 2 lemmas.\<close> | |
| 859 | ||
| 860 | lemma adm_disj_lemma1: | |
| 861 | assumes adm: "adm P" | |
| 862 | assumes chain: "chain Y" | |
| 863 | assumes P: "\<forall>i. \<exists>j\<ge>i. P (Y j)" | |
| 864 | shows "P (\<Squnion>i. Y i)" | |
| 865 | proof - | |
| 866 | define f where "f i = (LEAST j. i \<le> j \<and> P (Y j))" for i | |
| 867 | have chain': "chain (\<lambda>i. Y (f i))" | |
| 868 | unfolding f_def | |
| 869 | apply (rule chainI) | |
| 870 | apply (rule chain_mono [OF chain]) | |
| 871 | apply (rule Least_le) | |
| 872 | apply (rule LeastI2_ex) | |
| 873 | apply (simp_all add: P) | |
| 874 | done | |
| 875 | have f1: "\<And>i. i \<le> f i" and f2: "\<And>i. P (Y (f i))" | |
| 876 | using LeastI_ex [OF P [rule_format]] by (simp_all add: f_def) | |
| 877 | have lub_eq: "(\<Squnion>i. Y i) = (\<Squnion>i. Y (f i))" | |
| 878 | apply (rule below_antisym) | |
| 879 | apply (rule lub_mono [OF chain chain']) | |
| 880 | apply (rule chain_mono [OF chain f1]) | |
| 881 | apply (rule lub_range_mono [OF _ chain chain']) | |
| 882 | apply clarsimp | |
| 883 | done | |
| 884 | show "P (\<Squnion>i. Y i)" | |
| 885 | unfolding lub_eq using adm chain' f2 by (rule admD) | |
| 886 | qed | |
| 887 | ||
| 888 | lemma adm_disj_lemma2: "\<forall>n::nat. P n \<or> Q n \<Longrightarrow> (\<forall>i. \<exists>j\<ge>i. P j) \<or> (\<forall>i. \<exists>j\<ge>i. Q j)" | |
| 889 | apply (erule contrapos_pp) | |
| 890 | apply (clarsimp, rename_tac a b) | |
| 891 | apply (rule_tac x="max a b" in exI) | |
| 892 | apply simp | |
| 893 | done | |
| 894 | ||
| 895 | lemma adm_disj [simp]: "adm (\<lambda>x. P x) \<Longrightarrow> adm (\<lambda>x. Q x) \<Longrightarrow> adm (\<lambda>x. P x \<or> Q x)" | |
| 896 | apply (rule admI) | |
| 897 | apply (erule adm_disj_lemma2 [THEN disjE]) | |
| 898 | apply (erule (2) adm_disj_lemma1 [THEN disjI1]) | |
| 899 | apply (erule (2) adm_disj_lemma1 [THEN disjI2]) | |
| 900 | done | |
| 901 | ||
| 902 | lemma adm_imp [simp]: "adm (\<lambda>x. \<not> P x) \<Longrightarrow> adm (\<lambda>x. Q x) \<Longrightarrow> adm (\<lambda>x. P x \<longrightarrow> Q x)" | |
| 903 | by (subst imp_conv_disj) (rule adm_disj) | |
| 904 | ||
| 905 | lemma adm_iff [simp]: "adm (\<lambda>x. P x \<longrightarrow> Q x) \<Longrightarrow> adm (\<lambda>x. Q x \<longrightarrow> P x) \<Longrightarrow> adm (\<lambda>x. P x \<longleftrightarrow> Q x)" | |
| 906 | by (subst iff_conv_conj_imp) (rule adm_conj) | |
| 907 | ||
| 81582 | 908 | end | 
| 909 | ||
| 81575 | 910 | text \<open>admissibility and continuity\<close> | 
| 911 | ||
| 912 | lemma adm_below [simp]: "cont (\<lambda>x. u x) \<Longrightarrow> cont (\<lambda>x. v x) \<Longrightarrow> adm (\<lambda>x. u x \<sqsubseteq> v x)" | |
| 913 | by (simp add: adm_def cont2contlubE lub_mono ch2ch_cont) | |
| 914 | ||
| 915 | lemma adm_eq [simp]: "cont (\<lambda>x. u x) \<Longrightarrow> cont (\<lambda>x. v x) \<Longrightarrow> adm (\<lambda>x. u x = v x)" | |
| 916 | by (simp add: po_eq_conv) | |
| 917 | ||
| 918 | lemma adm_subst: "cont (\<lambda>x. t x) \<Longrightarrow> adm P \<Longrightarrow> adm (\<lambda>x. P (t x))" | |
| 919 | by (simp add: adm_def cont2contlubE ch2ch_cont) | |
| 920 | ||
| 921 | lemma adm_not_below [simp]: "cont (\<lambda>x. t x) \<Longrightarrow> adm (\<lambda>x. t x \<notsqsubseteq> u)" | |
| 922 | by (rule admI) (simp add: cont2contlubE ch2ch_cont lub_below_iff) | |
| 923 | ||
| 924 | ||
| 925 | subsection \<open>Compactness\<close> | |
| 926 | ||
| 81582 | 927 | context cpo | 
| 928 | begin | |
| 929 | ||
| 930 | definition compact :: "'a \<Rightarrow> bool" | |
| 81575 | 931 | where "compact k = adm (\<lambda>x. k \<notsqsubseteq> x)" | 
| 932 | ||
| 933 | lemma compactI: "adm (\<lambda>x. k \<notsqsubseteq> x) \<Longrightarrow> compact k" | |
| 934 | unfolding compact_def . | |
| 935 | ||
| 936 | lemma compactD: "compact k \<Longrightarrow> adm (\<lambda>x. k \<notsqsubseteq> x)" | |
| 937 | unfolding compact_def . | |
| 938 | ||
| 939 | lemma compactI2: "(\<And>Y. \<lbrakk>chain Y; x \<sqsubseteq> (\<Squnion>i. Y i)\<rbrakk> \<Longrightarrow> \<exists>i. x \<sqsubseteq> Y i) \<Longrightarrow> compact x" | |
| 940 | unfolding compact_def adm_def by fast | |
| 941 | ||
| 942 | lemma compactD2: "compact x \<Longrightarrow> chain Y \<Longrightarrow> x \<sqsubseteq> (\<Squnion>i. Y i) \<Longrightarrow> \<exists>i. x \<sqsubseteq> Y i" | |
| 943 | unfolding compact_def adm_def by fast | |
| 944 | ||
| 945 | lemma compact_below_lub_iff: "compact x \<Longrightarrow> chain Y \<Longrightarrow> x \<sqsubseteq> (\<Squnion>i. Y i) \<longleftrightarrow> (\<exists>i. x \<sqsubseteq> Y i)" | |
| 946 | by (fast intro: compactD2 elim: below_lub) | |
| 947 | ||
| 81582 | 948 | end | 
| 949 | ||
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 950 | lemma compact_chfin [simp]: "compact x" for x :: "'a::chfin" | 
| 81575 | 951 | by (rule compactI [OF adm_chfin]) | 
| 952 | ||
| 953 | lemma compact_imp_max_in_chain: "chain Y \<Longrightarrow> compact (\<Squnion>i. Y i) \<Longrightarrow> \<exists>i. max_in_chain i Y" | |
| 954 | apply (drule (1) compactD2, simp) | |
| 955 | apply (erule exE, rule_tac x=i in exI) | |
| 956 | apply (rule max_in_chainI) | |
| 957 | apply (rule below_antisym) | |
| 958 | apply (erule (1) chain_mono) | |
| 959 | apply (erule (1) below_trans [OF is_ub_thelub]) | |
| 960 | done | |
| 961 | ||
| 962 | text \<open>admissibility and compactness\<close> | |
| 963 | ||
| 964 | lemma adm_compact_not_below [simp]: | |
| 965 | "compact k \<Longrightarrow> cont (\<lambda>x. t x) \<Longrightarrow> adm (\<lambda>x. k \<notsqsubseteq> t x)" | |
| 966 | unfolding compact_def by (rule adm_subst) | |
| 967 | ||
| 968 | lemma adm_neq_compact [simp]: "compact k \<Longrightarrow> cont (\<lambda>x. t x) \<Longrightarrow> adm (\<lambda>x. t x \<noteq> k)" | |
| 969 | by (simp add: po_eq_conv) | |
| 970 | ||
| 971 | lemma adm_compact_neq [simp]: "compact k \<Longrightarrow> cont (\<lambda>x. t x) \<Longrightarrow> adm (\<lambda>x. k \<noteq> t x)" | |
| 972 | by (simp add: po_eq_conv) | |
| 973 | ||
| 974 | lemma compact_bottom [simp, intro]: "compact \<bottom>" | |
| 975 | by (rule compactI) simp | |
| 976 | ||
| 977 | text \<open>Any upward-closed predicate is admissible.\<close> | |
| 978 | ||
| 979 | lemma adm_upward: | |
| 980 | assumes P: "\<And>x y. \<lbrakk>P x; x \<sqsubseteq> y\<rbrakk> \<Longrightarrow> P y" | |
| 981 | shows "adm P" | |
| 982 | by (rule admI, drule spec, erule P, erule is_ub_thelub) | |
| 983 | ||
| 984 | lemmas adm_lemmas = | |
| 985 | adm_const adm_conj adm_all adm_ball adm_disj adm_imp adm_iff | |
| 986 | adm_below adm_eq adm_not_below | |
| 987 | adm_compact_not_below adm_compact_neq adm_neq_compact | |
| 988 | ||
| 81577 | 989 | |
| 81575 | 990 | section \<open>Class instances for the full function space\<close> | 
| 991 | ||
| 992 | subsection \<open>Full function space is a partial order\<close> | |
| 993 | ||
| 994 | instantiation "fun" :: (type, below) below | |
| 995 | begin | |
| 996 | ||
| 997 | definition below_fun_def: "(\<sqsubseteq>) \<equiv> (\<lambda>f g. \<forall>x. f x \<sqsubseteq> g x)" | |
| 998 | ||
| 999 | instance .. | |
| 1000 | end | |
| 1001 | ||
| 1002 | instance "fun" :: (type, po) po | |
| 1003 | proof | |
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 1004 | fix f g h :: "'a \<Rightarrow> 'b" | 
| 81575 | 1005 | show "f \<sqsubseteq> f" | 
| 1006 | by (simp add: below_fun_def) | |
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 1007 | show "f \<sqsubseteq> g \<Longrightarrow> g \<sqsubseteq> f \<Longrightarrow> f = g" | 
| 81575 | 1008 | by (simp add: below_fun_def fun_eq_iff below_antisym) | 
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 1009 | show "f \<sqsubseteq> g \<Longrightarrow> g \<sqsubseteq> h \<Longrightarrow> f \<sqsubseteq> h" | 
| 81575 | 1010 | unfolding below_fun_def by (fast elim: below_trans) | 
| 1011 | qed | |
| 1012 | ||
| 1013 | lemma fun_below_iff: "f \<sqsubseteq> g \<longleftrightarrow> (\<forall>x. f x \<sqsubseteq> g x)" | |
| 1014 | by (simp add: below_fun_def) | |
| 1015 | ||
| 1016 | lemma fun_belowI: "(\<And>x. f x \<sqsubseteq> g x) \<Longrightarrow> f \<sqsubseteq> g" | |
| 1017 | by (simp add: below_fun_def) | |
| 1018 | ||
| 1019 | lemma fun_belowD: "f \<sqsubseteq> g \<Longrightarrow> f x \<sqsubseteq> g x" | |
| 1020 | by (simp add: below_fun_def) | |
| 1021 | ||
| 1022 | ||
| 1023 | subsection \<open>Full function space is chain complete\<close> | |
| 1024 | ||
| 1025 | text \<open>Properties of chains of functions.\<close> | |
| 1026 | ||
| 1027 | lemma fun_chain_iff: "chain S \<longleftrightarrow> (\<forall>x. chain (\<lambda>i. S i x))" | |
| 1028 | by (auto simp: chain_def fun_below_iff) | |
| 1029 | ||
| 1030 | lemma ch2ch_fun: "chain S \<Longrightarrow> chain (\<lambda>i. S i x)" | |
| 1031 | by (simp add: chain_def below_fun_def) | |
| 1032 | ||
| 1033 | lemma ch2ch_lambda: "(\<And>x. chain (\<lambda>i. S i x)) \<Longrightarrow> chain S" | |
| 1034 | by (simp add: chain_def below_fun_def) | |
| 1035 | ||
| 1036 | text \<open>Type \<^typ>\<open>'a::type \<Rightarrow> 'b::cpo\<close> is chain complete\<close> | |
| 1037 | ||
| 1038 | lemma is_lub_lambda: "(\<And>x. range (\<lambda>i. Y i x) <<| f x) \<Longrightarrow> range Y <<| f" | |
| 1039 | by (simp add: is_lub_def is_ub_def below_fun_def) | |
| 1040 | ||
| 1041 | lemma is_lub_fun: "chain S \<Longrightarrow> range S <<| (\<lambda>x. \<Squnion>i. S i x)" | |
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 1042 | for S :: "nat \<Rightarrow> 'a::type \<Rightarrow> 'b" | 
| 81575 | 1043 | apply (rule is_lub_lambda) | 
| 1044 | apply (rule cpo_lubI) | |
| 1045 | apply (erule ch2ch_fun) | |
| 1046 | done | |
| 1047 | ||
| 1048 | lemma lub_fun: "chain S \<Longrightarrow> (\<Squnion>i. S i) = (\<lambda>x. \<Squnion>i. S i x)" | |
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 1049 | for S :: "nat \<Rightarrow> 'a::type \<Rightarrow> 'b" | 
| 81575 | 1050 | by (rule is_lub_fun [THEN lub_eqI]) | 
| 1051 | ||
| 1052 | instance "fun" :: (type, cpo) cpo | |
| 1053 | by intro_classes (rule exI, erule is_lub_fun) | |
| 1054 | ||
| 1055 | instance "fun" :: (type, discrete_cpo) discrete_cpo | |
| 1056 | proof | |
| 1057 | fix f g :: "'a \<Rightarrow> 'b" | |
| 1058 | show "f \<sqsubseteq> g \<longleftrightarrow> f = g" | |
| 1059 | by (simp add: fun_below_iff fun_eq_iff) | |
| 1060 | qed | |
| 1061 | ||
| 1062 | ||
| 1063 | subsection \<open>Full function space is pointed\<close> | |
| 1064 | ||
| 1065 | lemma minimal_fun: "(\<lambda>x. \<bottom>) \<sqsubseteq> f" | |
| 1066 | by (simp add: below_fun_def) | |
| 1067 | ||
| 1068 | instance "fun" :: (type, pcpo) pcpo | |
| 1069 | by standard (fast intro: minimal_fun) | |
| 1070 | ||
| 1071 | lemma inst_fun_pcpo: "\<bottom> = (\<lambda>x. \<bottom>)" | |
| 1072 | by (rule minimal_fun [THEN bottomI, symmetric]) | |
| 1073 | ||
| 1074 | lemma app_strict [simp]: "\<bottom> x = \<bottom>" | |
| 1075 | by (simp add: inst_fun_pcpo) | |
| 1076 | ||
| 1077 | lemma lambda_strict: "(\<lambda>x. \<bottom>) = \<bottom>" | |
| 1078 | by (rule bottomI, rule minimal_fun) | |
| 1079 | ||
| 1080 | ||
| 1081 | subsection \<open>Propagation of monotonicity and continuity\<close> | |
| 1082 | ||
| 1083 | text \<open>The lub of a chain of monotone functions is monotone.\<close> | |
| 1084 | ||
| 1085 | lemma adm_monofun: "adm monofun" | |
| 1086 | by (rule admI) (simp add: lub_fun fun_chain_iff monofun_def lub_mono) | |
| 1087 | ||
| 1088 | text \<open>The lub of a chain of continuous functions is continuous.\<close> | |
| 1089 | ||
| 1090 | lemma adm_cont: "adm cont" | |
| 1091 | by (rule admI) (simp add: lub_fun fun_chain_iff) | |
| 1092 | ||
| 1093 | text \<open>Function application preserves monotonicity and continuity.\<close> | |
| 1094 | ||
| 1095 | lemma mono2mono_fun: "monofun f \<Longrightarrow> monofun (\<lambda>x. f x y)" | |
| 1096 | by (simp add: monofun_def fun_below_iff) | |
| 1097 | ||
| 1098 | lemma cont2cont_fun: "cont f \<Longrightarrow> cont (\<lambda>x. f x y)" | |
| 1099 | apply (rule contI2) | |
| 1100 | apply (erule cont2mono [THEN mono2mono_fun]) | |
| 1101 | apply (simp add: cont2contlubE lub_fun ch2ch_cont) | |
| 1102 | done | |
| 1103 | ||
| 1104 | lemma cont_fun: "cont (\<lambda>f. f x)" | |
| 1105 | using cont_id by (rule cont2cont_fun) | |
| 1106 | ||
| 1107 | text \<open> | |
| 1108 | Lambda abstraction preserves monotonicity and continuity. | |
| 1109 | (Note \<open>(\<lambda>x. \<lambda>y. f x y) = f\<close>.) | |
| 1110 | \<close> | |
| 1111 | ||
| 1112 | lemma mono2mono_lambda: "(\<And>y. monofun (\<lambda>x. f x y)) \<Longrightarrow> monofun f" | |
| 1113 | by (simp add: monofun_def fun_below_iff) | |
| 1114 | ||
| 1115 | lemma cont2cont_lambda [simp]: | |
| 1116 | assumes f: "\<And>y. cont (\<lambda>x. f x y)" | |
| 1117 | shows "cont f" | |
| 1118 | by (rule contI, rule is_lub_lambda, rule contE [OF f]) | |
| 1119 | ||
| 1120 | text \<open>What D.A.Schmidt calls continuity of abstraction; never used here\<close> | |
| 1121 | ||
| 1122 | lemma contlub_lambda: "(\<And>x. chain (\<lambda>i. S i x)) \<Longrightarrow> (\<lambda>x. \<Squnion>i. S i x) = (\<Squnion>i. (\<lambda>x. S i x))" | |
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 1123 | for S :: "nat \<Rightarrow> 'a::type \<Rightarrow> 'b" | 
| 81575 | 1124 | by (simp add: lub_fun ch2ch_lambda) | 
| 1125 | ||
| 81577 | 1126 | |
| 81575 | 1127 | section \<open>The cpo of cartesian products\<close> | 
| 1128 | ||
| 1129 | subsection \<open>Unit type is a pcpo\<close> | |
| 1130 | ||
| 1131 | instantiation unit :: discrete_cpo | |
| 1132 | begin | |
| 1133 | ||
| 1134 | definition below_unit_def [simp]: "x \<sqsubseteq> (y::unit) \<longleftrightarrow> True" | |
| 1135 | ||
| 1136 | instance | |
| 1137 | by standard simp | |
| 1138 | ||
| 1139 | end | |
| 1140 | ||
| 1141 | instance unit :: pcpo | |
| 1142 | by standard simp | |
| 1143 | ||
| 1144 | ||
| 1145 | subsection \<open>Product type is a partial order\<close> | |
| 1146 | ||
| 1147 | instantiation prod :: (below, below) below | |
| 1148 | begin | |
| 1149 | ||
| 1150 | definition below_prod_def: "(\<sqsubseteq>) \<equiv> \<lambda>p1 p2. (fst p1 \<sqsubseteq> fst p2 \<and> snd p1 \<sqsubseteq> snd p2)" | |
| 1151 | ||
| 1152 | instance .. | |
| 1153 | ||
| 1154 | end | |
| 1155 | ||
| 1156 | instance prod :: (po, po) po | |
| 1157 | proof | |
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 1158 | fix x y z :: "'a \<times> 'b" | 
| 81575 | 1159 | show "x \<sqsubseteq> x" | 
| 1160 | by (simp add: below_prod_def) | |
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 1161 | show "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> x \<Longrightarrow> x = y" | 
| 81575 | 1162 | unfolding below_prod_def prod_eq_iff | 
| 1163 | by (fast intro: below_antisym) | |
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 1164 | show "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> z" | 
| 81575 | 1165 | unfolding below_prod_def | 
| 1166 | by (fast intro: below_trans) | |
| 1167 | qed | |
| 1168 | ||
| 1169 | ||
| 1170 | subsection \<open>Monotonicity of \emph{Pair}, \emph{fst}, \emph{snd}\<close>
 | |
| 1171 | ||
| 1172 | lemma prod_belowI: "fst p \<sqsubseteq> fst q \<Longrightarrow> snd p \<sqsubseteq> snd q \<Longrightarrow> p \<sqsubseteq> q" | |
| 1173 | by (simp add: below_prod_def) | |
| 1174 | ||
| 1175 | lemma Pair_below_iff [simp]: "(a, b) \<sqsubseteq> (c, d) \<longleftrightarrow> a \<sqsubseteq> c \<and> b \<sqsubseteq> d" | |
| 1176 | by (simp add: below_prod_def) | |
| 1177 | ||
| 1178 | text \<open>Pair \<open>(_,_)\<close> is monotone in both arguments\<close> | |
| 1179 | ||
| 1180 | lemma monofun_pair1: "monofun (\<lambda>x. (x, y))" | |
| 1181 | by (simp add: monofun_def) | |
| 1182 | ||
| 1183 | lemma monofun_pair2: "monofun (\<lambda>y. (x, y))" | |
| 1184 | by (simp add: monofun_def) | |
| 1185 | ||
| 1186 | lemma monofun_pair: "x1 \<sqsubseteq> x2 \<Longrightarrow> y1 \<sqsubseteq> y2 \<Longrightarrow> (x1, y1) \<sqsubseteq> (x2, y2)" | |
| 1187 | by simp | |
| 1188 | ||
| 1189 | lemma ch2ch_Pair [simp]: "chain X \<Longrightarrow> chain Y \<Longrightarrow> chain (\<lambda>i. (X i, Y i))" | |
| 1190 | by (rule chainI, simp add: chainE) | |
| 1191 | ||
| 1192 | text \<open>\<^term>\<open>fst\<close> and \<^term>\<open>snd\<close> are monotone\<close> | |
| 1193 | ||
| 1194 | lemma fst_monofun: "x \<sqsubseteq> y \<Longrightarrow> fst x \<sqsubseteq> fst y" | |
| 1195 | by (simp add: below_prod_def) | |
| 1196 | ||
| 1197 | lemma snd_monofun: "x \<sqsubseteq> y \<Longrightarrow> snd x \<sqsubseteq> snd y" | |
| 1198 | by (simp add: below_prod_def) | |
| 1199 | ||
| 1200 | lemma monofun_fst: "monofun fst" | |
| 1201 | by (simp add: monofun_def below_prod_def) | |
| 1202 | ||
| 1203 | lemma monofun_snd: "monofun snd" | |
| 1204 | by (simp add: monofun_def below_prod_def) | |
| 1205 | ||
| 1206 | lemmas ch2ch_fst [simp] = ch2ch_monofun [OF monofun_fst] | |
| 1207 | ||
| 1208 | lemmas ch2ch_snd [simp] = ch2ch_monofun [OF monofun_snd] | |
| 1209 | ||
| 1210 | lemma prod_chain_cases: | |
| 1211 | assumes chain: "chain Y" | |
| 1212 | obtains A B | |
| 1213 | where "chain A" and "chain B" and "Y = (\<lambda>i. (A i, B i))" | |
| 1214 | proof | |
| 1215 | from chain show "chain (\<lambda>i. fst (Y i))" | |
| 1216 | by (rule ch2ch_fst) | |
| 1217 | from chain show "chain (\<lambda>i. snd (Y i))" | |
| 1218 | by (rule ch2ch_snd) | |
| 1219 | show "Y = (\<lambda>i. (fst (Y i), snd (Y i)))" | |
| 1220 | by simp | |
| 1221 | qed | |
| 1222 | ||
| 1223 | ||
| 1224 | subsection \<open>Product type is a cpo\<close> | |
| 1225 | ||
| 1226 | lemma is_lub_Pair: "range A <<| x \<Longrightarrow> range B <<| y \<Longrightarrow> range (\<lambda>i. (A i, B i)) <<| (x, y)" | |
| 1227 | by (simp add: is_lub_def is_ub_def below_prod_def) | |
| 1228 | ||
| 1229 | lemma lub_Pair: "chain A \<Longrightarrow> chain B \<Longrightarrow> (\<Squnion>i. (A i, B i)) = (\<Squnion>i. A i, \<Squnion>i. B i)" | |
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 1230 | for A :: "nat \<Rightarrow> 'a" and B :: "nat \<Rightarrow> 'b" | 
| 81575 | 1231 | by (fast intro: lub_eqI is_lub_Pair elim: thelubE) | 
| 1232 | ||
| 1233 | lemma is_lub_prod: | |
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 1234 |   fixes S :: "nat \<Rightarrow> ('a \<times> 'b)"
 | 
| 81575 | 1235 | assumes "chain S" | 
| 1236 | shows "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))" | |
| 1237 | using assms by (auto elim: prod_chain_cases simp: is_lub_Pair cpo_lubI) | |
| 1238 | ||
| 1239 | lemma lub_prod: "chain S \<Longrightarrow> (\<Squnion>i. S i) = (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))" | |
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 1240 | for S :: "nat \<Rightarrow> 'a \<times> 'b" | 
| 81575 | 1241 | by (rule is_lub_prod [THEN lub_eqI]) | 
| 1242 | ||
| 1243 | instance prod :: (cpo, cpo) cpo | |
| 1244 | proof | |
| 1245 |   fix S :: "nat \<Rightarrow> ('a \<times> 'b)"
 | |
| 1246 | assume "chain S" | |
| 1247 | then have "range S <<| (\<Squnion>i. fst (S i), \<Squnion>i. snd (S i))" | |
| 1248 | by (rule is_lub_prod) | |
| 1249 | then show "\<exists>x. range S <<| x" .. | |
| 1250 | qed | |
| 1251 | ||
| 1252 | instance prod :: (discrete_cpo, discrete_cpo) discrete_cpo | |
| 1253 | proof | |
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 1254 | show "x \<sqsubseteq> y \<longleftrightarrow> x = y" for x y :: "'a \<times> 'b" | 
| 81575 | 1255 | by (simp add: below_prod_def prod_eq_iff) | 
| 1256 | qed | |
| 1257 | ||
| 1258 | ||
| 1259 | subsection \<open>Product type is pointed\<close> | |
| 1260 | ||
| 1261 | lemma minimal_prod: "(\<bottom>, \<bottom>) \<sqsubseteq> p" | |
| 1262 | by (simp add: below_prod_def) | |
| 1263 | ||
| 1264 | instance prod :: (pcpo, pcpo) pcpo | |
| 1265 | by intro_classes (fast intro: minimal_prod) | |
| 1266 | ||
| 1267 | lemma inst_prod_pcpo: "\<bottom> = (\<bottom>, \<bottom>)" | |
| 1268 | by (rule minimal_prod [THEN bottomI, symmetric]) | |
| 1269 | ||
| 1270 | lemma Pair_bottom_iff [simp]: "(x, y) = \<bottom> \<longleftrightarrow> x = \<bottom> \<and> y = \<bottom>" | |
| 1271 | by (simp add: inst_prod_pcpo) | |
| 1272 | ||
| 1273 | lemma fst_strict [simp]: "fst \<bottom> = \<bottom>" | |
| 1274 | unfolding inst_prod_pcpo by (rule fst_conv) | |
| 1275 | ||
| 1276 | lemma snd_strict [simp]: "snd \<bottom> = \<bottom>" | |
| 1277 | unfolding inst_prod_pcpo by (rule snd_conv) | |
| 1278 | ||
| 1279 | lemma Pair_strict [simp]: "(\<bottom>, \<bottom>) = \<bottom>" | |
| 1280 | by simp | |
| 1281 | ||
| 1282 | lemma split_strict [simp]: "case_prod f \<bottom> = f \<bottom> \<bottom>" | |
| 1283 | by (simp add: split_def) | |
| 1284 | ||
| 1285 | ||
| 1286 | subsection \<open>Continuity of \emph{Pair}, \emph{fst}, \emph{snd}\<close>
 | |
| 1287 | ||
| 1288 | lemma cont_pair1: "cont (\<lambda>x. (x, y))" | |
| 1289 | apply (rule contI) | |
| 1290 | apply (rule is_lub_Pair) | |
| 1291 | apply (erule cpo_lubI) | |
| 1292 | apply (rule is_lub_const) | |
| 1293 | done | |
| 1294 | ||
| 1295 | lemma cont_pair2: "cont (\<lambda>y. (x, y))" | |
| 1296 | apply (rule contI) | |
| 1297 | apply (rule is_lub_Pair) | |
| 1298 | apply (rule is_lub_const) | |
| 1299 | apply (erule cpo_lubI) | |
| 1300 | done | |
| 1301 | ||
| 1302 | lemma cont_fst: "cont fst" | |
| 1303 | apply (rule contI) | |
| 1304 | apply (simp add: lub_prod) | |
| 1305 | apply (erule cpo_lubI [OF ch2ch_fst]) | |
| 1306 | done | |
| 1307 | ||
| 1308 | lemma cont_snd: "cont snd" | |
| 1309 | apply (rule contI) | |
| 1310 | apply (simp add: lub_prod) | |
| 1311 | apply (erule cpo_lubI [OF ch2ch_snd]) | |
| 1312 | done | |
| 1313 | ||
| 1314 | lemma cont2cont_Pair [simp, cont2cont]: | |
| 1315 | assumes f: "cont (\<lambda>x. f x)" | |
| 1316 | assumes g: "cont (\<lambda>x. g x)" | |
| 1317 | shows "cont (\<lambda>x. (f x, g x))" | |
| 1318 | apply (rule cont_apply [OF f cont_pair1]) | |
| 1319 | apply (rule cont_apply [OF g cont_pair2]) | |
| 1320 | apply (rule cont_const) | |
| 1321 | done | |
| 1322 | ||
| 1323 | lemmas cont2cont_fst [simp, cont2cont] = cont_compose [OF cont_fst] | |
| 1324 | ||
| 1325 | lemmas cont2cont_snd [simp, cont2cont] = cont_compose [OF cont_snd] | |
| 1326 | ||
| 1327 | lemma cont2cont_case_prod: | |
| 1328 | assumes f1: "\<And>a b. cont (\<lambda>x. f x a b)" | |
| 1329 | assumes f2: "\<And>x b. cont (\<lambda>a. f x a b)" | |
| 1330 | assumes f3: "\<And>x a. cont (\<lambda>b. f x a b)" | |
| 1331 | assumes g: "cont (\<lambda>x. g x)" | |
| 1332 | shows "cont (\<lambda>x. case g x of (a, b) \<Rightarrow> f x a b)" | |
| 1333 | unfolding split_def | |
| 1334 | apply (rule cont_apply [OF g]) | |
| 1335 | apply (rule cont_apply [OF cont_fst f2]) | |
| 1336 | apply (rule cont_apply [OF cont_snd f3]) | |
| 1337 | apply (rule cont_const) | |
| 1338 | apply (rule f1) | |
| 1339 | done | |
| 1340 | ||
| 1341 | lemma prod_contI: | |
| 1342 | assumes f1: "\<And>y. cont (\<lambda>x. f (x, y))" | |
| 1343 | assumes f2: "\<And>x. cont (\<lambda>y. f (x, y))" | |
| 1344 | shows "cont f" | |
| 1345 | proof - | |
| 1346 | have "cont (\<lambda>(x, y). f (x, y))" | |
| 1347 | by (intro cont2cont_case_prod f1 f2 cont2cont) | |
| 1348 | then show "cont f" | |
| 1349 | by (simp only: case_prod_eta) | |
| 1350 | qed | |
| 1351 | ||
| 1352 | lemma prod_cont_iff: "cont f \<longleftrightarrow> (\<forall>y. cont (\<lambda>x. f (x, y))) \<and> (\<forall>x. cont (\<lambda>y. f (x, y)))" | |
| 1353 | apply safe | |
| 1354 | apply (erule cont_compose [OF _ cont_pair1]) | |
| 1355 | apply (erule cont_compose [OF _ cont_pair2]) | |
| 1356 | apply (simp only: prod_contI) | |
| 1357 | done | |
| 1358 | ||
| 1359 | lemma cont2cont_case_prod' [simp, cont2cont]: | |
| 1360 | assumes f: "cont (\<lambda>p. f (fst p) (fst (snd p)) (snd (snd p)))" | |
| 1361 | assumes g: "cont (\<lambda>x. g x)" | |
| 1362 | shows "cont (\<lambda>x. case_prod (f x) (g x))" | |
| 1363 | using assms by (simp add: cont2cont_case_prod prod_cont_iff) | |
| 1364 | ||
| 1365 | text \<open>The simple version (due to Joachim Breitner) is needed if | |
| 1366 | either element type of the pair is not a cpo.\<close> | |
| 1367 | ||
| 1368 | lemma cont2cont_split_simple [simp, cont2cont]: | |
| 1369 | assumes "\<And>a b. cont (\<lambda>x. f x a b)" | |
| 1370 | shows "cont (\<lambda>x. case p of (a, b) \<Rightarrow> f x a b)" | |
| 1371 | using assms by (cases p) auto | |
| 1372 | ||
| 1373 | text \<open>Admissibility of predicates on product types.\<close> | |
| 1374 | ||
| 1375 | lemma adm_case_prod [simp]: | |
| 1376 | assumes "adm (\<lambda>x. P x (fst (f x)) (snd (f x)))" | |
| 1377 | shows "adm (\<lambda>x. case f x of (a, b) \<Rightarrow> P x a b)" | |
| 1378 | unfolding case_prod_beta using assms . | |
| 1379 | ||
| 1380 | ||
| 1381 | subsection \<open>Compactness and chain-finiteness\<close> | |
| 1382 | ||
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 1383 | lemma fst_below_iff: "fst x \<sqsubseteq> y \<longleftrightarrow> x \<sqsubseteq> (y, snd x)" for x :: "'a \<times> 'b" | 
| 81575 | 1384 | by (simp add: below_prod_def) | 
| 1385 | ||
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 1386 | lemma snd_below_iff: "snd x \<sqsubseteq> y \<longleftrightarrow> x \<sqsubseteq> (fst x, y)" for x :: "'a \<times> 'b" | 
| 81575 | 1387 | by (simp add: below_prod_def) | 
| 1388 | ||
| 1389 | lemma compact_fst: "compact x \<Longrightarrow> compact (fst x)" | |
| 1390 | by (rule compactI) (simp add: fst_below_iff) | |
| 1391 | ||
| 1392 | lemma compact_snd: "compact x \<Longrightarrow> compact (snd x)" | |
| 1393 | by (rule compactI) (simp add: snd_below_iff) | |
| 1394 | ||
| 1395 | lemma compact_Pair: "compact x \<Longrightarrow> compact y \<Longrightarrow> compact (x, y)" | |
| 1396 | by (rule compactI) (simp add: below_prod_def) | |
| 1397 | ||
| 1398 | lemma compact_Pair_iff [simp]: "compact (x, y) \<longleftrightarrow> compact x \<and> compact y" | |
| 1399 | apply (safe intro!: compact_Pair) | |
| 1400 | apply (drule compact_fst, simp) | |
| 1401 | apply (drule compact_snd, simp) | |
| 1402 | done | |
| 1403 | ||
| 1404 | instance prod :: (chfin, chfin) chfin | |
| 1405 | apply intro_classes | |
| 1406 | apply (erule compact_imp_max_in_chain) | |
| 1407 | apply (case_tac "\<Squnion>i. Y i", simp) | |
| 1408 | done | |
| 1409 | ||
| 81577 | 1410 | |
| 81575 | 1411 | section \<open>Discrete cpo types\<close> | 
| 1412 | ||
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 1413 | datatype 'a discr = Discr "'a::type" | 
| 81575 | 1414 | |
| 1415 | subsection \<open>Discrete cpo class instance\<close> | |
| 1416 | ||
| 1417 | instantiation discr :: (type) discrete_cpo | |
| 1418 | begin | |
| 1419 | ||
| 1420 | definition "((\<sqsubseteq>) :: 'a discr \<Rightarrow> 'a discr \<Rightarrow> bool) = (=)" | |
| 1421 | ||
| 1422 | instance | |
| 1423 | by standard (simp add: below_discr_def) | |
| 1424 | ||
| 1425 | end | |
| 1426 | ||
| 1427 | ||
| 1428 | subsection \<open>\emph{undiscr}\<close>
 | |
| 1429 | ||
| 81583 
b6df83045178
clarified default_sort: "cpo" for bootstrap, "domain" for main HOLCF;
 wenzelm parents: 
81582diff
changeset | 1430 | definition undiscr :: "'a::type discr \<Rightarrow> 'a" | 
| 81575 | 1431 | where "undiscr x = (case x of Discr y \<Rightarrow> y)" | 
| 1432 | ||
| 1433 | lemma undiscr_Discr [simp]: "undiscr (Discr x) = x" | |
| 1434 | by (simp add: undiscr_def) | |
| 1435 | ||
| 1436 | lemma Discr_undiscr [simp]: "Discr (undiscr y) = y" | |
| 1437 | by (induct y) simp | |
| 1438 | ||
| 1439 | end |