src/HOL/UNITY/Follows.thy
changeset 6706 d8067e272d4f
child 6809 5b8912f7bb69
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/UNITY/Follows.thy	Mon May 24 15:47:06 1999 +0200
@@ -0,0 +1,20 @@
+(*  Title:      HOL/UNITY/Follows
+    ID:         $Id$
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Copyright   1998  University of Cambridge
+
+The Follows relation of Charpentier and Sivilotte
+*)
+
+Follows = Union +
+
+constdefs
+
+  Follows :: "['a => 'b::{order}, 'a => 'b::{order}] => 'a program set"
+                 (infixl 65)
+   "f Follows g == Increasing g Int Increasing f Int
+                   Always {s. f s <= g s} Int
+                   (INT k. {s. k <= g s} LeadsTo {s. k <= f s})"
+
+
+end