src/ZF/Fixedpt.thy
author paulson
Wed, 09 Oct 1996 13:32:33 +0200
changeset 2073 fb0655539d05
parent 1478 2b8c2a7547ab
child 2469 b50b8c0eec01
permissions -rw-r--r--
New unified treatment of sequent calculi by Sara Kalvala combines the old LK and Modal with the new ILL (Int. Linear Logic)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      ZF/fixedpt.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
Least and greatest fixed points
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
124
858ab9a9b047 made pseudo theories for all ML files;
clasohm
parents: 0
diff changeset
     9
Fixedpt = ZF + "domrange" +
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
consts
1401
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 753
diff changeset
    11
  bnd_mono    :: [i,i=>i]=>o
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 753
diff changeset
    12
  lfp, gfp    :: [i,i=>i]=>i
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
753
ec86863e87c8 replaced "rules" by "defs"
lcp
parents: 124
diff changeset
    14
defs
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
  (*monotone operator from Pow(D) to itself*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    16
  bnd_mono_def 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    17
      "bnd_mono(D,h) == h(D)<=D & (ALL W X. W<=X --> X<=D --> h(W) <= h(X))"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    18
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
  lfp_def     "lfp(D,h) == Inter({X: Pow(D). h(X) <= X})"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    20
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    21
  gfp_def     "gfp(D,h) == Union({X: Pow(D). X <= h(X)})"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    22
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    23
end