3071
|
1 |
(* Title: HOLCF/IOA/meta_theory/ShortExecutions.thy
|
3275
|
2 |
ID: $Id$
|
3071
|
3 |
Author: Olaf M"uller
|
|
4 |
Copyright 1997 TU Muenchen
|
|
5 |
|
3275
|
6 |
Some properties about (Cut ex), defined as follows:
|
3071
|
7 |
|
3275
|
8 |
For every execution ex there is another shorter execution (Cut ex)
|
3071
|
9 |
that has the same trace as ex, but its schedule ends with an external action.
|
|
10 |
|
|
11 |
*)
|
|
12 |
|
|
13 |
|
|
14 |
ShortExecutions = Traces +
|
|
15 |
|
|
16 |
consts
|
|
17 |
|
3275
|
18 |
LastActExtex ::"('a,'s)ioa => ('a,'s) pairs => bool"
|
|
19 |
LastActExtsch ::"('a,'s)ioa => 'a Seq => bool"
|
3071
|
20 |
|
3275
|
21 |
Cut ::"('a => bool) => 'a Seq => 'a Seq"
|
|
22 |
|
|
23 |
oraclebuild ::"('a => bool) => 'a Seq -> 'a Seq -> 'a Seq"
|
3071
|
24 |
|
|
25 |
defs
|
|
26 |
|
|
27 |
LastActExtsch_def
|
3847
|
28 |
"LastActExtsch A sch == (Cut (%x. x: ext A) sch = sch)"
|
3275
|
29 |
|
|
30 |
LastActExtex_def
|
|
31 |
"LastActExtex A ex == LastActExtsch A (filter_act`ex)"
|
|
32 |
|
|
33 |
Cut_def
|
|
34 |
"Cut P s == oraclebuild P`s`(Filter P`s)"
|
|
35 |
|
|
36 |
oraclebuild_def
|
|
37 |
"oraclebuild P == (fix`(LAM h s t. case t of
|
|
38 |
nil => nil
|
|
39 |
| x##xs =>
|
|
40 |
(case x of
|
|
41 |
Undef => UU
|
|
42 |
| Def y => (Takewhile (%x.~P x)`s)
|
|
43 |
@@ (y>>(h`(TL`(Dropwhile (%x.~ P x)`s))`xs))
|
|
44 |
)
|
|
45 |
))"
|
|
46 |
|
|
47 |
|
3071
|
48 |
|
|
49 |
|
|
50 |
rules
|
|
51 |
|
3361
|
52 |
Cut_prefixcl_Finite
|
|
53 |
"Finite s ==> (? y. s = Cut P s @@ y)"
|
3071
|
54 |
|
|
55 |
LastActExtsmall1
|
3275
|
56 |
"LastActExtsch A sch ==> LastActExtsch A (TL`(Dropwhile P`sch))"
|
3071
|
57 |
|
|
58 |
LastActExtsmall2
|
3275
|
59 |
"[| Finite sch1; LastActExtsch A (sch1 @@ sch2) |] ==> LastActExtsch A sch2"
|
3071
|
60 |
|
|
61 |
end
|
|
62 |
|