author | clasohm |
Fri, 24 Nov 1995 11:46:23 +0100 | |
changeset 1367 | 78bdb2d04771 |
parent 1274 | ea0668a1c0ba |
child 1479 | 21eb5e156d91 |
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