diff -r 66b1a7c42d94 -r 1f9362e769c1 src/HOL/UNITY/SubstAx.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/UNITY/SubstAx.thy Fri Apr 03 12:34:33 1998 +0200 @@ -0,0 +1,19 @@ +(* Title: HOL/UNITY/SubstAx + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1998 University of Cambridge + +My treatment of the Substitution Axiom -- not as an axiom! +*) + +SubstAx = WFair + + +constdefs + + LeadsTo :: "['a set, ('a * 'a)set set, 'a set, 'a set] => bool" + "LeadsTo Init Acts A B == + leadsTo Acts (reachable Init Acts Int A) + (reachable Init Acts Int B)" + + +end