| author | oheimb |
| Fri, 11 Dec 1998 18:56:30 +0100 | |
| changeset 6027 | 9dd06eeda95c |
| parent 5648 | fe887910e32e |
| child 6536 | 281d44905cab |
| permissions | -rw-r--r-- |
(* Title: HOL/UNITY/SubstAx ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1998 University of Cambridge LeadsTo relation: restricted to the set of reachable states. *) SubstAx = WFair + Constrains + constdefs LeadsTo :: "['a set, 'a set] => 'a program set" "LeadsTo A B == {F. F : leadsTo (reachable F Int A) (reachable F Int B)}" end