6008
|
1 |
(* Title: HOL/IOA/example/Correctness.thy
|
|
2 |
ID: $Id$
|
12218
|
3 |
Author: Olaf Müller
|
6008
|
4 |
*)
|
|
5 |
|
17244
|
6 |
header {* Correctness Proof *}
|
6008
|
7 |
|
17244
|
8 |
theory Correctness
|
|
9 |
imports SimCorrectness Spec Impl
|
|
10 |
begin
|
6008
|
11 |
|
17244
|
12 |
defaultsort type
|
6008
|
13 |
|
17244
|
14 |
constdefs
|
|
15 |
sim_relation :: "((nat * bool) * (nat set * bool)) set"
|
|
16 |
"sim_relation == {qua. let c = fst qua; a = snd qua ;
|
6008
|
17 |
k = fst c; b = snd c;
|
|
18 |
used = fst a; c = snd a
|
|
19 |
in
|
|
20 |
(! l:used. l < k) & b=c }"
|
|
21 |
|
17244
|
22 |
ML {* use_legacy_bindings (the_context ()) *}
|
|
23 |
|
|
24 |
end
|