| author | wenzelm | 
| Sat, 29 Sep 2012 21:24:20 +0200 | |
| changeset 49663 | b84fafaea4bb | 
| parent 42174 | d0be2722ce9f | 
| child 58889 | 5b7a9633cfa8 | 
| permissions | -rw-r--r-- | 
| 4530 | 1  | 
(* Title: HOL/IOA/Asig.thy  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
2  | 
Author: Tobias Nipkow & Konrad Slind  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
3  | 
Copyright 1994 TU Muenchen  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
4  | 
*)  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
5  | 
|
| 17288 | 6  | 
header {* Action signatures *}
 | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
7  | 
|
| 17288 | 8  | 
theory Asig  | 
9  | 
imports Main  | 
|
10  | 
begin  | 
|
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
11  | 
|
| 42174 | 12  | 
type_synonym  | 
| 17288 | 13  | 
  'a signature = "('a set * 'a set * 'a set)"
 | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
14  | 
|
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
15  | 
consts  | 
| 17288 | 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  | 
||
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
22  | 
is_asig ::"'action signature => bool"  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
23  | 
mk_ext_asig ::"'action signature => 'action signature"  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
24  | 
|
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
25  | 
|
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
26  | 
defs  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
27  | 
|
| 17288 | 28  | 
asig_inputs_def: "inputs == fst"  | 
29  | 
asig_outputs_def: "outputs == (fst o snd)"  | 
|
30  | 
asig_internals_def: "internals == (snd o snd)"  | 
|
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
31  | 
|
| 17288 | 32  | 
actions_def:  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
33  | 
"actions(asig) == (inputs(asig) Un outputs(asig) Un internals(asig))"  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
34  | 
|
| 17288 | 35  | 
externals_def:  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
36  | 
"externals(asig) == (inputs(asig) Un outputs(asig))"  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
37  | 
|
| 17288 | 38  | 
is_asig_def:  | 
39  | 
"is_asig(triple) ==  | 
|
40  | 
      ((inputs(triple) Int outputs(triple) = {})    &
 | 
|
41  | 
       (outputs(triple) Int internals(triple) = {}) &
 | 
|
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
42  | 
       (inputs(triple) Int internals(triple) = {}))"
 | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
43  | 
|
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
44  | 
|
| 17288 | 45  | 
mk_ext_asig_def:  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
46  | 
  "mk_ext_asig(triple) == (inputs(triple), outputs(triple), {})"
 | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
47  | 
|
| 19801 | 48  | 
|
49  | 
lemmas asig_projections = asig_inputs_def asig_outputs_def asig_internals_def  | 
|
50  | 
||
51  | 
lemma int_and_ext_is_act: "[| a~:internals(S) ;a~:externals(S)|] ==> a~:actions(S)"  | 
|
52  | 
apply (simp add: externals_def actions_def)  | 
|
53  | 
done  | 
|
54  | 
||
55  | 
lemma ext_is_act: "[|a:externals(S)|] ==> a:actions(S)"  | 
|
56  | 
apply (simp add: externals_def actions_def)  | 
|
57  | 
done  | 
|
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
58  | 
|
| 17288 | 59  | 
end  |