| author | nipkow | 
| Wed, 29 Mar 2000 15:09:51 +0200 | |
| changeset 8604 | c99e0024050c | 
| parent 8122 | b43ad07660b9 | 
| child 9685 | 6d123a7e30bd | 
| permissions | -rw-r--r-- | 
| 4776 | 1 | (* Title: HOL/UNITY/SubstAx | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1998 University of Cambridge | |
| 5 | ||
| 6536 | 6 | Weak LeadsTo relation (restricted to the set of reachable states) | 
| 4776 | 7 | *) | 
| 8 | ||
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5277diff
changeset | 9 | SubstAx = WFair + Constrains + | 
| 4776 | 10 | |
| 8041 | 11 | constdefs | 
| 8122 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8041diff
changeset | 12 | Ensures :: "['a set, 'a set] => 'a program set" (infixl 60) | 
| 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8041diff
changeset | 13 |     "A Ensures B == {F. F : (reachable F Int A) ensures B}"
 | 
| 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 paulson parents: 
8041diff
changeset | 14 | |
| 8041 | 15 | LeadsTo :: "['a set, 'a set] => 'a program set" (infixl 60) | 
| 16 |     "A LeadsTo B == {F. F : (reachable F Int A) leadsTo B}"
 | |
| 4776 | 17 | |
| 18 | end |