author | oheimb |
Mon, 21 Sep 1998 23:25:27 +0200 | |
changeset 5526 | e7617b57a3e6 |
parent 4776 | 1f9362e769c1 |
child 5648 | fe887910e32e |
permissions | -rw-r--r-- |
(* Title: HOL/UNITY/FP ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1998 University of Cambridge Fixed Point of a Program From Misra, "A Logic for Concurrent Programming", 1994 *) FP = UNITY + constdefs FP_Orig :: "('a * 'a)set set => 'a set" "FP_Orig Acts == Union{A. ALL B. stable Acts (A Int B)}" FP :: "('a * 'a)set set => 'a set" "FP Acts == {s. stable Acts {s}}" end