| author | nipkow |
| Fri, 15 Nov 2002 18:02:25 +0100 | |
| changeset 13716 | 73de0ef7cb25 |
| parent 12218 | 6597093b77e7 |
| child 14981 | e73f8140af78 |
| permissions | -rw-r--r-- |
(* Title: HOL/IOA/NTP/Correctness.thy ID: $Id$ Author: Tobias Nipkow & Konrad Slind License: GPL (GNU GENERAL PUBLIC LICENSE) The main correctness proof: Impl implements Spec. *) Correctness = Impl + Spec + constdefs hom :: 'm impl_state => 'm list "hom(s) == rq(rec(s)) @ (if rbit(rec s) = sbit(sen s) then sq(sen s) else tl(sq(sen s)))" end