| author | nipkow |
| Tue, 23 Jun 1998 18:07:45 +0200 | |
| changeset 5071 | 548f398d770b |
| parent 4776 | 1f9362e769c1 |
| child 5111 | 8f4b72f0c15d |
| permissions | -rw-r--r-- |
(* 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