src/HOL/UNITY/FP.thy
changeset 4776 1f9362e769c1
child 5648 fe887910e32e
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/UNITY/FP.thy	Fri Apr 03 12:34:33 1998 +0200
@@ -0,0 +1,21 @@
+(*  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