| author | wenzelm | 
| Thu, 06 Dec 2001 17:16:30 +0100 | |
| changeset 12412 | d0857ea70f23 | 
| parent 12338 | de0f4a63baa5 | 
| child 14096 | f79d139c7e46 | 
| permissions | -rw-r--r-- | 
| 2640 | 1 | (* Title: HOLCF/Lift.thy | 
| 2 | ID: $Id$ | |
| 12026 | 3 | Author: Olaf Mueller | 
| 4 | License: GPL (GNU GENERAL PUBLIC LICENSE) | |
| 2640 | 5 | *) | 
| 6 | ||
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12026diff
changeset | 7 | header {* Lifting types of class type to flat pcpo's *}
 | 
| 12026 | 8 | |
| 9 | theory Lift = Cprod3: | |
| 10 | ||
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12026diff
changeset | 11 | defaultsort type | 
| 12026 | 12 | |
| 13 | ||
| 14 | typedef 'a lift = "UNIV :: 'a option set" .. | |
| 15 | ||
| 16 | constdefs | |
| 17 | Undef :: "'a lift" | |
| 18 | "Undef == Abs_lift None" | |
| 19 | Def :: "'a => 'a lift" | |
| 20 | "Def x == Abs_lift (Some x)" | |
| 21 | ||
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12026diff
changeset | 22 | instance lift :: (type) sq_ord .. | 
| 12026 | 23 | |
| 24 | defs (overloaded) | |
| 25 | less_lift_def: "x << y == (x=y | x=Undef)" | |
| 26 | ||
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12026diff
changeset | 27 | instance lift :: (type) po | 
| 12026 | 28 | proof | 
| 29 | fix x y z :: "'a lift" | |
| 30 | show "x << x" by (unfold less_lift_def) blast | |
| 31 |   { assume "x << y" and "y << x" thus "x = y" by (unfold less_lift_def) blast }
 | |
| 32 |   { assume "x << y" and "y << z" thus "x << z" by (unfold less_lift_def) blast }
 | |
| 33 | qed | |
| 34 | ||
| 35 | lemma inst_lift_po: "(op <<) = (\<lambda>x y. x = y | x = Undef)" | |
| 36 |   -- {* For compatibility with old HOLCF-Version. *}
 | |
| 37 | by (simp only: less_lift_def [symmetric]) | |
| 38 | ||
| 39 | ||
| 40 | subsection {* Type lift is pointed *}
 | |
| 41 | ||
| 42 | lemma minimal_lift [iff]: "Undef << x" | |
| 43 | by (simp add: inst_lift_po) | |
| 44 | ||
| 45 | lemma UU_lift_def: "(SOME u. \<forall>y. u \<sqsubseteq> y) = Undef" | |
| 46 | apply (rule minimal2UU [symmetric]) | |
| 47 | apply (rule minimal_lift) | |
| 48 | done | |
| 49 | ||
| 50 | lemma least_lift: "EX x::'a lift. ALL y. x << y" | |
| 51 | apply (rule_tac x = Undef in exI) | |
| 52 | apply (rule minimal_lift [THEN allI]) | |
| 53 | done | |
| 54 | ||
| 55 | ||
| 56 | subsection {* Type lift is a cpo *}
 | |
| 57 | ||
| 58 | text {*
 | |
| 59 |   The following lemmas have already been proved in @{text Pcpo.ML} and
 | |
| 60 |   @{text Fix.ML}, but there class @{text pcpo} is assumed, although
 | |
| 61 |   only @{text po} is necessary and a least element. Therefore they are
 | |
| 62 |   redone here for the @{text po} lift with least element @{text
 | |
| 63 | Undef}. | |
| 64 | *} | |
| 65 | ||
| 66 | lemma notUndef_I: "[| x<<y; x ~= Undef |] ==> y ~= Undef" | |
| 67 |   -- {* Tailoring @{text notUU_I} of @{text Pcpo.ML} to @{text Undef} *}
 | |
| 68 | by (blast intro: antisym_less) | |
| 69 | ||
| 70 | lemma chain_mono2_po: "[| EX j.~Y(j)=Undef; chain(Y::nat=>('a)lift) |]
 | |
| 71 | ==> EX j. ALL i. j<i-->~Y(i)=Undef" | |
| 72 |   -- {* Tailoring @{text chain_mono2} of @{text Pcpo.ML} to @{text Undef} *}
 | |
| 73 | apply safe | |
| 74 | apply (rule exI) | |
| 75 | apply (intro strip) | |
| 76 | apply (rule notUndef_I) | |
| 77 | apply (erule (1) chain_mono) | |
| 78 | apply assumption | |
| 79 | done | |
| 80 | ||
| 81 | lemma flat_imp_chfin_poo: "(ALL Y. chain(Y::nat=>('a)lift)-->(EX n. max_in_chain n Y))"
 | |
| 82 |   -- {* Tailoring @{text flat_imp_chfin} of @{text Fix.ML} to @{text lift} *}
 | |
| 83 | apply (unfold max_in_chain_def) | |
| 84 | apply (intro strip) | |
| 85 | apply (rule_tac P = "ALL i. Y (i) = Undef" in case_split) | |
| 86 | apply (rule_tac x = 0 in exI) | |
| 87 | apply (intro strip) | |
| 88 | apply (rule trans) | |
| 89 | apply (erule spec) | |
| 90 | apply (rule sym) | |
| 91 | apply (erule spec) | |
| 92 |   apply (subgoal_tac "ALL x y. x << (y:: ('a) lift) --> x=Undef | x=y")
 | |
| 93 | prefer 2 apply (simp add: inst_lift_po) | |
| 94 | apply (rule chain_mono2_po [THEN exE]) | |
| 95 | apply fast | |
| 96 | apply assumption | |
| 97 | apply (rule_tac x = "Suc x" in exI) | |
| 98 | apply (intro strip) | |
| 99 | apply (rule disjE) | |
| 100 | prefer 3 apply assumption | |
| 101 | apply (rule mp) | |
| 102 | apply (drule spec) | |
| 103 | apply (erule spec) | |
| 104 | apply (erule le_imp_less_or_eq [THEN disjE]) | |
| 105 | apply (erule chain_mono) | |
| 106 | apply auto | |
| 107 | done | |
| 108 | ||
| 109 | theorem cpo_lift: "chain (Y::nat => 'a lift) ==> EX x. range Y <<| x" | |
| 110 | apply (cut_tac flat_imp_chfin_poo) | |
| 111 | apply (blast intro: lub_finch1) | |
| 112 | done | |
| 113 | ||
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12026diff
changeset | 114 | instance lift :: (type) pcpo | 
| 12026 | 115 | apply intro_classes | 
| 116 | apply (erule cpo_lift) | |
| 117 | apply (rule least_lift) | |
| 118 | done | |
| 119 | ||
| 120 | lemma inst_lift_pcpo: "UU = Undef" | |
| 121 |   -- {* For compatibility with old HOLCF-Version. *}
 | |
| 122 | by (simp add: UU_def UU_lift_def) | |
| 123 | ||
| 124 | ||
| 125 | subsection {* Lift as a datatype *}
 | |
| 126 | ||
| 127 | lemma lift_distinct1: "UU ~= Def x" | |
| 128 | by (simp add: Undef_def Def_def Abs_lift_inject lift_def inst_lift_pcpo) | |
| 129 | ||
| 130 | lemma lift_distinct2: "Def x ~= UU" | |
| 131 | by (simp add: Undef_def Def_def Abs_lift_inject lift_def inst_lift_pcpo) | |
| 132 | ||
| 133 | lemma Def_inject: "(Def x = Def x') = (x = x')" | |
| 134 | by (simp add: Def_def Abs_lift_inject lift_def) | |
| 135 | ||
| 136 | lemma lift_induct: "P UU ==> (!!x. P (Def x)) ==> P y" | |
| 137 | apply (induct y) | |
| 138 | apply (induct_tac y) | |
| 139 | apply (simp_all add: Undef_def Def_def inst_lift_pcpo) | |
| 140 | done | |
| 141 | ||
| 142 | rep_datatype lift | |
| 143 | distinct lift_distinct1 lift_distinct2 | |
| 144 | inject Def_inject | |
| 145 | induction lift_induct | |
| 146 | ||
| 147 | lemma Def_not_UU: "Def a ~= UU" | |
| 148 | by simp | |
| 149 | ||
| 150 | ||
| 151 | subsection {* Further operations *}
 | |
| 152 | ||
| 153 | consts | |
| 154 |  flift1      :: "('a => 'b::pcpo) => ('a lift -> 'b)"
 | |
| 155 |  flift2      :: "('a => 'b)       => ('a lift -> 'b lift)"
 | |
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12026diff
changeset | 156 |  liftpair    ::"'a::type lift * 'b::type lift => ('a * 'b) lift"
 | 
| 2640 | 157 | |
| 12026 | 158 | defs | 
| 159 | flift1_def: | |
| 160 | "flift1 f == (LAM x. (case x of | |
| 161 | UU => UU | |
| 162 | | Def y => (f y)))" | |
| 163 | flift2_def: | |
| 164 | "flift2 f == (LAM x. (case x of | |
| 165 | UU => UU | |
| 166 | | Def y => Def (f y)))" | |
| 167 | liftpair_def: | |
| 168 | "liftpair x == (case (cfst$x) of | |
| 169 | UU => UU | |
| 170 | | Def x1 => (case (csnd$x) of | |
| 171 | UU => UU | |
| 172 | | Def x2 => Def (x1,x2)))" | |
| 173 | ||
| 174 | ||
| 175 | declare inst_lift_pcpo [symmetric, simp] | |
| 176 | ||
| 177 | ||
| 178 | lemma less_lift: "(x::'a lift) << y = (x=y | x=UU)" | |
| 179 | by (simp add: inst_lift_po) | |
| 180 | ||
| 181 | ||
| 182 | text {* @{text UU} and @{text Def} *}
 | |
| 183 | ||
| 184 | lemma Lift_exhaust: "x = UU | (EX y. x = Def y)" | |
| 185 | by (induct x) simp_all | |
| 186 | ||
| 187 | lemma Lift_cases: "[| x = UU ==> P; ? a. x = Def a ==> P |] ==> P" | |
| 188 | by (insert Lift_exhaust) blast | |
| 189 | ||
| 190 | lemma not_Undef_is_Def: "(x ~= UU) = (EX y. x = Def y)" | |
| 191 | by (cases x) simp_all | |
| 192 | ||
| 193 | text {*
 | |
| 194 |   For @{term "x ~= UU"} in assumptions @{text def_tac} replaces @{text
 | |
| 195 |   x} by @{text "Def a"} in conclusion. *}
 | |
| 196 | ||
| 197 | ML {*
 | |
| 198 | local val not_Undef_is_Def = thm "not_Undef_is_Def" | |
| 199 | in val def_tac = SIMPSET' (fn ss => | |
| 200 | etac (not_Undef_is_Def RS iffD1 RS exE) THEN' asm_simp_tac ss) | |
| 201 | end; | |
| 202 | *} | |
| 203 | ||
| 204 | lemma Undef_eq_UU: "Undef = UU" | |
| 205 | by (rule inst_lift_pcpo [symmetric]) | |
| 206 | ||
| 207 | lemma DefE: "Def x = UU ==> R" | |
| 208 | by simp | |
| 209 | ||
| 210 | lemma DefE2: "[| x = Def s; x = UU |] ==> R" | |
| 211 | by simp | |
| 212 | ||
| 213 | lemma Def_inject_less_eq: "Def x << Def y = (x = y)" | |
| 214 | by (simp add: less_lift_def) | |
| 215 | ||
| 216 | lemma Def_less_is_eq [simp]: "Def x << y = (Def x = y)" | |
| 217 | by (simp add: less_lift) | |
| 218 | ||
| 219 | ||
| 220 | subsection {* Lift is flat *}
 | |
| 221 | ||
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
12026diff
changeset | 222 | instance lift :: (type) flat | 
| 12026 | 223 | proof | 
| 224 | show "ALL x y::'a lift. x << y --> x = UU | x = y" | |
| 225 | by (simp add: less_lift) | |
| 226 | qed | |
| 227 | ||
| 228 | defaultsort pcpo | |
| 229 | ||
| 230 | ||
| 231 | text {*
 | |
| 232 | \medskip Two specific lemmas for the combination of LCF and HOL | |
| 233 | terms. | |
| 234 | *} | |
| 235 | ||
| 236 | lemma cont_Rep_CFun_app: "[|cont g; cont f|] ==> cont(%x. ((f x)$(g x)) s)" | |
| 237 | apply (rule cont2cont_CF1L) | |
| 238 | apply (tactic "resolve_tac cont_lemmas1 1")+ | |
| 239 | apply auto | |
| 240 | done | |
| 241 | ||
| 242 | lemma cont_Rep_CFun_app_app: "[|cont g; cont f|] ==> cont(%x. ((f x)$(g x)) s t)" | |
| 243 | apply (rule cont2cont_CF1L) | |
| 244 | apply (erule cont_Rep_CFun_app) | |
| 245 | apply assumption | |
| 246 | done | |
| 2640 | 247 | |
| 12026 | 248 | text {* Continuity of if-then-else. *}
 | 
| 249 | ||
| 250 | lemma cont_if: "[| cont f1; cont f2 |] ==> cont (%x. if b then f1 x else f2 x)" | |
| 251 | by (cases b) simp_all | |
| 252 | ||
| 253 | ||
| 254 | subsection {* Continuity Proofs for flift1, flift2, if *}
 | |
| 255 | ||
| 256 | text {* Need the instance of @{text flat}. *}
 | |
| 257 | ||
| 258 | lemma cont_flift1_arg: "cont (lift_case UU f)" | |
| 259 |   -- {* @{text flift1} is continuous in its argument itself. *}
 | |
| 260 | apply (rule flatdom_strict2cont) | |
| 261 | apply simp | |
| 262 | done | |
| 263 | ||
| 264 | lemma cont_flift1_not_arg: "!!f. [| !! a. cont (%y. (f y) a) |] ==> | |
| 265 | cont (%y. lift_case UU (f y))" | |
| 266 |   -- {* @{text flift1} is continuous in a variable that occurs only
 | |
| 267 |     in the @{text Def} branch. *}
 | |
| 268 | apply (rule cont2cont_CF1L_rev) | |
| 269 | apply (intro strip) | |
| 270 | apply (case_tac y) | |
| 271 | apply simp | |
| 272 | apply simp | |
| 273 | done | |
| 274 | ||
| 275 | lemma cont_flift1_arg_and_not_arg: "!!f. [| !! a. cont (%y. (f y) a); cont g|] ==> | |
| 276 | cont (%y. lift_case UU (f y) (g y))" | |
| 277 |   -- {* @{text flift1} is continuous in a variable that occurs either
 | |
| 278 |     in the @{text Def} branch or in the argument. *}
 | |
| 279 | apply (rule_tac tt = g in cont2cont_app) | |
| 280 | apply (rule cont_flift1_not_arg) | |
| 281 | apply auto | |
| 282 | apply (rule cont_flift1_arg) | |
| 283 | done | |
| 284 | ||
| 285 | lemma cont_flift2_arg: "cont (lift_case UU (%y. Def (f y)))" | |
| 286 |   -- {* @{text flift2} is continuous in its argument itself. *}
 | |
| 287 | apply (rule flatdom_strict2cont) | |
| 288 | apply simp | |
| 289 | done | |
| 290 | ||
| 291 | text {*
 | |
| 292 | \medskip Extension of cont_tac and installation of simplifier. | |
| 293 | *} | |
| 294 | ||
| 295 | lemma cont2cont_CF1L_rev2: "(!!y. cont (%x. c1 x y)) ==> cont c1" | |
| 296 | apply (rule cont2cont_CF1L_rev) | |
| 297 | apply simp | |
| 298 | done | |
| 299 | ||
| 300 | lemmas cont_lemmas_ext [simp] = | |
| 301 | cont_flift1_arg cont_flift2_arg | |
| 302 | cont_flift1_arg_and_not_arg cont2cont_CF1L_rev2 | |
| 303 | cont_Rep_CFun_app cont_Rep_CFun_app_app cont_if | |
| 304 | ||
| 305 | ML_setup {*
 | |
| 306 | val cont_lemmas2 = cont_lemmas1 @ thms "cont_lemmas_ext"; | |
| 307 | ||
| 308 | fun cont_tac i = resolve_tac cont_lemmas2 i; | |
| 309 | fun cont_tacR i = REPEAT (cont_tac i); | |
| 310 | ||
| 311 | local val flift1_def = thm "flift1_def" and flift2_def = thm "flift2_def" | |
| 312 | in fun cont_tacRs i = | |
| 313 | simp_tac (simpset() addsimps [flift1_def, flift2_def]) i THEN | |
| 314 | REPEAT (cont_tac i) | |
| 315 | end; | |
| 316 | ||
| 317 | simpset_ref() := simpset() addSolver | |
| 318 | (mk_solver "cont_tac" (K (DEPTH_SOLVE_1 o cont_tac))); | |
| 319 | *} | |
| 320 | ||
| 321 | ||
| 322 | subsection {* flift1, flift2 *}
 | |
| 323 | ||
| 324 | lemma flift1_Def [simp]: "flift1 f$(Def x) = (f x)" | |
| 325 | by (simp add: flift1_def) | |
| 326 | ||
| 327 | lemma flift2_Def [simp]: "flift2 f$(Def x) = Def (f x)" | |
| 328 | by (simp add: flift2_def) | |
| 329 | ||
| 330 | lemma flift1_UU [simp]: "flift1 f$UU = UU" | |
| 331 | by (simp add: flift1_def) | |
| 332 | ||
| 333 | lemma flift2_UU [simp]: "flift2 f$UU = UU" | |
| 334 | by (simp add: flift2_def) | |
| 335 | ||
| 336 | lemma flift2_nUU [simp]: "x~=UU ==> (flift2 f)$x~=UU" | |
| 337 | by (tactic "def_tac 1") | |
| 2640 | 338 | |
| 339 | end |