src/HOL/UNITY/SubstAx.thy
author paulson
Thu, 02 Jul 1998 16:53:55 +0200
changeset 5111 8f4b72f0c15d
parent 4776 1f9362e769c1
child 5253 82a5ca6290aa
permissions -rw-r--r--
Uncurried functions LeadsTo and reach Deleted leading parameters thanks to new Goal command

(*  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