| author | wenzelm | 
| Wed, 17 Jun 2009 23:22:52 +0200 | |
| changeset 31695 | 36c5c15597f2 | 
| parent 19739 | c58ef2aa5430 | 
| child 35174 | e15040ae75d7 | 
| permissions | -rw-r--r-- | 
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 1 | (* Title: HOL/IOA/NTP/Action.thy | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 2 | ID: $Id$ | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 3 | Author: Tobias Nipkow & Konrad Slind | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 4 | *) | 
| 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 5 | |
| 17244 | 6 | header {* The set of all actions of the system *}
 | 
| 7 | ||
| 8 | theory Action | |
| 9 | imports Packet | |
| 10 | begin | |
| 11 | ||
| 12 | datatype 'm action = S_msg 'm | R_msg 'm | |
| 13 | | S_pkt "'m packet" | R_pkt "'m packet" | |
| 14 | | S_ack bool | R_ack bool | |
| 15 | | C_m_s | C_m_r | C_r_s | C_r_r 'm | |
| 16 | ||
| 3073 
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 17 | end |