| author | wenzelm | 
| Tue, 29 Mar 2011 17:47:11 +0200 | |
| changeset 42151 | 4da4fc77664b | 
| parent 40774 | 0437dbc127b3 | 
| child 58880 | 0baae4311a9f | 
| permissions | -rw-r--r-- | 
| 42151 | 1  | 
(* Title: HOL/HOLCF/IOA/NTP/Spec.thy  | 
| 
3073
 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
2  | 
Author: Tobias Nipkow & Konrad Slind  | 
| 
 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
3  | 
*)  | 
| 
 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
4  | 
|
| 17244 | 5  | 
header {* The specification of reliable transmission *}
 | 
6  | 
||
7  | 
theory Spec  | 
|
8  | 
imports IOA Action  | 
|
9  | 
begin  | 
|
| 
3073
 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
10  | 
|
| 27361 | 11  | 
definition  | 
12  | 
spec_sig :: "'m action signature" where  | 
|
13  | 
  sig_def: "spec_sig = (UN m.{S_msg(m)}, 
 | 
|
14  | 
                        UN m.{R_msg(m)}, 
 | 
|
15  | 
                        {})"
 | 
|
| 
3073
 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
16  | 
|
| 27361 | 17  | 
definition  | 
18  | 
  spec_trans :: "('m action, 'm list)transition set" where
 | 
|
19  | 
trans_def: "spec_trans =  | 
|
20  | 
   {tr. let s = fst(tr);                            
 | 
|
21  | 
t = snd(snd(tr))  | 
|
22  | 
in  | 
|
23  | 
case fst(snd(tr))  | 
|
24  | 
of  | 
|
25  | 
S_msg(m) => t = s@[m] |  | 
|
26  | 
R_msg(m) => s = (m#t) |  | 
|
27  | 
S_pkt(pkt) => False |  | 
|
28  | 
R_pkt(pkt) => False |  | 
|
29  | 
S_ack(b) => False |  | 
|
30  | 
R_ack(b) => False |  | 
|
31  | 
C_m_s => False |  | 
|
32  | 
C_m_r => False |  | 
|
33  | 
C_r_s => False |  | 
|
34  | 
C_r_r(m) => False}"  | 
|
| 
3073
 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
35  | 
|
| 27361 | 36  | 
definition  | 
37  | 
  spec_ioa :: "('m action, 'm list)ioa" where
 | 
|
38  | 
  ioa_def: "spec_ioa = (spec_sig, {[]}, spec_trans,{},{})"
 | 
|
| 17244 | 39  | 
|
| 
3073
 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
40  | 
end  |