| author | wenzelm | 
| Fri, 24 May 2019 20:16:35 +0200 | |
| changeset 70288 | 2e101846ad8f | 
| parent 62002 | f1599e98c4d0 | 
| permissions | -rw-r--r-- | 
| 42151 | 1 | (* Title: HOL/HOLCF/IOA/ABP/Action.thy | 
| 40945 | 2 | Author: Olaf Müller | 
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 3 | *) | 
| 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 4 | |
| 62002 | 5 | section \<open>The set of all actions of the system\<close> | 
| 17244 | 6 | |
| 7 | theory Action | |
| 8 | imports Packet | |
| 9 | begin | |
| 10 | ||
| 58310 | 11 | datatype 'm action = | 
| 17244 | 12 | Next | S_msg 'm | R_msg 'm | 
| 13 | | S_pkt "'m packet" | R_pkt "'m packet" | |
| 14 | | S_ack bool | R_ack bool | |
| 15 | ||
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 16 | end |