| author | wenzelm | 
| Sun, 21 Oct 2007 16:27:42 +0200 | |
| changeset 25135 | 4f8176c940cf | 
| parent 17244 | 0b2ff9541727 | 
| child 35174 | e15040ae75d7 | 
| permissions | -rw-r--r-- | 
| 
3072
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOLCF/IOA/ABP/Spec.thy  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 12218 | 3  | 
Author: Olaf Müller  | 
| 
3072
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
4  | 
*)  | 
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
5  | 
|
| 17244 | 6  | 
header {* The specification of reliable transmission *}
 | 
7  | 
||
8  | 
theory Spec  | 
|
9  | 
imports IOA Action  | 
|
10  | 
begin  | 
|
| 
3072
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
11  | 
|
| 
25135
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
17244 
diff
changeset
 | 
12  | 
definition  | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
17244 
diff
changeset
 | 
13  | 
spec_sig :: "'m action signature" where  | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
17244 
diff
changeset
 | 
14  | 
  sig_def: "spec_sig = (UN m.{S_msg(m)},
 | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
17244 
diff
changeset
 | 
15  | 
                       UN m.{R_msg(m)} Un {Next},
 | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
17244 
diff
changeset
 | 
16  | 
                       {})"
 | 
| 
3072
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
17  | 
|
| 
25135
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
17244 
diff
changeset
 | 
18  | 
definition  | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
17244 
diff
changeset
 | 
19  | 
  spec_trans :: "('m action, 'm list)transition set" where
 | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
17244 
diff
changeset
 | 
20  | 
trans_def: "spec_trans =  | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
17244 
diff
changeset
 | 
21  | 
   {tr. let s = fst(tr);
 | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
17244 
diff
changeset
 | 
22  | 
t = snd(snd(tr))  | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
17244 
diff
changeset
 | 
23  | 
in  | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
17244 
diff
changeset
 | 
24  | 
case fst(snd(tr))  | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
17244 
diff
changeset
 | 
25  | 
of  | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
17244 
diff
changeset
 | 
26  | 
Next => t=s | (* Note that there is condition as in Sender *)  | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
17244 
diff
changeset
 | 
27  | 
S_msg(m) => t = s@[m] |  | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
17244 
diff
changeset
 | 
28  | 
R_msg(m) => s = (m#t) |  | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
17244 
diff
changeset
 | 
29  | 
S_pkt(pkt) => False |  | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
17244 
diff
changeset
 | 
30  | 
R_pkt(pkt) => False |  | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
17244 
diff
changeset
 | 
31  | 
S_ack(b) => False |  | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
17244 
diff
changeset
 | 
32  | 
R_ack(b) => False}"  | 
| 
3072
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
33  | 
|
| 
25135
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
17244 
diff
changeset
 | 
34  | 
definition  | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
17244 
diff
changeset
 | 
35  | 
  spec_ioa :: "('m action, 'm list)ioa" where
 | 
| 
 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 
wenzelm 
parents: 
17244 
diff
changeset
 | 
36  | 
  ioa_def: "spec_ioa = (spec_sig, {[]}, spec_trans)"
 | 
| 
3072
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
37  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
38  | 
end  |