src/HOL/UNITY/Follows.thy
author paulson
Wed, 22 Dec 1999 17:18:03 +0100
changeset 8074 36a6c38e0eca
parent 6809 5b8912f7bb69
child 8122 b43ad07660b9
permissions -rw-r--r--
Working version after a FAILED attempt to base Follows upon LeadsETo
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
8074
36a6c38e0eca Working version after a FAILED attempt to base Follows upon LeadsETo
paulson
parents: 6809
diff changeset
     7
36a6c38e0eca Working version after a FAILED attempt to base Follows upon LeadsETo
paulson
parents: 6809
diff changeset
     8
The safety conditions ensures that "givenBy f" is implementable in the
36a6c38e0eca Working version after a FAILED attempt to base Follows upon LeadsETo
paulson
parents: 6809
diff changeset
     9
  progress part: g cannot do anything silly.
6706
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
    10
*)
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
    11
8074
36a6c38e0eca Working version after a FAILED attempt to base Follows upon LeadsETo
paulson
parents: 6809
diff changeset
    12
Follows = ELT +
6706
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
    13
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
    14
constdefs
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
    15
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
    16
  Follows :: "['a => 'b::{order}, 'a => 'b::{order}] => 'a program set"
6809
5b8912f7bb69 shortened Follows to Fols
paulson
parents: 6706
diff changeset
    17
                 (infixl "Fols" 65)
5b8912f7bb69 shortened Follows to Fols
paulson
parents: 6706
diff changeset
    18
   "f Fols g == Increasing g Int Increasing f Int
5b8912f7bb69 shortened Follows to Fols
paulson
parents: 6706
diff changeset
    19
                Always {s. f s <= g s} Int
5b8912f7bb69 shortened Follows to Fols
paulson
parents: 6706
diff changeset
    20
                (INT k. {s. k <= g s} LeadsTo {s. k <= f s})"
6706
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
    21
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
    22
d8067e272d4f Theory of the "Follows" relation
paulson
parents:
diff changeset
    23
end