IOA/example/Packet.thy
changeset 168 44ff2275d44f
parent 156 fd1be45b64bf
equal deleted inserted replaced
167:37a6e2f55230 168:44ff2275d44f
       
     1 (*  Title:      HOL/IOA/example/Packet.thy
       
     2     ID:         $Id$
       
     3     Author:     Tobias Nipkow & Konrad Slind
       
     4     Copyright   1994  TU Muenchen
       
     5 
       
     6 Packets
       
     7 *)
       
     8 
     1 Packet = Arith +
     9 Packet = Arith +
     2 
    10 
     3 types
    11 types
     4 
    12 
     5    'msg packet = "bool * 'msg"
    13    'msg packet = "bool * 'msg"
     7 consts
    15 consts
     8 
    16 
     9   hdr  :: "'msg packet => bool"
    17   hdr  :: "'msg packet => bool"
    10   msg :: "'msg packet => 'msg"
    18   msg :: "'msg packet => 'msg"
    11 
    19 
    12 rules
    20 defs
    13 
    21 
    14   hdr_def "hdr == fst"
    22   hdr_def "hdr == fst"
    15 
       
    16   msg_def "msg == snd"
    23   msg_def "msg == snd"
    17 
    24 
    18 end
    25 end