| author | wenzelm | 
| Tue, 07 Aug 2001 21:27:00 +0200 | |
| changeset 11472 | d08d4e17a5f6 | 
| parent 5648 | fe887910e32e | 
| child 13796 | 19f50fa807ae | 
| permissions | -rw-r--r-- | 
| 4776 | 1 | (* Title: HOL/UNITY/FP | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1998 University of Cambridge | |
| 5 | ||
| 6 | Fixed Point of a Program | |
| 7 | ||
| 8 | From Misra, "A Logic for Concurrent Programming", 1994 | |
| 9 | *) | |
| 10 | ||
| 11 | FP = UNITY + | |
| 12 | ||
| 13 | constdefs | |
| 14 | ||
| 5648 | 15 | FP_Orig :: "'a program => 'a set" | 
| 16 |     "FP_Orig F == Union{A. ALL B. F : stable (A Int B)}"
 | |
| 4776 | 17 | |
| 5648 | 18 | FP :: "'a program => 'a set" | 
| 19 |     "FP F == {s. F : stable {s}}"
 | |
| 4776 | 20 | |
| 21 | end |