author | lcp |
Tue, 25 Apr 1995 11:14:03 +0200 | |
changeset 1072 | 0140ff702b23 |
parent 1051 | 4fcd0638e61d |
child 1151 | c820b3cc3df0 |
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 = Solve + Impl + Spec + consts hom :: "'m impl_state => 'm list" defs hom_def "hom(s) == rq(rec(s)) @ (if rbit(rec s) = sbit(sen s) then sq(sen s) \ \ else ttl(sq(sen s)))" end