| author | wenzelm |
| Wed, 06 Oct 1999 18:11:37 +0200 | |
| changeset 7754 | 4b1bc1266c8c |
| parent 6809 | 5b8912f7bb69 |
| child 8074 | 36a6c38e0eca |
| permissions | -rw-r--r-- |
(* 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 "Fols" 65) "f Fols 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