author | nipkow |
Sun, 29 Jan 1995 14:02:17 +0100 | |
changeset 204 | 21c405b4039f |
parent 168 | 44ff2275d44f |
child 249 | 492493334e0f |
permissions | -rw-r--r-- |
(* Title: HOL/IOA/example/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)), \ \ sq(sen(s)), \ \ ttl(sq(sen(s))))" end