| author | wenzelm | 
| Thu, 13 Jul 2000 23:09:25 +0200 | |
| changeset 9311 | ab5b24cbaa16 | 
| parent 9248 | e1dee89de037 | 
| child 10230 | 5eb935d6cc69 | 
| permissions | -rw-r--r-- | 
| 2278 | 1 | (* Title: HOLCF/Up3.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Franz Regensburger | |
| 4 | Copyright 1993 Technische Universitaet Muenchen | |
| 5 | ||
| 9169 | 6 | Class instance of  ('a)u for class pcpo
 | 
| 2278 | 7 | *) | 
| 8 | ||
| 2640 | 9 | (* for compatibility with old HOLCF-Version *) | 
| 9169 | 10 | Goal "UU = Abs_Up(Inl ())"; | 
| 11 | by (simp_tac (HOL_ss addsimps [UU_def,UU_up_def]) 1); | |
| 12 | qed "inst_up_pcpo"; | |
| 2640 | 13 | |
| 2278 | 14 | (* -------------------------------------------------------------------------*) | 
| 15 | (* some lemmas restated for class pcpo *) | |
| 16 | (* ------------------------------------------------------------------------ *) | |
| 17 | ||
| 9169 | 18 | Goal "~ Iup(x) << UU"; | 
| 19 | by (stac inst_up_pcpo 1); | |
| 20 | by (rtac less_up2b 1); | |
| 21 | qed "less_up3b"; | |
| 2278 | 22 | |
| 9169 | 23 | Goal "Iup(x) ~= UU"; | 
| 24 | by (stac inst_up_pcpo 1); | |
| 25 | by (rtac defined_Iup 1); | |
| 26 | qed "defined_Iup2"; | |
| 2278 | 27 | |
| 28 | (* ------------------------------------------------------------------------ *) | |
| 29 | (* continuity for Iup *) | |
| 30 | (* ------------------------------------------------------------------------ *) | |
| 31 | ||
| 9169 | 32 | Goal "contlub(Iup)"; | 
| 33 | by (rtac contlubI 1); | |
| 34 | by (strip_tac 1); | |
| 35 | by (rtac trans 1); | |
| 36 | by (rtac (thelub_up1a RS sym) 2); | |
| 37 | by (fast_tac HOL_cs 3); | |
| 38 | by (etac (monofun_Iup RS ch2ch_monofun) 2); | |
| 39 | by (res_inst_tac [("f","Iup")] arg_cong  1);
 | |
| 40 | by (rtac lub_equal 1); | |
| 41 | by (atac 1); | |
| 42 | by (rtac (monofun_Ifup2 RS ch2ch_monofun) 1); | |
| 43 | by (etac (monofun_Iup RS ch2ch_monofun) 1); | |
| 44 | by (asm_simp_tac Up0_ss 1); | |
| 45 | qed "contlub_Iup"; | |
| 2278 | 46 | |
| 9169 | 47 | Goal "cont(Iup)"; | 
| 48 | by (rtac monocontlub2cont 1); | |
| 49 | by (rtac monofun_Iup 1); | |
| 50 | by (rtac contlub_Iup 1); | |
| 51 | qed "cont_Iup"; | |
| 2278 | 52 | |
| 53 | ||
| 54 | (* ------------------------------------------------------------------------ *) | |
| 55 | (* continuity for Ifup *) | |
| 56 | (* ------------------------------------------------------------------------ *) | |
| 57 | ||
| 9169 | 58 | Goal "contlub(Ifup)"; | 
| 59 | by (rtac contlubI 1); | |
| 60 | by (strip_tac 1); | |
| 61 | by (rtac trans 1); | |
| 62 | by (rtac (thelub_fun RS sym) 2); | |
| 63 | by (etac (monofun_Ifup1 RS ch2ch_monofun) 2); | |
| 64 | by (rtac ext 1); | |
| 65 | by (res_inst_tac [("p","x")] upE 1);
 | |
| 66 | by (asm_simp_tac Up0_ss 1); | |
| 67 | by (rtac (lub_const RS thelubI RS sym) 1); | |
| 68 | by (asm_simp_tac Up0_ss 1); | |
| 69 | by (etac contlub_cfun_fun 1); | |
| 70 | qed "contlub_Ifup1"; | |
| 2278 | 71 | |
| 72 | ||
| 8161 | 73 | Goal "contlub(Ifup(f))"; | 
| 74 | by (rtac contlubI 1); | |
| 75 | by (strip_tac 1); | |
| 76 | by (rtac disjE 1); | |
| 77 | by (stac thelub_up1a 2); | |
| 78 | by (atac 2); | |
| 79 | by (atac 2); | |
| 80 | by (asm_simp_tac Up0_ss 2); | |
| 81 | by (stac thelub_up1b 3); | |
| 82 | by (atac 3); | |
| 83 | by (atac 3); | |
| 84 | by (fast_tac HOL_cs 1); | |
| 85 | by (asm_simp_tac Up0_ss 2); | |
| 86 | by (rtac (chain_UU_I_inverse RS sym) 2); | |
| 87 | by (rtac allI 2); | |
| 88 | by  (res_inst_tac [("p","Y(i)")] upE 2);
 | |
| 89 | by (asm_simp_tac Up0_ss 2); | |
| 90 | by (rtac notE 2); | |
| 91 | by (dtac spec 2); | |
| 92 | by (etac spec 2); | |
| 93 | by (atac 2); | |
| 94 | by (stac contlub_cfun_arg 1); | |
| 95 | by (etac (monofun_Ifup2 RS ch2ch_monofun) 1); | |
| 96 | by (rtac lub_equal2 1); | |
| 97 | by (rtac (monofun_Rep_CFun2 RS ch2ch_monofun) 2); | |
| 98 | by (etac (monofun_Ifup2 RS ch2ch_monofun) 2); | |
| 99 | by (etac (monofun_Ifup2 RS ch2ch_monofun) 2); | |
| 100 | by (rtac (chain_mono2 RS exE) 1); | |
| 101 | by (atac 2); | |
| 102 | by (etac exE 1); | |
| 103 | by (etac exE 1); | |
| 104 | by (rtac exI 1); | |
| 105 | by  (res_inst_tac [("s","Iup(x)"),("t","Y(i)")] ssubst 1);
 | |
| 106 | by (atac 1); | |
| 107 | by (rtac defined_Iup2 1); | |
| 108 | by (rtac exI 1); | |
| 109 | by (strip_tac 1); | |
| 110 | by (res_inst_tac [("p","Y(i)")] upE 1);
 | |
| 111 | by (asm_simp_tac Up0_ss 2); | |
| 112 | by (res_inst_tac [("P","Y(i) = UU")] notE 1);
 | |
| 113 | by (fast_tac HOL_cs 1); | |
| 114 | by (stac inst_up_pcpo 1); | |
| 115 | by (atac 1); | |
| 116 | qed "contlub_Ifup2"; | |
| 2278 | 117 | |
| 9169 | 118 | Goal "cont(Ifup)"; | 
| 119 | by (rtac monocontlub2cont 1); | |
| 120 | by (rtac monofun_Ifup1 1); | |
| 121 | by (rtac contlub_Ifup1 1); | |
| 122 | qed "cont_Ifup1"; | |
| 2278 | 123 | |
| 9169 | 124 | Goal "cont(Ifup(f))"; | 
| 125 | by (rtac monocontlub2cont 1); | |
| 126 | by (rtac monofun_Ifup2 1); | |
| 127 | by (rtac contlub_Ifup2 1); | |
| 128 | qed "cont_Ifup2"; | |
| 2278 | 129 | |
| 130 | ||
| 131 | (* ------------------------------------------------------------------------ *) | |
| 132 | (* continuous versions of lemmas for ('a)u                                  *)
 | |
| 133 | (* ------------------------------------------------------------------------ *) | |
| 134 | ||
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 135 | Goalw [up_def] "z = UU | (? x. z = up`x)"; | 
| 9245 | 136 | by (simp_tac (Up0_ss addsimps [cont_Iup]) 1); | 
| 137 | by (stac inst_up_pcpo 1); | |
| 138 | by (rtac Exh_Up 1); | |
| 139 | qed "Exh_Up1"; | |
| 2278 | 140 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 141 | Goalw [up_def] "up`x=up`y ==> x=y"; | 
| 9245 | 142 | by (rtac inject_Iup 1); | 
| 143 | by (etac box_equals 1); | |
| 144 | by (simp_tac (Up0_ss addsimps [cont_Iup]) 1); | |
| 145 | by (simp_tac (Up0_ss addsimps [cont_Iup]) 1); | |
| 146 | qed "inject_up"; | |
| 2278 | 147 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 148 | Goalw [up_def] " up`x ~= UU"; | 
| 9245 | 149 | by (simp_tac (Up0_ss addsimps [cont_Iup]) 1); | 
| 150 | by (rtac defined_Iup2 1); | |
| 151 | qed "defined_up"; | |
| 2278 | 152 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 153 | val prems = Goalw [up_def] | 
| 9245 | 154 | "[| p=UU ==> Q; !!x. p=up`x==>Q|] ==>Q"; | 
| 155 | by (rtac upE 1); | |
| 156 | by (resolve_tac prems 1); | |
| 157 | by (etac (inst_up_pcpo RS ssubst) 1); | |
| 158 | by (resolve_tac (tl prems) 1); | |
| 159 | by (asm_simp_tac (Up0_ss addsimps [cont_Iup]) 1); | |
| 160 | qed "upE1"; | |
| 2278 | 161 | |
| 4098 | 162 | val tac = (simp_tac (simpset() addsimps [cont_Iup,cont_Ifup1, | 
| 2566 | 163 | cont_Ifup2,cont2cont_CF1L]) 1); | 
| 2278 | 164 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 165 | Goalw [up_def,fup_def] "fup`f`UU=UU"; | 
| 9245 | 166 | by (stac inst_up_pcpo 1); | 
| 167 | by (stac beta_cfun 1); | |
| 168 | by tac; | |
| 169 | by (stac beta_cfun 1); | |
| 170 | by tac; | |
| 171 | by (simp_tac (Up0_ss addsimps [cont_Iup,cont_Ifup1,cont_Ifup2]) 1); | |
| 172 | qed "fup1"; | |
| 2278 | 173 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 174 | Goalw [up_def,fup_def] "fup`f`(up`x)=f`x"; | 
| 9245 | 175 | by (stac beta_cfun 1); | 
| 176 | by (rtac cont_Iup 1); | |
| 177 | by (stac beta_cfun 1); | |
| 178 | by tac; | |
| 179 | by (stac beta_cfun 1); | |
| 180 | by (rtac cont_Ifup2 1); | |
| 181 | by (simp_tac (Up0_ss addsimps [cont_Iup,cont_Ifup1,cont_Ifup2]) 1); | |
| 182 | qed "fup2"; | |
| 2278 | 183 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 184 | Goalw [up_def,fup_def] "~ up`x << UU"; | 
| 9245 | 185 | by (simp_tac (Up0_ss addsimps [cont_Iup]) 1); | 
| 186 | by (rtac less_up3b 1); | |
| 187 | qed "less_up4b"; | |
| 2278 | 188 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 189 | Goalw [up_def,fup_def] | 
| 9245 | 190 | "(up`x << up`y) = (x<<y)"; | 
| 191 | by (simp_tac (Up0_ss addsimps [cont_Iup]) 1); | |
| 192 | by (rtac less_up2c 1); | |
| 193 | qed "less_up4c"; | |
| 2278 | 194 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 195 | Goalw [up_def,fup_def] | 
| 4721 
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
 oheimb parents: 
4098diff
changeset | 196 | "[| chain(Y); ? i x. Y(i) = up`x |] ==>\ | 
| 9245 | 197 | \ lub(range(Y)) = up`(lub(range(%i. fup`(LAM x. x)`(Y i))))"; | 
| 198 | by (stac beta_cfun 1); | |
| 199 | by tac; | |
| 200 | by (stac beta_cfun 1); | |
| 201 | by tac; | |
| 202 | by (stac (beta_cfun RS ext) 1); | |
| 203 | by tac; | |
| 204 | by (rtac thelub_up1a 1); | |
| 205 | by (atac 1); | |
| 206 | by (etac exE 1); | |
| 207 | by (etac exE 1); | |
| 208 | by (rtac exI 1); | |
| 209 | by (rtac exI 1); | |
| 210 | by (etac box_equals 1); | |
| 211 | by (rtac refl 1); | |
| 212 | by (simp_tac (Up0_ss addsimps [cont_Iup]) 1); | |
| 213 | qed "thelub_up2a"; | |
| 2278 | 214 | |
| 215 | ||
| 216 | ||
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 217 | Goalw [up_def,fup_def] | 
| 9245 | 218 | "[| chain(Y); ! i x. Y(i) ~= up`x |] ==> lub(range(Y)) = UU"; | 
| 219 | by (stac inst_up_pcpo 1); | |
| 220 | by (rtac thelub_up1b 1); | |
| 221 | by (atac 1); | |
| 222 | by (strip_tac 1); | |
| 223 | by (dtac spec 1); | |
| 224 | by (dtac spec 1); | |
| 225 | by (rtac swap 1); | |
| 226 | by (atac 1); | |
| 227 | by (dtac notnotD 1); | |
| 228 | by (etac box_equals 1); | |
| 229 | by (rtac refl 1); | |
| 230 | by (simp_tac (Up0_ss addsimps [cont_Iup]) 1); | |
| 231 | qed "thelub_up2b"; | |
| 2278 | 232 | |
| 233 | ||
| 9169 | 234 | Goal "(? x. z = up`x) = (z~=UU)"; | 
| 235 | by (rtac iffI 1); | |
| 236 | by (etac exE 1); | |
| 237 | by (hyp_subst_tac 1); | |
| 238 | by (rtac defined_up 1); | |
| 239 | by (res_inst_tac [("p","z")] upE1 1);
 | |
| 240 | by (etac notE 1); | |
| 241 | by (atac 1); | |
| 242 | by (etac exI 1); | |
| 243 | qed "up_lemma2"; | |
| 2278 | 244 | |
| 245 | ||
| 9169 | 246 | Goal "[| chain(Y); lub(range(Y)) = up`x |] ==> ? i x. Y(i) = up`x"; | 
| 247 | by (rtac exE 1); | |
| 248 | by (rtac chain_UU_I_inverse2 1); | |
| 249 | by (rtac (up_lemma2 RS iffD1) 1); | |
| 250 | by (etac exI 1); | |
| 251 | by (rtac exI 1); | |
| 252 | by (rtac (up_lemma2 RS iffD2) 1); | |
| 253 | by (atac 1); | |
| 254 | qed "thelub_up2a_rev"; | |
| 2278 | 255 | |
| 9169 | 256 | Goal "[| chain(Y); lub(range(Y)) = UU |] ==> ! i x. Y(i) ~= up`x"; | 
| 257 | by (rtac allI 1); | |
| 258 | by (rtac (not_ex RS iffD1) 1); | |
| 259 | by (rtac contrapos 1); | |
| 260 | by (etac (up_lemma2 RS iffD1) 2); | |
| 261 | by (fast_tac (HOL_cs addSDs [chain_UU_I RS spec]) 1); | |
| 262 | qed "thelub_up2b_rev"; | |
| 2278 | 263 | |
| 264 | ||
| 9169 | 265 | Goal "chain(Y) ==> lub(range(Y)) = UU | \ | 
| 266 | \ lub(range(Y)) = up`(lub(range(%i. fup`(LAM x. x)`(Y i))))"; | |
| 267 | by (rtac disjE 1); | |
| 268 | by (rtac disjI1 2); | |
| 269 | by (rtac thelub_up2b 2); | |
| 270 | by (atac 2); | |
| 271 | by (atac 2); | |
| 272 | by (rtac disjI2 2); | |
| 273 | by (rtac thelub_up2a 2); | |
| 274 | by (atac 2); | |
| 275 | by (atac 2); | |
| 276 | by (fast_tac HOL_cs 1); | |
| 277 | qed "thelub_up3"; | |
| 2278 | 278 | |
| 9169 | 279 | Goal "fup`up`x=x"; | 
| 280 | by (res_inst_tac [("p","x")] upE1 1);
 | |
| 281 | by (asm_simp_tac ((simpset_of Cfun3.thy) addsimps [fup1,fup2]) 1); | |
| 282 | by (asm_simp_tac ((simpset_of Cfun3.thy) addsimps [fup1,fup2]) 1); | |
| 283 | qed "fup3"; | |
| 2278 | 284 | |
| 285 | (* ------------------------------------------------------------------------ *) | |
| 286 | (* install simplifier for ('a)u                                             *)
 | |
| 287 | (* ------------------------------------------------------------------------ *) | |
| 288 | ||
| 3327 | 289 | Addsimps [fup1,fup2,defined_up]; |