src/ZF/UNITY/FP.thy
changeset 11479 697dcaaf478f
child 12195 ed2893765a08
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/ZF/UNITY/FP.thy	Wed Aug 08 14:33:10 2001 +0200
@@ -0,0 +1,22 @@
+(*  Title:      ZF/UNITY/FP.thy
+    ID:         $Id$
+    Author:     Sidi O Ehmety, Computer Laboratory
+    Copyright   2001  University of Cambridge
+
+Fixed Point of a Program
+
+From Misra, "A Logic for Concurrent Programming", 1994
+
+Theory ported from HOL.
+*)
+
+FP = UNITY +
+
+constdefs   
+  FP_Orig :: i=>i
+    "FP_Orig(F) == Union({A:condition. ALL B:condition. F : stable(A Int B)})"
+
+  FP :: i=>i
+    "FP(F) == {s:state. F : stable({s})}"
+
+end