| author | wenzelm | 
| Thu, 16 Jan 1997 14:53:37 +0100 | |
| changeset 2509 | 0a7169d89b7a | 
| parent 2439 | e73cb5924261 | 
| permissions | -rw-r--r-- | 
| 1274 | 1 | (* | 
| 2 | ID: $Id$ | |
| 1461 | 3 | Author: Franz Regensburger | 
| 1274 | 4 | Copyright 1993 Technische Universitaet Muenchen | 
| 5 | ||
| 6 | Lemmas for stream.thy | |
| 7 | *) | |
| 8 | ||
| 9 | open Stream; | |
| 10 | ||
| 11 | (* ------------------------------------------------------------------------*) | |
| 12 | (* The isomorphisms stream_rep_iso and stream_abs_iso are strict *) | |
| 13 | (* ------------------------------------------------------------------------*) | |
| 14 | ||
| 15 | val stream_iso_strict= stream_rep_iso RS (stream_abs_iso RS | |
| 1461 | 16 | (allI RSN (2,allI RS iso_strict))); | 
| 1274 | 17 | |
| 18 | val stream_rews = [stream_iso_strict RS conjunct1, | |
| 1461 | 19 | stream_iso_strict RS conjunct2]; | 
| 1274 | 20 | |
| 21 | (* ------------------------------------------------------------------------*) | |
| 22 | (* Properties of stream_copy *) | |
| 23 | (* ------------------------------------------------------------------------*) | |
| 24 | ||
| 25 | fun prover defs thm = prove_goalw Stream.thy defs thm | |
| 26 | (fn prems => | |
| 1461 | 27 | [ | 
| 28 | (cut_facts_tac prems 1), | |
| 29 | (asm_simp_tac (!simpset addsimps | |
| 30 | (stream_rews @ [stream_abs_iso,stream_rep_iso])) 1) | |
| 31 | ]); | |
| 1274 | 32 | |
| 33 | val stream_copy = | |
| 1461 | 34 | [ | 
| 35 | prover [stream_copy_def] "stream_copy`f`UU=UU", | |
| 36 | prover [stream_copy_def,scons_def] | |
| 37 | "x~=UU ==> stream_copy`f`(scons`x`xs)= scons`x`(f`xs)" | |
| 38 | ]; | |
| 1274 | 39 | |
| 40 | val stream_rews = stream_copy @ stream_rews; | |
| 41 | ||
| 42 | (* ------------------------------------------------------------------------*) | |
| 43 | (* Exhaustion and elimination for streams *) | |
| 44 | (* ------------------------------------------------------------------------*) | |
| 45 | ||
| 46 | qed_goalw "Exh_stream" Stream.thy [scons_def] | |
| 1461 | 47 | "s = UU | (? x xs. x~=UU & s = scons`x`xs)" | 
| 1274 | 48 | (fn prems => | 
| 1461 | 49 | [ | 
| 50 | (Simp_tac 1), | |
| 51 | (rtac (stream_rep_iso RS subst) 1), | |
| 52 |         (res_inst_tac [("p","stream_rep`s")] sprodE 1),
 | |
| 53 | (asm_simp_tac (!simpset addsimps stream_rews) 1), | |
| 54 | (Asm_simp_tac 1), | |
| 2277 
9174de6c7143
moved Lift*.* to Up*.*, renaming of all constans and theorems concerned,
 oheimb parents: 
2033diff
changeset | 55 |         (res_inst_tac [("p","y")] upE1 1),
 | 
| 2421 | 56 | (contr_tac 1), | 
| 1461 | 57 | (rtac disjI2 1), | 
| 58 | (rtac exI 1), | |
| 59 | (etac conjI 1), | |
| 2421 | 60 | (rtac exI 1), | 
| 1461 | 61 | (Asm_simp_tac 1) | 
| 62 | ]); | |
| 1274 | 63 | |
| 64 | qed_goal "streamE" Stream.thy | |
| 1461 | 65 | "[| s=UU ==> Q; !!x xs.[|s=scons`x`xs;x~=UU|]==>Q|]==>Q" | 
| 1274 | 66 | (fn prems => | 
| 1461 | 67 | [ | 
| 68 | (rtac (Exh_stream RS disjE) 1), | |
| 69 | (eresolve_tac prems 1), | |
| 70 | (etac exE 1), | |
| 71 | (etac exE 1), | |
| 72 | (resolve_tac prems 1), | |
| 73 | (fast_tac HOL_cs 1), | |
| 74 | (fast_tac HOL_cs 1) | |
| 75 | ]); | |
| 1274 | 76 | |
| 77 | (* ------------------------------------------------------------------------*) | |
| 78 | (* Properties of stream_when *) | |
| 79 | (* ------------------------------------------------------------------------*) | |
| 80 | ||
| 81 | fun prover defs thm = prove_goalw Stream.thy defs thm | |
| 82 | (fn prems => | |
| 1461 | 83 | [ | 
| 84 | (cut_facts_tac prems 1), | |
| 85 | (asm_simp_tac (!simpset addsimps | |
| 86 | (stream_rews @ [stream_abs_iso,stream_rep_iso])) 1) | |
| 87 | ]); | |
| 1274 | 88 | |
| 89 | ||
| 90 | val stream_when = [ | |
| 1461 | 91 | prover [stream_when_def] "stream_when`f`UU=UU", | 
| 92 | prover [stream_when_def,scons_def] | |
| 93 | "x~=UU ==> stream_when`f`(scons`x`xs)= f`x`xs" | |
| 94 | ]; | |
| 1274 | 95 | |
| 96 | val stream_rews = stream_when @ stream_rews; | |
| 97 | ||
| 98 | (* ------------------------------------------------------------------------*) | |
| 99 | (* Rewrites for discriminators and selectors *) | |
| 100 | (* ------------------------------------------------------------------------*) | |
| 101 | ||
| 102 | fun prover defs thm = prove_goalw Stream.thy defs thm | |
| 103 | (fn prems => | |
| 1461 | 104 | [ | 
| 105 | (simp_tac (!simpset addsimps stream_rews) 1) | |
| 106 | ]); | |
| 1274 | 107 | |
| 108 | val stream_discsel = [ | |
| 1461 | 109 | prover [is_scons_def] "is_scons`UU=UU", | 
| 110 | prover [shd_def] "shd`UU=UU", | |
| 111 | prover [stl_def] "stl`UU=UU" | |
| 112 | ]; | |
| 1274 | 113 | |
| 114 | fun prover defs thm = prove_goalw Stream.thy defs thm | |
| 115 | (fn prems => | |
| 1461 | 116 | [ | 
| 117 | (cut_facts_tac prems 1), | |
| 118 | (asm_simp_tac (!simpset addsimps stream_rews) 1) | |
| 119 | ]); | |
| 1274 | 120 | |
| 121 | val stream_discsel = [ | |
| 122 | prover [is_scons_def,shd_def,stl_def] "x~=UU ==> is_scons`(scons`x`xs)=TT", | |
| 123 | prover [is_scons_def,shd_def,stl_def] "x~=UU ==> shd`(scons`x`xs)=x", | |
| 124 | prover [is_scons_def,shd_def,stl_def] "x~=UU ==> stl`(scons`x`xs)=xs" | |
| 1461 | 125 | ] @ stream_discsel; | 
| 1274 | 126 | |
| 127 | val stream_rews = stream_discsel @ stream_rews; | |
| 128 | ||
| 129 | (* ------------------------------------------------------------------------*) | |
| 130 | (* Definedness and strictness *) | |
| 131 | (* ------------------------------------------------------------------------*) | |
| 132 | ||
| 133 | fun prover contr thm = prove_goal Stream.thy thm | |
| 134 | (fn prems => | |
| 1461 | 135 | [ | 
| 2439 | 136 |         (res_inst_tac [("P1",contr)] classical2 1),
 | 
| 1461 | 137 | (simp_tac (!simpset addsimps stream_rews) 1), | 
| 138 | (dtac sym 1), | |
| 139 | (Asm_simp_tac 1), | |
| 140 | (simp_tac (!simpset addsimps (prems @ stream_rews)) 1) | |
| 141 | ]); | |
| 1274 | 142 | |
| 143 | val stream_constrdef = [ | |
| 1461 | 144 | prover "is_scons`(UU::'a stream)~=UU" "x~=UU ==> scons`(x::'a)`xs~=UU" | 
| 145 | ]; | |
| 1274 | 146 | |
| 147 | fun prover defs thm = prove_goalw Stream.thy defs thm | |
| 148 | (fn prems => | |
| 1461 | 149 | [ | 
| 150 | (simp_tac (!simpset addsimps stream_rews) 1) | |
| 151 | ]); | |
| 1274 | 152 | |
| 153 | val stream_constrdef = [ | |
| 1461 | 154 | prover [scons_def] "scons`UU`xs=UU" | 
| 155 | ] @ stream_constrdef; | |
| 1274 | 156 | |
| 157 | val stream_rews = stream_constrdef @ stream_rews; | |
| 158 | ||
| 159 | ||
| 160 | (* ------------------------------------------------------------------------*) | |
| 161 | (* Distinctness wrt. << and = *) | |
| 162 | (* ------------------------------------------------------------------------*) | |
| 163 | ||
| 164 | ||
| 165 | (* ------------------------------------------------------------------------*) | |
| 166 | (* Invertibility *) | |
| 167 | (* ------------------------------------------------------------------------*) | |
| 168 | ||
| 169 | val stream_invert = | |
| 1461 | 170 | [ | 
| 1274 | 171 | prove_goal Stream.thy "[|x1~=UU; y1~=UU;\ | 
| 172 | \ scons`x1`x2 << scons`y1`y2|] ==> x1<< y1 & x2 << y2" | |
| 173 | (fn prems => | |
| 1461 | 174 | [ | 
| 175 | (cut_facts_tac prems 1), | |
| 176 | (rtac conjI 1), | |
| 2421 | 177 |         (dres_inst_tac [("fo","stream_when`(LAM x l.x)")] monofun_cfun_arg 1),
 | 
| 1461 | 178 | (etac box_less 1), | 
| 179 | (asm_simp_tac (!simpset addsimps stream_when) 1), | |
| 180 | (asm_simp_tac (!simpset addsimps stream_when) 1), | |
| 2421 | 181 |         (dres_inst_tac [("fo","stream_when`(LAM x l.l)")] monofun_cfun_arg 1),
 | 
| 1461 | 182 | (etac box_less 1), | 
| 183 | (asm_simp_tac (!simpset addsimps stream_when) 1), | |
| 184 | (asm_simp_tac (!simpset addsimps stream_when) 1) | |
| 185 | ]) | |
| 186 | ]; | |
| 1274 | 187 | |
| 188 | (* ------------------------------------------------------------------------*) | |
| 189 | (* Injectivity *) | |
| 190 | (* ------------------------------------------------------------------------*) | |
| 191 | ||
| 192 | val stream_inject = | |
| 1461 | 193 | [ | 
| 1274 | 194 | prove_goal Stream.thy "[|x1~=UU; y1~=UU;\ | 
| 195 | \ scons`x1`x2 = scons`y1`y2 |] ==> x1= y1 & x2 = y2" | |
| 196 | (fn prems => | |
| 1461 | 197 | [ | 
| 198 | (cut_facts_tac prems 1), | |
| 199 | (rtac conjI 1), | |
| 200 |         (dres_inst_tac [("f","stream_when`(LAM x l.x)")] cfun_arg_cong 1),
 | |
| 201 | (etac box_equals 1), | |
| 202 | (asm_simp_tac (!simpset addsimps stream_when) 1), | |
| 203 | (asm_simp_tac (!simpset addsimps stream_when) 1), | |
| 204 |         (dres_inst_tac [("f","stream_when`(LAM x l.l)")] cfun_arg_cong 1),
 | |
| 205 | (etac box_equals 1), | |
| 206 | (asm_simp_tac (!simpset addsimps stream_when) 1), | |
| 207 | (asm_simp_tac (!simpset addsimps stream_when) 1) | |
| 208 | ]) | |
| 209 | ]; | |
| 1274 | 210 | |
| 211 | (* ------------------------------------------------------------------------*) | |
| 212 | (* definedness for discriminators and selectors *) | |
| 213 | (* ------------------------------------------------------------------------*) | |
| 214 | ||
| 215 | fun prover thm = prove_goal Stream.thy thm | |
| 216 | (fn prems => | |
| 1461 | 217 | [ | 
| 218 | (cut_facts_tac prems 1), | |
| 219 | (rtac streamE 1), | |
| 220 | (contr_tac 1), | |
| 221 | (REPEAT (asm_simp_tac (!simpset addsimps stream_discsel) 1)) | |
| 222 | ]); | |
| 1274 | 223 | |
| 224 | val stream_discsel_def = | |
| 1461 | 225 | [ | 
| 226 | prover "s~=UU ==> is_scons`s ~= UU", | |
| 227 | prover "s~=UU ==> shd`s ~=UU" | |
| 228 | ]; | |
| 1274 | 229 | |
| 230 | val stream_rews = stream_discsel_def @ stream_rews; | |
| 231 | ||
| 232 | ||
| 233 | (* ------------------------------------------------------------------------*) | |
| 234 | (* Properties stream_take *) | |
| 235 | (* ------------------------------------------------------------------------*) | |
| 236 | ||
| 237 | val stream_take = | |
| 1461 | 238 | [ | 
| 1274 | 239 | prove_goalw Stream.thy [stream_take_def] "stream_take n`UU = UU" | 
| 240 | (fn prems => | |
| 1461 | 241 | [ | 
| 242 |         (res_inst_tac [("n","n")] natE 1),
 | |
| 243 | (Asm_simp_tac 1), | |
| 244 | (Asm_simp_tac 1), | |
| 245 | (simp_tac (!simpset addsimps stream_rews) 1) | |
| 246 | ]), | |
| 1274 | 247 | prove_goalw Stream.thy [stream_take_def] "stream_take 0`xs=UU" | 
| 248 | (fn prems => | |
| 1461 | 249 | [ | 
| 250 | (Asm_simp_tac 1) | |
| 251 | ])]; | |
| 1274 | 252 | |
| 253 | fun prover thm = prove_goalw Stream.thy [stream_take_def] thm | |
| 254 | (fn prems => | |
| 1461 | 255 | [ | 
| 256 | (cut_facts_tac prems 1), | |
| 257 | (Simp_tac 1), | |
| 258 | (asm_simp_tac (!simpset addsimps stream_rews) 1) | |
| 259 | ]); | |
| 1274 | 260 | |
| 261 | val stream_take = [ | |
| 262 | prover | |
| 263 | "x~=UU ==> stream_take (Suc n)`(scons`x`xs) = scons`x`(stream_take n`xs)" | |
| 1461 | 264 | ] @ stream_take; | 
| 1274 | 265 | |
| 266 | val stream_rews = stream_take @ stream_rews; | |
| 267 | ||
| 268 | (* ------------------------------------------------------------------------*) | |
| 269 | (* enhance the simplifier *) | |
| 270 | (* ------------------------------------------------------------------------*) | |
| 271 | ||
| 272 | qed_goal "stream_copy2" Stream.thy | |
| 273 | "stream_copy`f`(scons`x`xs) = scons`x`(f`xs)" | |
| 274 | (fn prems => | |
| 1461 | 275 | [ | 
| 1675 | 276 | (case_tac "x=UU" 1), | 
| 1461 | 277 | (asm_simp_tac (!simpset addsimps stream_rews) 1), | 
| 278 | (asm_simp_tac (!simpset addsimps stream_rews) 1) | |
| 279 | ]); | |
| 1274 | 280 | |
| 281 | qed_goal "shd2" Stream.thy "shd`(scons`x`xs) = x" | |
| 282 | (fn prems => | |
| 1461 | 283 | [ | 
| 1675 | 284 | (case_tac "x=UU" 1), | 
| 1461 | 285 | (asm_simp_tac (!simpset addsimps stream_rews) 1), | 
| 286 | (asm_simp_tac (!simpset addsimps stream_rews) 1) | |
| 287 | ]); | |
| 1274 | 288 | |
| 289 | qed_goal "stream_take2" Stream.thy | |
| 290 | "stream_take (Suc n)`(scons`x`xs) = scons`x`(stream_take n`xs)" | |
| 291 | (fn prems => | |
| 1461 | 292 | [ | 
| 1675 | 293 | (case_tac "x=UU" 1), | 
| 1461 | 294 | (asm_simp_tac (!simpset addsimps stream_rews) 1), | 
| 295 | (asm_simp_tac (!simpset addsimps stream_rews) 1) | |
| 296 | ]); | |
| 1274 | 297 | |
| 298 | val stream_rews = [stream_iso_strict RS conjunct1, | |
| 1461 | 299 | stream_iso_strict RS conjunct2, | 
| 1274 | 300 | hd stream_copy, stream_copy2] | 
| 301 | @ stream_when | |
| 302 | @ [hd stream_discsel,shd2] @ (tl (tl stream_discsel)) | |
| 303 | @ stream_constrdef | |
| 304 | @ stream_discsel_def | |
| 305 | @ [ stream_take2] @ (tl stream_take); | |
| 306 | ||
| 307 | ||
| 308 | (* ------------------------------------------------------------------------*) | |
| 309 | (* take lemma for streams *) | |
| 310 | (* ------------------------------------------------------------------------*) | |
| 311 | ||
| 312 | fun prover reach defs thm = prove_goalw Stream.thy defs thm | |
| 313 | (fn prems => | |
| 2033 | 314 | [ | 
| 315 |         (res_inst_tac [("t","s1")] (reach RS subst) 1),
 | |
| 316 |         (res_inst_tac [("t","s2")] (reach RS subst) 1),
 | |
| 317 | (stac fix_def2 1), | |
| 318 | (stac contlub_cfun_fun 1), | |
| 319 | (rtac is_chain_iterate 1), | |
| 320 | (stac contlub_cfun_fun 1), | |
| 321 | (rtac is_chain_iterate 1), | |
| 322 | (rtac lub_equal 1), | |
| 323 | (rtac (is_chain_iterate RS ch2ch_fappL) 1), | |
| 324 | (rtac (is_chain_iterate RS ch2ch_fappL) 1), | |
| 325 | (rtac allI 1), | |
| 2421 | 326 | (resolve_tac prems 1) | 
| 2033 | 327 | ]); | 
| 1274 | 328 | |
| 329 | val stream_take_lemma = prover stream_reach [stream_take_def] | |
| 1461 | 330 | "(!!n.stream_take n`s1 = stream_take n`s2) ==> s1=s2"; | 
| 1274 | 331 | |
| 332 | ||
| 333 | qed_goal "stream_reach2" Stream.thy "lub(range(%i.stream_take i`s))=s" | |
| 334 | (fn prems => | |
| 1461 | 335 | [ | 
| 336 |         (res_inst_tac [("t","s")] (stream_reach RS subst) 1),
 | |
| 2033 | 337 | (stac fix_def2 1), | 
| 1461 | 338 | (rewtac stream_take_def), | 
| 2033 | 339 | (stac contlub_cfun_fun 1), | 
| 1461 | 340 | (rtac is_chain_iterate 1), | 
| 341 | (rtac refl 1) | |
| 342 | ]); | |
| 1274 | 343 | |
| 344 | (* ------------------------------------------------------------------------*) | |
| 345 | (* Co -induction for streams *) | |
| 346 | (* ------------------------------------------------------------------------*) | |
| 347 | ||
| 348 | qed_goalw "stream_coind_lemma" Stream.thy [stream_bisim_def] | |
| 349 | "stream_bisim R ==> ! p q. R p q --> stream_take n`p = stream_take n`q" | |
| 350 | (fn prems => | |
| 1461 | 351 | [ | 
| 352 | (cut_facts_tac prems 1), | |
| 353 | (nat_ind_tac "n" 1), | |
| 354 | (simp_tac (!simpset addsimps stream_rews) 1), | |
| 355 | (strip_tac 1), | |
| 356 | ((etac allE 1) THEN (etac allE 1) THEN (etac (mp RS disjE) 1)), | |
| 357 | (atac 1), | |
| 358 | (asm_simp_tac (!simpset addsimps stream_rews) 1), | |
| 359 | (etac exE 1), | |
| 360 | (etac exE 1), | |
| 361 | (etac exE 1), | |
| 362 | (asm_simp_tac (!simpset addsimps stream_rews) 1), | |
| 363 | (REPEAT (etac conjE 1)), | |
| 364 | (rtac cfun_arg_cong 1), | |
| 365 | (fast_tac HOL_cs 1) | |
| 366 | ]); | |
| 1274 | 367 | |
| 368 | qed_goal "stream_coind" Stream.thy "[|stream_bisim R ;R p q|] ==> p = q" | |
| 369 | (fn prems => | |
| 1461 | 370 | [ | 
| 371 | (rtac stream_take_lemma 1), | |
| 372 | (rtac (stream_coind_lemma RS spec RS spec RS mp) 1), | |
| 373 | (resolve_tac prems 1), | |
| 374 | (resolve_tac prems 1) | |
| 375 | ]); | |
| 1274 | 376 | |
| 377 | (* ------------------------------------------------------------------------*) | |
| 378 | (* structural induction for admissible predicates *) | |
| 379 | (* ------------------------------------------------------------------------*) | |
| 380 | ||
| 381 | qed_goal "stream_finite_ind" Stream.thy | |
| 382 | "[|P(UU);\ | |
| 383 | \ !! x s1.[|x~=UU;P(s1)|] ==> P(scons`x`s1)\ | |
| 384 | \ |] ==> !s.P(stream_take n`s)" | |
| 385 | (fn prems => | |
| 1461 | 386 | [ | 
| 387 | (nat_ind_tac "n" 1), | |
| 388 | (simp_tac (!simpset addsimps stream_rews) 1), | |
| 389 | (resolve_tac prems 1), | |
| 390 | (rtac allI 1), | |
| 391 |         (res_inst_tac [("s","s")] streamE 1),
 | |
| 392 | (asm_simp_tac (!simpset addsimps stream_rews) 1), | |
| 393 | (resolve_tac prems 1), | |
| 394 | (asm_simp_tac (!simpset addsimps stream_rews) 1), | |
| 395 | (resolve_tac prems 1), | |
| 396 | (atac 1), | |
| 397 | (etac spec 1) | |
| 398 | ]); | |
| 1274 | 399 | |
| 400 | qed_goalw "stream_finite_ind2" Stream.thy [stream_finite_def] | |
| 401 | "(!!n.P(stream_take n`s)) ==> stream_finite(s) -->P(s)" | |
| 402 | (fn prems => | |
| 1461 | 403 | [ | 
| 404 | (strip_tac 1), | |
| 405 | (etac exE 1), | |
| 406 | (etac subst 1), | |
| 407 | (resolve_tac prems 1) | |
| 408 | ]); | |
| 1274 | 409 | |
| 410 | qed_goal "stream_finite_ind3" Stream.thy | |
| 411 | "[|P(UU);\ | |
| 412 | \ !! x s1.[|x~=UU;P(s1)|] ==> P(scons`x`s1)\ | |
| 413 | \ |] ==> stream_finite(s) --> P(s)" | |
| 414 | (fn prems => | |
| 1461 | 415 | [ | 
| 416 | (rtac stream_finite_ind2 1), | |
| 417 | (rtac (stream_finite_ind RS spec) 1), | |
| 418 | (REPEAT (resolve_tac prems 1)), | |
| 419 | (REPEAT (atac 1)) | |
| 420 | ]); | |
| 1274 | 421 | |
| 422 | (* prove induction using definition of admissibility | |
| 423 | stream_reach rsp. stream_reach2 | |
| 424 | and finite induction stream_finite_ind *) | |
| 425 | ||
| 426 | qed_goal "stream_ind" Stream.thy | |
| 427 | "[|adm(P);\ | |
| 428 | \ P(UU);\ | |
| 429 | \ !! x s1.[|x~=UU;P(s1)|] ==> P(scons`x`s1)\ | |
| 430 | \ |] ==> P(s)" | |
| 431 | (fn prems => | |
| 1461 | 432 | [ | 
| 433 | (rtac (stream_reach2 RS subst) 1), | |
| 434 | (rtac (adm_def2 RS iffD1 RS spec RS mp RS mp) 1), | |
| 435 | (resolve_tac prems 1), | |
| 436 | (SELECT_GOAL (rewtac stream_take_def) 1), | |
| 437 | (rtac ch2ch_fappL 1), | |
| 438 | (rtac is_chain_iterate 1), | |
| 439 | (rtac allI 1), | |
| 440 | (rtac (stream_finite_ind RS spec) 1), | |
| 441 | (REPEAT (resolve_tac prems 1)), | |
| 442 | (REPEAT (atac 1)) | |
| 443 | ]); | |
| 1274 | 444 | |
| 445 | (* prove induction with usual LCF-Method using fixed point induction *) | |
| 446 | qed_goal "stream_ind" Stream.thy | |
| 447 | "[|adm(P);\ | |
| 448 | \ P(UU);\ | |
| 449 | \ !! x s1.[|x~=UU;P(s1)|] ==> P(scons`x`s1)\ | |
| 450 | \ |] ==> P(s)" | |
| 451 | (fn prems => | |
| 1461 | 452 | [ | 
| 453 | (rtac (stream_reach RS subst) 1), | |
| 454 |         (res_inst_tac [("x","s")] spec 1),
 | |
| 455 | (rtac wfix_ind 1), | |
| 456 | (rtac adm_impl_admw 1), | |
| 457 | (REPEAT (resolve_tac adm_thms 1)), | |
| 458 | (rtac adm_subst 1), | |
| 459 | (cont_tacR 1), | |
| 460 | (resolve_tac prems 1), | |
| 461 | (rtac allI 1), | |
| 462 | (rtac (rewrite_rule [stream_take_def] stream_finite_ind) 1), | |
| 463 | (REPEAT (resolve_tac prems 1)), | |
| 464 | (REPEAT (atac 1)) | |
| 465 | ]); | |
| 1274 | 466 | |
| 467 | ||
| 468 | (* ------------------------------------------------------------------------*) | |
| 469 | (* simplify use of Co-induction *) | |
| 470 | (* ------------------------------------------------------------------------*) | |
| 471 | ||
| 472 | qed_goal "surjectiv_scons" Stream.thy "scons`(shd`s)`(stl`s)=s" | |
| 473 | (fn prems => | |
| 1461 | 474 | [ | 
| 475 |         (res_inst_tac [("s","s")] streamE 1),
 | |
| 476 | (asm_simp_tac (!simpset addsimps stream_rews) 1), | |
| 477 | (asm_simp_tac (!simpset addsimps stream_rews) 1) | |
| 478 | ]); | |
| 1274 | 479 | |
| 480 | ||
| 481 | qed_goalw "stream_coind_lemma2" Stream.thy [stream_bisim_def] | |
| 482 | "!s1 s2. R s1 s2 --> shd`s1 = shd`s2 & R (stl`s1) (stl`s2) ==> stream_bisim R" | |
| 483 | (fn prems => | |
| 1461 | 484 | [ | 
| 485 | (cut_facts_tac prems 1), | |
| 486 | (strip_tac 1), | |
| 487 | (etac allE 1), | |
| 488 | (etac allE 1), | |
| 489 | (dtac mp 1), | |
| 490 | (atac 1), | |
| 491 | (etac conjE 1), | |
| 1675 | 492 | (case_tac "s1 = UU & s2 = UU" 1), | 
| 1461 | 493 | (rtac disjI1 1), | 
| 494 | (fast_tac HOL_cs 1), | |
| 495 | (rtac disjI2 1), | |
| 496 | (rtac disjE 1), | |
| 2421 | 497 | (etac (de_Morgan_conj RS subst) 1), | 
| 1461 | 498 |         (res_inst_tac [("x","shd`s1")] exI 1),
 | 
| 499 |         (res_inst_tac [("x","stl`s1")] exI 1),
 | |
| 500 |         (res_inst_tac [("x","stl`s2")] exI 1),
 | |
| 501 | (rtac conjI 1), | |
| 502 | (eresolve_tac stream_discsel_def 1), | |
| 503 | (asm_simp_tac (!simpset addsimps stream_rews addsimps [surjectiv_scons]) 1), | |
| 504 |         (eres_inst_tac [("s","shd`s1"),("t","shd`s2")] subst 1),
 | |
| 505 | (simp_tac (!simpset addsimps stream_rews addsimps [surjectiv_scons]) 1), | |
| 506 |         (res_inst_tac [("x","shd`s2")] exI 1),
 | |
| 507 |         (res_inst_tac [("x","stl`s1")] exI 1),
 | |
| 508 |         (res_inst_tac [("x","stl`s2")] exI 1),
 | |
| 509 | (rtac conjI 1), | |
| 510 | (eresolve_tac stream_discsel_def 1), | |
| 511 | (asm_simp_tac (!simpset addsimps stream_rews addsimps [surjectiv_scons]) 1), | |
| 512 |         (res_inst_tac [("s","shd`s1"),("t","shd`s2")] ssubst 1),
 | |
| 513 | (etac sym 1), | |
| 514 | (simp_tac (!simpset addsimps stream_rews addsimps [surjectiv_scons]) 1) | |
| 515 | ]); | |
| 1274 | 516 | |
| 517 | ||
| 518 | (* ------------------------------------------------------------------------*) | |
| 519 | (* theorems about finite and infinite streams *) | |
| 520 | (* ------------------------------------------------------------------------*) | |
| 521 | ||
| 522 | (* ----------------------------------------------------------------------- *) | |
| 523 | (* 2 lemmas about stream_finite *) | |
| 524 | (* ----------------------------------------------------------------------- *) | |
| 525 | ||
| 526 | qed_goalw "stream_finite_UU" Stream.thy [stream_finite_def] | |
| 1461 | 527 | "stream_finite(UU)" | 
| 1274 | 528 | (fn prems => | 
| 1461 | 529 | [ | 
| 530 | (rtac exI 1), | |
| 531 | (simp_tac (!simpset addsimps stream_rews) 1) | |
| 532 | ]); | |
| 1274 | 533 | |
| 534 | qed_goal "inf_stream_not_UU" Stream.thy "~stream_finite(s) ==> s ~= UU" | |
| 535 | (fn prems => | |
| 1461 | 536 | [ | 
| 537 | (cut_facts_tac prems 1), | |
| 538 | (etac swap 1), | |
| 539 | (dtac notnotD 1), | |
| 540 | (hyp_subst_tac 1), | |
| 541 | (rtac stream_finite_UU 1) | |
| 542 | ]); | |
| 1274 | 543 | |
| 544 | (* ----------------------------------------------------------------------- *) | |
| 545 | (* a lemma about shd *) | |
| 546 | (* ----------------------------------------------------------------------- *) | |
| 547 | ||
| 548 | qed_goal "stream_shd_lemma1" Stream.thy "shd`s=UU --> s=UU" | |
| 549 | (fn prems => | |
| 1461 | 550 | [ | 
| 551 |         (res_inst_tac [("s","s")] streamE 1),
 | |
| 552 | (asm_simp_tac (!simpset addsimps stream_rews) 1), | |
| 553 | (hyp_subst_tac 1), | |
| 554 | (asm_simp_tac (!simpset addsimps stream_rews) 1) | |
| 555 | ]); | |
| 1274 | 556 | |
| 557 | ||
| 558 | (* ----------------------------------------------------------------------- *) | |
| 559 | (* lemmas about stream_take *) | |
| 560 | (* ----------------------------------------------------------------------- *) | |
| 561 | ||
| 562 | qed_goal "stream_take_lemma1" Stream.thy | |
| 563 | "!x xs.x~=UU --> \ | |
| 564 | \ stream_take (Suc n)`(scons`x`xs) = scons`x`xs --> stream_take n`xs=xs" | |
| 565 | (fn prems => | |
| 1461 | 566 | [ | 
| 567 | (rtac allI 1), | |
| 568 | (rtac allI 1), | |
| 569 | (rtac impI 1), | |
| 570 | (asm_simp_tac (!simpset addsimps stream_rews) 1), | |
| 571 | (strip_tac 1), | |
| 572 | (rtac ((hd stream_inject) RS conjunct2) 1), | |
| 573 | (atac 1), | |
| 574 | (atac 1), | |
| 575 | (atac 1) | |
| 576 | ]); | |
| 1274 | 577 | |
| 578 | ||
| 579 | qed_goal "stream_take_lemma2" Stream.thy | |
| 580 | "! s2. stream_take n`s2 = s2 --> stream_take (Suc n)`s2=s2" | |
| 581 | (fn prems => | |
| 1461 | 582 | [ | 
| 583 | (nat_ind_tac "n" 1), | |
| 584 | (simp_tac (!simpset addsimps stream_rews) 1), | |
| 2421 | 585 | (strip_tac 1), | 
| 1461 | 586 |         (res_inst_tac [("s","s2")] streamE 1),
 | 
| 2421 | 587 | (asm_simp_tac (!simpset addsimps stream_rews) 1), | 
| 588 | (hyp_subst_tac 1), | |
| 589 | (rotate_tac ~1 1), | |
| 590 | (asm_full_simp_tac (!simpset addsimps stream_rews) 1), | |
| 1461 | 591 | (subgoal_tac "stream_take n1`xs = xs" 1), | 
| 2421 | 592 | (rtac ((hd stream_inject) RS conjunct2) 2), | 
| 593 | (atac 4), | |
| 594 | (atac 2), | |
| 595 | (atac 2), | |
| 1461 | 596 | (rtac cfun_arg_cong 1), | 
| 597 | (fast_tac HOL_cs 1) | |
| 598 | ]); | |
| 1274 | 599 | |
| 600 | qed_goal "stream_take_lemma3" Stream.thy | |
| 601 | "!x xs.x~=UU --> \ | |
| 602 | \ stream_take n`(scons`x`xs) = scons`x`xs --> stream_take n`xs=xs" | |
| 603 | (fn prems => | |
| 1461 | 604 | [ | 
| 605 | (nat_ind_tac "n" 1), | |
| 606 | (asm_simp_tac (!simpset addsimps stream_rews) 1), | |
| 607 | (strip_tac 1 ), | |
| 608 |         (res_inst_tac [("P","scons`x`xs=UU")] notE 1),
 | |
| 609 | (eresolve_tac stream_constrdef 1), | |
| 610 | (etac sym 1), | |
| 611 | (strip_tac 1 ), | |
| 612 | (rtac (stream_take_lemma2 RS spec RS mp) 1), | |
| 613 |         (res_inst_tac [("x1.1","x")] ((hd stream_inject) RS conjunct2) 1),
 | |
| 614 | (atac 1), | |
| 615 | (atac 1), | |
| 616 | (etac (stream_take2 RS subst) 1) | |
| 617 | ]); | |
| 1274 | 618 | |
| 619 | qed_goal "stream_take_lemma4" Stream.thy | |
| 620 | "!x xs.\ | |
| 621 | \stream_take n`xs=xs --> stream_take (Suc n)`(scons`x`xs) = scons`x`xs" | |
| 622 | (fn prems => | |
| 1461 | 623 | [ | 
| 624 | (nat_ind_tac "n" 1), | |
| 625 | (simp_tac (!simpset addsimps stream_rews) 1), | |
| 626 | (simp_tac (!simpset addsimps stream_rews) 1) | |
| 627 | ]); | |
| 1274 | 628 | |
| 629 | (* ---- *) | |
| 630 | ||
| 631 | qed_goal "stream_take_lemma5" Stream.thy | |
| 632 | "!s. stream_take n`s=s --> iterate n stl s=UU" | |
| 633 | (fn prems => | |
| 1461 | 634 | [ | 
| 635 | (nat_ind_tac "n" 1), | |
| 636 | (Simp_tac 1), | |
| 637 | (simp_tac (!simpset addsimps stream_rews) 1), | |
| 638 | (strip_tac 1), | |
| 639 |         (res_inst_tac [("s","s")] streamE 1),
 | |
| 640 | (hyp_subst_tac 1), | |
| 2033 | 641 | (stac iterate_Suc2 1), | 
| 1461 | 642 | (asm_simp_tac (!simpset addsimps stream_rews) 1), | 
| 2033 | 643 | (stac iterate_Suc2 1), | 
| 1461 | 644 | (asm_simp_tac (!simpset addsimps stream_rews) 1), | 
| 645 | (etac allE 1), | |
| 646 | (etac mp 1), | |
| 647 | (hyp_subst_tac 1), | |
| 648 | (etac (stream_take_lemma1 RS spec RS spec RS mp RS mp) 1), | |
| 649 | (atac 1) | |
| 650 | ]); | |
| 1274 | 651 | |
| 652 | qed_goal "stream_take_lemma6" Stream.thy | |
| 653 | "!s.iterate n stl s =UU --> stream_take n`s=s" | |
| 654 | (fn prems => | |
| 1461 | 655 | [ | 
| 656 | (nat_ind_tac "n" 1), | |
| 657 | (Simp_tac 1), | |
| 658 | (strip_tac 1), | |
| 659 | (asm_simp_tac (!simpset addsimps stream_rews) 1), | |
| 660 | (rtac allI 1), | |
| 661 |         (res_inst_tac [("s","s")] streamE 1),
 | |
| 662 | (hyp_subst_tac 1), | |
| 663 | (asm_simp_tac (!simpset addsimps stream_rews) 1), | |
| 664 | (hyp_subst_tac 1), | |
| 2033 | 665 | (stac iterate_Suc2 1), | 
| 1461 | 666 | (asm_simp_tac (!simpset addsimps stream_rews) 1) | 
| 667 | ]); | |
| 1274 | 668 | |
| 669 | qed_goal "stream_take_lemma7" Stream.thy | |
| 670 | "(iterate n stl s=UU) = (stream_take n`s=s)" | |
| 671 | (fn prems => | |
| 1461 | 672 | [ | 
| 673 | (rtac iffI 1), | |
| 674 | (etac (stream_take_lemma6 RS spec RS mp) 1), | |
| 675 | (etac (stream_take_lemma5 RS spec RS mp) 1) | |
| 676 | ]); | |
| 1274 | 677 | |
| 678 | ||
| 679 | qed_goal "stream_take_lemma8" Stream.thy | |
| 680 | "[|adm(P); !n. ? m. n < m & P (stream_take m`s)|] ==> P(s)" | |
| 681 | (fn prems => | |
| 1461 | 682 | [ | 
| 683 | (cut_facts_tac prems 1), | |
| 684 | (rtac (stream_reach2 RS subst) 1), | |
| 1991 | 685 | (rtac adm_lemma11 1), | 
| 1461 | 686 | (atac 1), | 
| 687 | (atac 2), | |
| 688 | (rewtac stream_take_def), | |
| 689 | (rtac ch2ch_fappL 1), | |
| 690 | (rtac is_chain_iterate 1) | |
| 691 | ]); | |
| 1274 | 692 | |
| 693 | (* ----------------------------------------------------------------------- *) | |
| 694 | (* lemmas stream_finite *) | |
| 695 | (* ----------------------------------------------------------------------- *) | |
| 696 | ||
| 697 | qed_goalw "stream_finite_lemma1" Stream.thy [stream_finite_def] | |
| 698 | "stream_finite(xs) ==> stream_finite(scons`x`xs)" | |
| 699 | (fn prems => | |
| 1461 | 700 | [ | 
| 701 | (cut_facts_tac prems 1), | |
| 702 | (etac exE 1), | |
| 703 | (rtac exI 1), | |
| 704 | (etac (stream_take_lemma4 RS spec RS spec RS mp) 1) | |
| 705 | ]); | |
| 1274 | 706 | |
| 707 | qed_goalw "stream_finite_lemma2" Stream.thy [stream_finite_def] | |
| 708 | "[|x~=UU; stream_finite(scons`x`xs)|] ==> stream_finite(xs)" | |
| 709 | (fn prems => | |
| 1461 | 710 | [ | 
| 711 | (cut_facts_tac prems 1), | |
| 712 | (etac exE 1), | |
| 713 | (rtac exI 1), | |
| 714 | (etac (stream_take_lemma3 RS spec RS spec RS mp RS mp) 1), | |
| 715 | (atac 1) | |
| 716 | ]); | |
| 1274 | 717 | |
| 718 | qed_goal "stream_finite_lemma3" Stream.thy | |
| 719 | "x~=UU ==> stream_finite(scons`x`xs) = stream_finite(xs)" | |
| 720 | (fn prems => | |
| 1461 | 721 | [ | 
| 722 | (cut_facts_tac prems 1), | |
| 723 | (rtac iffI 1), | |
| 724 | (etac stream_finite_lemma2 1), | |
| 725 | (atac 1), | |
| 726 | (etac stream_finite_lemma1 1) | |
| 727 | ]); | |
| 1274 | 728 | |
| 729 | ||
| 730 | qed_goalw "stream_finite_lemma5" Stream.thy [stream_finite_def] | |
| 731 | "(!n. s1 << s2 --> stream_take n`s2 = s2 --> stream_finite(s1))\ | |
| 732 | \=(s1 << s2 --> stream_finite(s2) --> stream_finite(s1))" | |
| 733 | (fn prems => | |
| 1461 | 734 | [ | 
| 735 | (rtac iffI 1), | |
| 736 | (fast_tac HOL_cs 1), | |
| 737 | (fast_tac HOL_cs 1) | |
| 738 | ]); | |
| 1274 | 739 | |
| 740 | qed_goal "stream_finite_lemma6" Stream.thy | |
| 741 | "!s1 s2. s1 << s2 --> stream_take n`s2 = s2 --> stream_finite(s1)" | |
| 742 | (fn prems => | |
| 1461 | 743 | [ | 
| 744 | (nat_ind_tac "n" 1), | |
| 745 | (simp_tac (!simpset addsimps stream_rews) 1), | |
| 746 | (strip_tac 1 ), | |
| 747 | (hyp_subst_tac 1), | |
| 748 | (dtac UU_I 1), | |
| 749 | (hyp_subst_tac 1), | |
| 750 | (rtac stream_finite_UU 1), | |
| 751 | (rtac allI 1), | |
| 752 | (rtac allI 1), | |
| 753 |         (res_inst_tac [("s","s1")] streamE 1),
 | |
| 754 | (hyp_subst_tac 1), | |
| 755 | (strip_tac 1 ), | |
| 756 | (rtac stream_finite_UU 1), | |
| 757 | (hyp_subst_tac 1), | |
| 758 |         (res_inst_tac [("s","s2")] streamE 1),
 | |
| 759 | (hyp_subst_tac 1), | |
| 760 | (strip_tac 1 ), | |
| 761 | (dtac UU_I 1), | |
| 762 | (asm_simp_tac(!simpset addsimps (stream_rews @ [stream_finite_UU])) 1), | |
| 763 | (hyp_subst_tac 1), | |
| 764 | (simp_tac (!simpset addsimps stream_rews) 1), | |
| 765 | (strip_tac 1 ), | |
| 766 | (rtac stream_finite_lemma1 1), | |
| 767 | (subgoal_tac "xs << xsa" 1), | |
| 768 | (subgoal_tac "stream_take n1`xsa = xsa" 1), | |
| 769 | (fast_tac HOL_cs 1), | |
| 770 |         (res_inst_tac  [("x1.1","xa"),("y1.1","xa")] 
 | |
| 1274 | 771 | ((hd stream_inject) RS conjunct2) 1), | 
| 1461 | 772 | (atac 1), | 
| 773 | (atac 1), | |
| 774 | (atac 1), | |
| 775 |         (res_inst_tac [("x1.1","x"),("y1.1","xa")]
 | |
| 776 | ((hd stream_invert) RS conjunct2) 1), | |
| 777 | (atac 1), | |
| 778 | (atac 1), | |
| 779 | (atac 1) | |
| 780 | ]); | |
| 1274 | 781 | |
| 782 | qed_goal "stream_finite_lemma7" Stream.thy | |
| 783 | "s1 << s2 --> stream_finite(s2) --> stream_finite(s1)" | |
| 784 | (fn prems => | |
| 1461 | 785 | [ | 
| 786 | (rtac (stream_finite_lemma5 RS iffD1) 1), | |
| 787 | (rtac allI 1), | |
| 788 | (rtac (stream_finite_lemma6 RS spec RS spec) 1) | |
| 789 | ]); | |
| 1274 | 790 | |
| 791 | qed_goalw "stream_finite_lemma8" Stream.thy [stream_finite_def] | |
| 792 | "stream_finite(s) = (? n. iterate n stl s = UU)" | |
| 793 | (fn prems => | |
| 1461 | 794 | [ | 
| 795 | (simp_tac (!simpset addsimps [stream_take_lemma7]) 1) | |
| 796 | ]); | |
| 1274 | 797 | |
| 798 | ||
| 799 | (* ----------------------------------------------------------------------- *) | |
| 800 | (* admissibility of ~stream_finite *) | |
| 801 | (* ----------------------------------------------------------------------- *) | |
| 802 | ||
| 803 | qed_goalw "adm_not_stream_finite" Stream.thy [adm_def] | |
| 804 | "adm(%s. ~ stream_finite(s))" | |
| 805 | (fn prems => | |
| 1461 | 806 | [ | 
| 807 | (strip_tac 1 ), | |
| 2439 | 808 |         (res_inst_tac [("P1","!i. ~ stream_finite(Y(i))")] classical2 1),
 | 
| 1461 | 809 | (atac 2), | 
| 810 | (subgoal_tac "!i.stream_finite(Y(i))" 1), | |
| 811 | (fast_tac HOL_cs 1), | |
| 812 | (rtac allI 1), | |
| 813 | (rtac (stream_finite_lemma7 RS mp RS mp) 1), | |
| 814 | (etac is_ub_thelub 1), | |
| 815 | (atac 1) | |
| 816 | ]); | |
| 1274 | 817 | |
| 818 | (* ----------------------------------------------------------------------- *) | |
| 819 | (* alternative prove for admissibility of ~stream_finite *) | |
| 820 | (* show that stream_finite(s) = (? n. iterate n stl s = UU) *) | |
| 821 | (* and prove adm. of ~(? n. iterate n stl s = UU) *) | |
| 822 | (* proof uses theorems stream_take_lemma5-7; stream_finite_lemma8 *) | |
| 823 | (* ----------------------------------------------------------------------- *) | |
| 824 | ||
| 825 | ||
| 826 | qed_goal "adm_not_stream_finite" Stream.thy "adm(%s. ~ stream_finite(s))" | |
| 827 | (fn prems => | |
| 1461 | 828 | [ | 
| 829 | (subgoal_tac "(!s.(~stream_finite(s))=(!n.iterate n stl s ~=UU))" 1), | |
| 830 | (etac (adm_cong RS iffD2)1), | |
| 831 | (REPEAT(resolve_tac adm_thms 1)), | |
| 832 | (rtac cont_iterate2 1), | |
| 833 | (rtac allI 1), | |
| 2033 | 834 | (stac stream_finite_lemma8 1), | 
| 1461 | 835 | (fast_tac HOL_cs 1) | 
| 836 | ]); | |
| 1274 | 837 | |
| 838 |