src/HOL/UNITY/Follows.thy
author paulson
Mon, 24 May 1999 15:47:06 +0200
changeset 6706 d8067e272d4f
child 6809 5b8912f7bb69
permissions -rw-r--r--
Theory of the "Follows" relation
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
6706
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/Follows
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
     2
    ID:         $Id$
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
     5
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
     6
The Follows relation of Charpentier and Sivilotte
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
     7
*)
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
     8
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
     9
Follows = Union +
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
    10
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
    11
constdefs
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
    12
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
    13
  Follows :: "['a => 'b::{order}, 'a => 'b::{order}] => 'a program set"
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
    14
                 (infixl 65)
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
    15
   "f Follows g == Increasing g Int Increasing f Int
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
    16
                   Always {s. f s <= g s} Int
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
    17
                   (INT k. {s. k <= g s} LeadsTo {s. k <= f s})"
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
    18
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
    19
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
    20
end