src/HOLCF/IOA/ABP/Packet.thy
author wenzelm
Tue, 10 Jul 2007 23:29:43 +0200
changeset 23719 ccd9cb15c062
parent 19738 1ac610922636
child 25131 2c8caac48ade
permissions -rw-r--r--
more markup for inner and outer syntax; added enclose;

(*  Title:      HOLCF/IOA/ABP/Packet.thy
    ID:         $Id$
    Author:     Olaf Müller
*)

header {* Packets *}

theory Packet
imports Main
begin

types
  'msg packet = "bool * 'msg"

constdefs
  hdr  :: "'msg packet => bool"
  "hdr == fst"

  msg :: "'msg packet => 'msg"
  "msg == snd"

end