author | huffman |
Wed, 12 Oct 2005 03:02:18 +0200 | |
changeset 17840 | 11bcd77cfa22 |
parent 17291 | 94f6113fe9ed |
child 19742 | 86f21beabafc |
permissions | -rw-r--r-- |
(* Title: HOLCF/ex/Fix2.thy ID: $Id$ Author: Franz Regensburger Show that fix is the unique least fixed-point operator. From axioms gix1_def,gix2_def it follows that fix = gix *) theory Fix2 imports HOLCF begin consts gix :: "('a->'a)->'a" axioms gix1_def: "F$(gix$F) = gix$F" gix2_def: "F$y=y ==> gix$F << y" ML {* use_legacy_bindings (the_context ()) *} end