| author | wenzelm |
| Tue, 13 Jun 2000 18:33:34 +0200 | |
| changeset 9058 | 7856a01119fb |
| parent 9019 | 9c1118619d6c |
| child 10265 | 4e004b548049 |
| permissions | -rw-r--r-- |
| 6706 | 1 |
(* Title: HOL/UNITY/Follows |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1998 University of Cambridge |
|
5 |
||
| 8128 | 6 |
The "Follows" relation of Charpentier and Sivilotte |
|
9019
9c1118619d6c
new parent MultisetOrder and new results about multiset unions
paulson
parents:
8948
diff
changeset
|
7 |
|
|
9c1118619d6c
new parent MultisetOrder and new results about multiset unions
paulson
parents:
8948
diff
changeset
|
8 |
add_path "../Induct"; |
| 6706 | 9 |
*) |
10 |
||
|
9019
9c1118619d6c
new parent MultisetOrder and new results about multiset unions
paulson
parents:
8948
diff
changeset
|
11 |
Follows = SubstAx + ListOrder + MultisetOrder + |
| 6706 | 12 |
|
13 |
constdefs |
|
14 |
||
15 |
Follows :: "['a => 'b::{order}, 'a => 'b::{order}] => 'a program set"
|
|
| 6809 | 16 |
(infixl "Fols" 65) |
17 |
"f Fols g == Increasing g Int Increasing f Int |
|
18 |
Always {s. f s <= g s} Int
|
|
19 |
(INT k. {s. k <= g s} LeadsTo {s. k <= f s})"
|
|
| 6706 | 20 |
|
21 |
||
22 |
end |