author | lcp |
Tue, 25 Apr 1995 11:14:03 +0200 | |
changeset 1072 | 0140ff702b23 |
parent 1051 | 4fcd0638e61d |
child 1151 | c820b3cc3df0 |
permissions | -rw-r--r-- |
1051 | 1 |
(* Title: HOL/IOA/NTP/Correctness.thy |
2 |
ID: $Id$ |
|
3 |
Author: Tobias Nipkow & Konrad Slind |
|
4 |
Copyright 1994 TU Muenchen |
|
5 |
||
6 |
The main correctness proof: Impl implements Spec |
|
7 |
*) |
|
8 |
||
9 |
Correctness = Solve + Impl + Spec + |
|
10 |
||
11 |
consts |
|
12 |
||
13 |
hom :: "'m impl_state => 'm list" |
|
14 |
||
15 |
defs |
|
16 |
||
17 |
hom_def |
|
18 |
"hom(s) == rq(rec(s)) @ (if rbit(rec s) = sbit(sen s) then sq(sen s) \ |
|
19 |
\ else ttl(sq(sen s)))" |
|
20 |
||
21 |
end |