| author | kleing |
| Mon, 16 May 2005 09:35:05 +0200 | |
| changeset 15964 | f2074e12d1d4 |
| parent 14981 | e73f8140af78 |
| child 17233 | 41eee2e7b465 |
| permissions | -rw-r--r-- |
| 3071 | 1 |
(* Title: HOLCF/IOA/meta_theory/RefMappings.thy |
| 3275 | 2 |
ID: $Id$ |
| 12218 | 3 |
Author: Olaf Müller |
| 3071 | 4 |
|
| 12218 | 5 |
Refinement Mappings in HOLCF/IOA. |
| 3071 | 6 |
*) |
7 |
||
8 |
RefMappings = Traces + |
|
9 |
||
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12218
diff
changeset
|
10 |
default type |
| 3071 | 11 |
|
12 |
consts |
|
|
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3433
diff
changeset
|
13 |
|
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
14 |
move ::"[('a,'s)ioa,('a,'s)pairs,'s,'a,'s] => bool"
|
| 3071 | 15 |
is_ref_map ::"[('s1=>'s2),('a,'s1)ioa,('a,'s2)ioa] => bool"
|
16 |
is_weak_ref_map ::"[('s1=>'s2),('a,'s1)ioa,('a,'s2)ioa] => bool"
|
|
17 |
||
|
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3433
diff
changeset
|
18 |
|
| 3071 | 19 |
defs |
20 |
||
21 |
||
22 |
move_def |
|
23 |
"move ioa ex s a t == |
|
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
24 |
(is_exec_frag ioa (s,ex) & Finite ex & |
|
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
25 |
laststate (s,ex)=t & |
| 10835 | 26 |
mk_trace ioa$ex = (if a:ext(ioa) then a>>nil else nil))" |
| 3071 | 27 |
|
28 |
is_ref_map_def |
|
29 |
"is_ref_map f C A == |
|
30 |
(!s:starts_of(C). f(s):starts_of(A)) & |
|
31 |
(!s t a. reachable C s & |
|
32 |
s -a--C-> t |
|
33 |
--> (? ex. move A ex (f s) a (f t)))" |
|
34 |
||
35 |
is_weak_ref_map_def |
|
36 |
"is_weak_ref_map f C A == |
|
37 |
(!s:starts_of(C). f(s):starts_of(A)) & |
|
38 |
(!s t a. reachable C s & |
|
39 |
s -a--C-> t |
|
40 |
--> (if a:ext(C) |
|
41 |
then (f s) -a--A-> (f t) |
|
42 |
else (f s)=(f t)))" |
|
43 |
||
|
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
3433
diff
changeset
|
44 |
|
| 3071 | 45 |
end |