| author | wenzelm | 
| Mon, 08 Oct 2007 18:13:10 +0200 | |
| changeset 24911 | 4efb68e5576d | 
| parent 19739 | c58ef2aa5430 | 
| child 35174 | e15040ae75d7 | 
| permissions | -rw-r--r-- | 
(* Title: HOL/IOA/NTP/Action.thy ID: $Id$ Author: Tobias Nipkow & Konrad Slind *) header {* The set of all actions of the system *} theory Action imports Packet begin datatype 'm action = S_msg 'm | R_msg 'm | S_pkt "'m packet" | R_pkt "'m packet" | S_ack bool | R_ack bool | C_m_s | C_m_r | C_r_s | C_r_r 'm end