author | aspinall |
Fri, 30 Sep 2005 18:18:34 +0200 | |
changeset 17740 | fc385ce6187d |
parent 17244 | 0b2ff9541727 |
child 19740 | 6b38551d0798 |
permissions | -rw-r--r-- |
(* Title: HOL/IOA/example/Correctness.thy ID: $Id$ Author: Olaf Müller *) header {* Correctness Proof *} theory Correctness imports SimCorrectness Spec Impl begin defaultsort type constdefs sim_relation :: "((nat * bool) * (nat set * bool)) set" "sim_relation == {qua. let c = fst qua; a = snd qua ; k = fst c; b = snd c; used = fst a; c = snd a in (! l:used. l < k) & b=c }" ML {* use_legacy_bindings (the_context ()) *} end