| author | wenzelm | 
| Mon, 31 Jul 2000 14:37:18 +0200 | |
| changeset 9481 | b16624f1ea38 | 
| parent 2570 | 24d7e8fb8261 | 
| child 10835 | f4745d77e620 | 
| permissions | -rw-r--r-- | 
(* Title: HOLCF/ex/Fix2.thy ID: $Id$ Author: Franz Regensburger Copyright 1995 Technische Universitaet Muenchen Show that fix is the unique least fixed-point operator. From axioms gix1_def,gix2_def it follows that fix = gix *) Fix2 = Fix + consts gix :: "('a->'a)->'a" rules gix1_def "F`(gix`F) = gix`F" gix2_def "F`y=y ==> gix`F << y" end