| author | nipkow |
| Mon, 04 Mar 1996 14:37:33 +0100 | |
| changeset 1531 | e5eb247ad13c |
| parent 1475 | 7f5a4cd08209 |
| child 1558 | 9c6ebfab4e05 |
| permissions | -rw-r--r-- |
| 1475 | 1 |
(* Title: HOL/Lfp.thy |
| 923 | 2 |
ID: $Id$ |
| 1475 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
| 923 | 4 |
Copyright 1992 University of Cambridge |
5 |
||
6 |
The Knaster-Tarski Theorem |
|
7 |
*) |
|
8 |
||
| 1114 | 9 |
Lfp = mono + Prod + |
|
1370
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1264
diff
changeset
|
10 |
consts lfp :: ['a set=>'a set] => 'a set |
| 923 | 11 |
defs |
12 |
(*least fixed point*) |
|
13 |
lfp_def "lfp(f) == Inter({u. f(u) <= u})"
|
|
14 |
end |