| author | paulson |
| Thu, 06 Aug 1998 15:47:26 +0200 | |
| changeset 5277 | e4297d03e5d2 |
| parent 5253 | 82a5ca6290aa |
| child 5313 | 1861a564d7e2 |
| 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 |
||
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5253
diff
changeset
|
6 |
LeadsTo relation: restricted to the set of reachable states. |
| 4776 | 7 |
*) |
8 |
||
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5253
diff
changeset
|
9 |
SubstAx = WFair + Traces + |
| 4776 | 10 |
|
11 |
constdefs |
|
12 |
||
| 5253 | 13 |
LeadsTo :: "['a program, 'a set, 'a set] => bool" |
14 |
"LeadsTo prg A B == |
|
15 |
leadsTo (Acts prg) |
|
16 |
(reachable prg Int A) |
|
17 |
(reachable prg Int B)" |
|
| 4776 | 18 |
end |