| author | haftmann | 
| Fri, 24 Oct 2008 17:48:36 +0200 | |
| changeset 28684 | 48faac324061 | 
| parent 25131 | 2c8caac48ade | 
| child 35174 | e15040ae75d7 | 
| permissions | -rw-r--r-- | 
(* Title: HOL/IOA/NTP/Packet.thy ID: $Id$ Author: Tobias Nipkow & Konrad Slind *) theory Packet imports Multiset begin types 'msg packet = "bool * 'msg" definition hdr :: "'msg packet => bool" where "hdr == fst" definition msg :: "'msg packet => 'msg" where "msg == snd" text {* Instantiation of a tautology? *} lemma eq_packet_imp_eq_hdr: "!x. x = packet --> hdr(x) = hdr(packet)" by simp declare hdr_def [simp] msg_def [simp] end