| author | wenzelm |
| Sat, 13 May 2006 02:51:37 +0200 | |
| changeset 19628 | de019ddcd89e |
| parent 17233 | 41eee2e7b465 |
| child 19741 | f65265d71426 |
| permissions | -rw-r--r-- |
| 3071 | 1 |
(* Title: HOL/IOA/meta_theory/Asig.thy |
2 |
ID: $Id$ |
|
| 12218 | 3 |
Author: Olaf Müller, Tobias Nipkow & Konrad Slind |
| 3071 | 4 |
*) |
5 |
||
| 17233 | 6 |
header {* Action signatures *}
|
| 3071 | 7 |
|
| 17233 | 8 |
theory Asig |
9 |
imports Main |
|
10 |
begin |
|
| 3071 | 11 |
|
| 17233 | 12 |
types |
13 |
'a signature = "('a set * 'a set * 'a set)"
|
|
| 3071 | 14 |
|
15 |
consts |
|
| 17233 | 16 |
actions :: "'action signature => 'action set" |
17 |
inputs :: "'action signature => 'action set" |
|
18 |
outputs :: "'action signature => 'action set" |
|
19 |
internals :: "'action signature => 'action set" |
|
20 |
externals :: "'action signature => 'action set" |
|
21 |
locals :: "'action signature => 'action set" |
|
| 3071 | 22 |
is_asig ::"'action signature => bool" |
23 |
mk_ext_asig ::"'action signature => 'action signature" |
|
24 |
||
25 |
defs |
|
26 |
||
| 17233 | 27 |
asig_inputs_def: "inputs == fst" |
28 |
asig_outputs_def: "outputs == (fst o snd)" |
|
29 |
asig_internals_def: "internals == (snd o snd)" |
|
| 3071 | 30 |
|
| 17233 | 31 |
actions_def: |
| 3071 | 32 |
"actions(asig) == (inputs(asig) Un outputs(asig) Un internals(asig))" |
33 |
||
| 17233 | 34 |
externals_def: |
| 3071 | 35 |
"externals(asig) == (inputs(asig) Un outputs(asig))" |
36 |
||
| 17233 | 37 |
locals_def: |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
38 |
"locals asig == ((internals asig) Un (outputs asig))" |
|
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
39 |
|
| 17233 | 40 |
is_asig_def: |
41 |
"is_asig(triple) == |
|
42 |
((inputs(triple) Int outputs(triple) = {}) &
|
|
43 |
(outputs(triple) Int internals(triple) = {}) &
|
|
| 3071 | 44 |
(inputs(triple) Int internals(triple) = {}))"
|
45 |
||
46 |
||
| 17233 | 47 |
mk_ext_asig_def: |
| 3071 | 48 |
"mk_ext_asig(triple) == (inputs(triple), outputs(triple), {})"
|
49 |
||
| 17233 | 50 |
ML {* use_legacy_bindings (the_context ()) *}
|
| 3071 | 51 |
|
| 17233 | 52 |
end |