IOA/example/Spec.thy
changeset 252 a4dc62a46ee4
parent 251 f04b33ce250f
child 253 132634d24019
equal deleted inserted replaced
251:f04b33ce250f 252:a4dc62a46ee4
     1 (*  Title:      HOL/IOA/example/Spec.thy
       
     2     ID:         $Id$
       
     3     Author:     Tobias Nipkow & Konrad Slind
       
     4     Copyright   1994  TU Muenchen
       
     5 
       
     6 The specification of reliable transmission
       
     7 *)
       
     8 
       
     9 Spec = List + IOA + Action +
       
    10 
       
    11 consts
       
    12 
       
    13 spec_sig   :: "'m action signature"
       
    14 spec_trans :: "('m action, 'm list)transition set"
       
    15 spec_ioa   :: "('m action, 'm list)ioa"
       
    16 
       
    17 defs
       
    18 
       
    19 sig_def "spec_sig == <UN m.{S_msg(m)}, 
       
    20                      UN m.{R_msg(m)}, 
       
    21                      {}>"
       
    22 
       
    23 trans_def "spec_trans ==                           
       
    24  {tr. let s = fst(tr);                            
       
    25           t = snd(snd(tr))                        
       
    26       in                                          
       
    27       case fst(snd(tr))                           
       
    28       of                                          
       
    29       S_msg(m) => t = s@[m]  |                    
       
    30       R_msg(m) => s = (m#t)  |                    
       
    31       S_pkt(pkt) => False |                    
       
    32       R_pkt(pkt) => False |                    
       
    33       S_ack(b) => False |                      
       
    34       R_ack(b) => False |                      
       
    35       C_m_s => False |                            
       
    36       C_m_r => False |                            
       
    37       C_r_s => False |                            
       
    38       C_r_r(m) => False}"
       
    39 
       
    40 ioa_def "spec_ioa == <spec_sig, {[]}, spec_trans>"
       
    41 
       
    42 end