| author | oheimb | 
| Tue, 04 Nov 1997 14:37:51 +0100 | |
| changeset 4121 | 390e10ddadf2 | 
| parent 3898 | f6bf42312e9e | 
| child 4816 | 64f075872f69 | 
| permissions | -rw-r--r-- | 
(* Title: HOL/IOA/NTP/Correctness.thy ID: $Id$ Author: Tobias Nipkow & Konrad Slind Copyright 1994 TU Muenchen The main correctness proof: Impl implements Spec *) Correctness = IOA + 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