author | wenzelm |
Fri, 08 Sep 2006 19:44:43 +0200 | |
changeset 20494 | 99ad217b6974 |
parent 19739 | c58ef2aa5430 |
child 25131 | 2c8caac48ade |
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" constdefs hdr :: "'msg packet => bool" "hdr == fst" msg :: "'msg packet => 'msg" "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