src/HOL/HOLCF/IOA/NTP/Receiver.thy
changeset 41476 0fa9629aa399
parent 40774 0437dbc127b3
child 42151 4da4fc77664b
equal deleted inserted replaced
41468:0e4f6cf20cdf 41476:0fa9629aa399
     6 
     6 
     7 theory Receiver
     7 theory Receiver
     8 imports IOA Action
     8 imports IOA Action
     9 begin
     9 begin
    10 
    10 
    11 types
    11 type_synonym
    12 
    12 
    13 'm receiver_state
    13 'm receiver_state
    14 = "'m list * bool multiset * 'm packet multiset * bool * bool"
    14 = "'m list * bool multiset * 'm packet multiset * bool * bool"
    15 (* messages  #replies        #received            header mode *)
    15 (* messages  #replies        #received            header mode *)
    16 
    16