src/ZF/Fixedpt.thy
author paulson
Mon, 13 Jul 1998 16:43:57 +0200
changeset 5137 60205b0de9b9
parent 3940 1d5bee4d047f
child 6093 87bf8c03b169
permissions -rw-r--r--
Huge tidy-up: removal of leading \!\! addsplits split_tac[split_if] now default nat_succI now included by default

(*  Title:      ZF/fixedpt.thy
    ID:         $Id$
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1992  University of Cambridge

Least and greatest fixed points
*)

Fixedpt = domrange +

global

consts
  bnd_mono    :: [i,i=>i]=>o
  lfp, gfp    :: [i,i=>i]=>i

local

defs
  (*monotone operator from Pow(D) to itself*)
  bnd_mono_def 
      "bnd_mono(D,h) == h(D)<=D & (ALL W X. W<=X --> X<=D --> h(W) <= h(X))"

  lfp_def     "lfp(D,h) == Inter({X: Pow(D). h(X) <= X})"

  gfp_def     "gfp(D,h) == Union({X: Pow(D). X <= h(X)})"

end