| author | wenzelm |
| Thu, 27 Sep 2001 22:23:20 +0200 | |
| changeset 11603 | c3724decadef |
| parent 6008 | d0e9b1619468 |
| child 12218 | 6597093b77e7 |
| permissions | -rw-r--r-- |
(* Title: HOL/IOA/example/Correctness.thy ID: $Id$ Author: Olaf Mueller Copyright 1997 TU Muenchen Correctness Proof *) Correctness = SimCorrectness + Spec + Impl + default term consts sim_relation :: "((nat * bool) * (nat set * bool)) set" defs sim_relation_def "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 }" end