| author | blanchet | 
| Thu, 06 Jun 2013 21:18:39 +0200 | |
| changeset 52324 | 095c88b93e8d | 
| parent 44389 | a3b5fdfb04a3 | 
| child 72835 | 66ca5016b008 | 
| permissions | -rw-r--r-- | 
| 42151 | 1  | 
(* Title: HOL/HOLCF/IOA/ABP/Check.ML  | 
| 
19360
 
f47412f922ab
converted Müller to Mueller to make smlnj 110.58 work
 
kleing 
parents: 
17239 
diff
changeset
 | 
2  | 
Author: Olaf Mueller  | 
| 
3072
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
3  | 
|
| 12218 | 4  | 
The Model Checker.  | 
| 
3072
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
6  | 
|
| 17239 | 7  | 
structure Check =  | 
8  | 
struct  | 
|
| 
3072
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
9  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
10  | 
(* ----------------------------------------------------------------  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
11  | 
P r o t o t y p e M o d e l C h e c k e r  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
12  | 
----------------------------------------------------------------*)  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
13  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
14  | 
fun check(extacts,intacts,string_of_a,startsI,string_of_s,  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
15  | 
nexts,hom,transA,startsS) =  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
16  | 
let fun check_s(s,unchecked,checked) =  | 
| 33245 | 17  | 
let fun check_sa a unchecked =  | 
18  | 
let fun check_sas t unchecked =  | 
|
| 
36692
 
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
 
haftmann 
parents: 
35174 
diff
changeset
 | 
19  | 
(if member (op =) extacts a then  | 
| 
3072
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
20  | 
(if transA(hom s,a,hom t) then ( )  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
21  | 
                           else (writeln("Error: Mapping of Externals!");
 | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
22  | 
string_of_s s; writeln"";  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
23  | 
string_of_a a; writeln"";  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
24  | 
string_of_s t;writeln"";writeln"" ))  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
25  | 
else (if hom(s)=hom(t) then ( )  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
26  | 
                           else (writeln("Error: Mapping of Internals!");
 | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
27  | 
string_of_s s; writeln"";  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
28  | 
string_of_a a; writeln"";  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
29  | 
string_of_s t;writeln"";writeln"" ));  | 
| 
36692
 
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
 
haftmann 
parents: 
35174 
diff
changeset
 | 
30  | 
if member (op =) checked t then unchecked else insert (op =) t unchecked)  | 
| 33245 | 31  | 
in fold check_sas (nexts s a) unchecked end;  | 
32  | 
val unchecked' = fold check_sa (extacts @ intacts) unchecked  | 
|
| 
36692
 
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
 
haftmann 
parents: 
35174 
diff
changeset
 | 
33  | 
in (if member (op =) startsI s then  | 
| 
 
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
 
haftmann 
parents: 
35174 
diff
changeset
 | 
34  | 
(if member (op =) startsS (hom s) then ()  | 
| 
3072
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
35  | 
                     else writeln("Error: At start states!"))
 | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
36  | 
else ();  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
37  | 
checks(unchecked',s::checked)) end  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
38  | 
and checks([],_) = ()  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
39  | 
| checks(s::unchecked,checked) = check_s(s,unchecked,checked)  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
40  | 
in checks(startsI,[]) end;  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
41  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
42  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
43  | 
(* ------------------------------------------------------  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
44  | 
A B P E x a m p l e  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
45  | 
-------------------------------------------------------*)  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
46  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
47  | 
datatype msg = m | n | l;  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
48  | 
datatype act = Next | S_msg of msg | R_msg of msg  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
49  | 
| S_pkt of bool * msg | R_pkt of bool * msg  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
50  | 
| S_ack of bool | R_ack of bool;  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
51  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
52  | 
(* -------------------- Transition relation of Specification -----------*)  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
53  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
54  | 
fun transA((u,s),a,(v,t)) =  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
55  | 
(case a of  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
56  | 
Next => v andalso t = s |  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
57  | 
S_msg(q) => u andalso not(v) andalso t = s@[q] |  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
58  | 
R_msg(q) => u = v andalso s = (q::t) |  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
59  | 
S_pkt(b,q) => false |  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
60  | 
R_pkt(b,q) => false |  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
61  | 
S_ack(b) => false |  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
62  | 
R_ack(b) => false);  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
63  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
64  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
65  | 
(* ---------------------- Abstraction function --------------------------*)  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
66  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
67  | 
fun hom((env,p,a,q,b,_,_)) = (env,q@(if (a=b) then tl(p) else p));  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
68  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
69  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
70  | 
(* --------------------- Transition relation of Implementation ----------*)  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
71  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
72  | 
fun nexts (s as (env,p,a,q,b,ch1,ch2)) action =  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
73  | 
(case action of  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
74  | 
Next => if p=[] then [(true,p,a,q,b,ch1,ch2)] else [] |  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
75  | 
S_msg(mornorl) => if env then [(false,p@[mornorl],a,q,b,ch1,ch2)] else [] |  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
76  | 
R_msg(mornorl) => if (q<>[] andalso mornorl=hd(q))  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
77  | 
then [(env,p,a,tl(q),b,ch1,ch2)]  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
78  | 
else [] |  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
79  | 
S_pkt(h,mornorl) => if (p<>[] andalso mornorl=hd(p) andalso h=a)  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
80  | 
then (if (ch1<>[] andalso hd(rev(ch1))=(h,mornorl))  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
81  | 
then [s]  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
82  | 
else [s,(env,p,a,q,b,ch1@[(h,mornorl)],ch2)])  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
83  | 
else [] |  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
84  | 
R_pkt(h,mornorl) => if (ch1<>[] andalso hd(ch1)=(h,mornorl))  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
85  | 
then (if (h<>b andalso q=[])  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
86  | 
then [(env,p,a,q@[mornorl],not(b),ch1,ch2),  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
87  | 
(env,p,a,q@[mornorl],not(b),tl(ch1),ch2)]  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
88  | 
else [s,(env,p,a,q,b,tl(ch1),ch2)])  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
89  | 
else [] |  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
90  | 
S_ack(h) => if (h=b)  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
91  | 
then (if (ch2<>[] andalso h=hd(rev(ch2)))  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
92  | 
then [s]  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
93  | 
else [s,(env,p,a,q,b,ch1,ch2@[h])])  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
94  | 
else [] |  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
95  | 
R_ack(h) => if (ch2<>[] andalso hd(ch2)=h)  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
96  | 
then (if h=a  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
97  | 
then [(env,tl(p),not(a),q,b,ch1,ch2),  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
98  | 
(env,tl(p),not(a),q,b,ch1,tl(ch2))]  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
99  | 
else [s,(env,p,a,q,b,ch1,tl(ch2))])  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
100  | 
else [])  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
101  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
102  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
103  | 
val extactions = [Next,S_msg(m),R_msg(m),S_msg(n),R_msg(n),S_msg(l),R_msg(l)];  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
104  | 
val intactions = [S_pkt(true,m),R_pkt(true,m),S_ack(true),R_ack(true),  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
105  | 
S_pkt(false,m),R_pkt(false,m),S_ack(false),R_ack(false),  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
106  | 
S_pkt(true,n),R_pkt(true,n),S_pkt(true,l),R_pkt(true,l),  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
107  | 
S_pkt(false,n),R_pkt(false,n),S_pkt(false,l),R_pkt(false,l)];  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
108  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
109  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
110  | 
(* ------------------------------------  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
111  | 
Input / Output utilities  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
112  | 
------------------------------------*)  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
113  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
114  | 
fun print_list (lpar, rpar, pre: 'a -> unit) (lll : 'a list) =  | 
| 
44389
 
a3b5fdfb04a3
tuned signature -- contrast physical output primitives versus Output.raw_message;
 
wenzelm 
parents: 
42151 
diff
changeset
 | 
115  | 
let fun prec x = (Output.physical_stdout ","; pre x)  | 
| 
3072
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
116  | 
in  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
117  | 
(case lll of  | 
| 
44389
 
a3b5fdfb04a3
tuned signature -- contrast physical output primitives versus Output.raw_message;
 
wenzelm 
parents: 
42151 
diff
changeset
 | 
118  | 
[] => (Output.physical_stdout lpar; Output.physical_stdout rpar)  | 
| 
 
a3b5fdfb04a3
tuned signature -- contrast physical output primitives versus Output.raw_message;
 
wenzelm 
parents: 
42151 
diff
changeset
 | 
119  | 
| x::lll => (Output.physical_stdout lpar; pre x; List.app prec lll; Output.physical_stdout rpar))  | 
| 
3072
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
120  | 
end;  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
121  | 
|
| 
44389
 
a3b5fdfb04a3
tuned signature -- contrast physical output primitives versus Output.raw_message;
 
wenzelm 
parents: 
42151 
diff
changeset
 | 
122  | 
fun pr_bool true = Output.physical_stdout "true"  | 
| 
 
a3b5fdfb04a3
tuned signature -- contrast physical output primitives versus Output.raw_message;
 
wenzelm 
parents: 
42151 
diff
changeset
 | 
123  | 
| pr_bool false = Output.physical_stdout "false";  | 
| 
3072
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
124  | 
|
| 
44389
 
a3b5fdfb04a3
tuned signature -- contrast physical output primitives versus Output.raw_message;
 
wenzelm 
parents: 
42151 
diff
changeset
 | 
125  | 
fun pr_msg m = Output.physical_stdout "m"  | 
| 
 
a3b5fdfb04a3
tuned signature -- contrast physical output primitives versus Output.raw_message;
 
wenzelm 
parents: 
42151 
diff
changeset
 | 
126  | 
| pr_msg n = Output.physical_stdout "n"  | 
| 
 
a3b5fdfb04a3
tuned signature -- contrast physical output primitives versus Output.raw_message;
 
wenzelm 
parents: 
42151 
diff
changeset
 | 
127  | 
| pr_msg l = Output.physical_stdout "l";  | 
| 
3072
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
128  | 
|
| 
44389
 
a3b5fdfb04a3
tuned signature -- contrast physical output primitives versus Output.raw_message;
 
wenzelm 
parents: 
42151 
diff
changeset
 | 
129  | 
fun pr_act a = Output.physical_stdout (case a of  | 
| 
3072
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
130  | 
Next => "Next"|  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
131  | 
S_msg(ma) => "S_msg(ma)" |  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
132  | 
R_msg(ma) => "R_msg(ma)" |  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
133  | 
S_pkt(b,ma) => "S_pkt(b,ma)" |  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
134  | 
R_pkt(b,ma) => "R_pkt(b,ma)" |  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
135  | 
S_ack(b) => "S_ack(b)" |  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
136  | 
R_ack(b) => "R_ack(b)");  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
137  | 
|
| 
44389
 
a3b5fdfb04a3
tuned signature -- contrast physical output primitives versus Output.raw_message;
 
wenzelm 
parents: 
42151 
diff
changeset
 | 
138  | 
fun pr_pkt (b,ma) = (Output.physical_stdout "<"; pr_bool b;Output.physical_stdout ", "; pr_msg ma; Output.physical_stdout ">");  | 
| 
3072
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
139  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
140  | 
val pr_bool_list  = print_list("[","]",pr_bool);
 | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
141  | 
val pr_msg_list   = print_list("[","]",pr_msg);
 | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
142  | 
val pr_pkt_list   = print_list("[","]",pr_pkt);
 | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
143  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
144  | 
fun pr_tuple (env,p,a,q,b,ch1,ch2) =  | 
| 
44389
 
a3b5fdfb04a3
tuned signature -- contrast physical output primitives versus Output.raw_message;
 
wenzelm 
parents: 
42151 
diff
changeset
 | 
145  | 
        (Output.physical_stdout "{"; pr_bool env; Output.physical_stdout ", "; pr_msg_list p;  Output.physical_stdout ", ";
 | 
| 
 
a3b5fdfb04a3
tuned signature -- contrast physical output primitives versus Output.raw_message;
 
wenzelm 
parents: 
42151 
diff
changeset
 | 
146  | 
pr_bool a; Output.physical_stdout ", "; pr_msg_list q; Output.physical_stdout ", ";  | 
| 
 
a3b5fdfb04a3
tuned signature -- contrast physical output primitives versus Output.raw_message;
 
wenzelm 
parents: 
42151 
diff
changeset
 | 
147  | 
pr_bool b; Output.physical_stdout ", "; pr_pkt_list ch1; Output.physical_stdout ", ";  | 
| 
 
a3b5fdfb04a3
tuned signature -- contrast physical output primitives versus Output.raw_message;
 
wenzelm 
parents: 
42151 
diff
changeset
 | 
148  | 
pr_bool_list ch2; Output.physical_stdout "}");  | 
| 
3072
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
149  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
150  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
151  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
152  | 
(* ---------------------------------  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
153  | 
Main function call  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
154  | 
---------------------------------*)  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
155  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
156  | 
(*  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
157  | 
check(extactions,intactions,pr_act, [(true,[],true,[],false,[],[])],  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
158  | 
pr_tuple, nexts, hom, transA, [(true,[])]);  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
159  | 
*)  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
160  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
161  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
162  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
163  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
164  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
165  | 
(*  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
166  | 
Little test example  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
167  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
168  | 
datatype act = A;  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
169  | 
fun transA(s,a,t) = (not(s)=t);  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
170  | 
fun hom(i) = i mod 2 = 0;  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
171  | 
fun nexts s A = [(s+1) mod 4];  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
172  | 
check([A],[],K"A", [0], string_of_int, nexts, hom, transA, [true]);  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
173  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
174  | 
fun nexts s A = [(s+1) mod 5];  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
175  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
176  | 
*)  | 
| 17239 | 177  | 
|
178  | 
end;  |