author | wenzelm |
Sun, 25 Oct 1998 12:33:27 +0100 | |
changeset 5769 | 6a422b22ba02 |
parent 3433 | 2de17c994071 |
child 7661 | 8c3190b173aa |
permissions | -rw-r--r-- |
3071 | 1 |
(* Title: HOL/IOA/meta_theory/Asig.thy |
2 |
ID: $Id$ |
|
3275 | 3 |
Author: Olaf Mueller, Tobias Nipkow & Konrad Slind |
4 |
Copyright 1994, 1996 TU Muenchen |
|
3071 | 5 |
|
6 |
Action signatures |
|
7 |
*) |
|
8 |
||
9 |
Asig = Arith + |
|
10 |
||
11 |
types |
|
12 |
||
13 |
'a signature = "('a set * 'a set * 'a set)" |
|
14 |
||
15 |
consts |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
16 |
actions,inputs,outputs,internals,externals,locals |
3071 | 17 |
::"'action signature => 'action set" |
18 |
is_asig ::"'action signature => bool" |
|
19 |
mk_ext_asig ::"'action signature => 'action signature" |
|
20 |
||
21 |
||
22 |
defs |
|
23 |
||
24 |
asig_inputs_def "inputs == fst" |
|
25 |
asig_outputs_def "outputs == (fst o snd)" |
|
26 |
asig_internals_def "internals == (snd o snd)" |
|
27 |
||
28 |
actions_def |
|
29 |
"actions(asig) == (inputs(asig) Un outputs(asig) Un internals(asig))" |
|
30 |
||
31 |
externals_def |
|
32 |
"externals(asig) == (inputs(asig) Un outputs(asig))" |
|
33 |
||
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
34 |
locals_def |
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
35 |
"locals asig == ((internals asig) Un (outputs asig))" |
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
36 |
|
3071 | 37 |
is_asig_def |
38 |
"is_asig(triple) == |
|
39 |
((inputs(triple) Int outputs(triple) = {}) & |
|
40 |
(outputs(triple) Int internals(triple) = {}) & |
|
41 |
(inputs(triple) Int internals(triple) = {}))" |
|
42 |
||
43 |
||
44 |
mk_ext_asig_def |
|
45 |
"mk_ext_asig(triple) == (inputs(triple), outputs(triple), {})" |
|
46 |
||
47 |
||
48 |
end |