author | wenzelm |
Mon, 06 Feb 2006 20:59:10 +0100 | |
changeset 18943 | 947d3a694654 |
parent 17291 | 94f6113fe9ed |
child 19742 | 86f21beabafc |
permissions | -rw-r--r-- |
1479 | 1 |
(* Title: HOLCF/ex/Fix2.thy |
1274 | 2 |
ID: $Id$ |
1479 | 3 |
Author: Franz Regensburger |
1274 | 4 |
|
17291 | 5 |
Show that fix is the unique least fixed-point operator. |
12036 | 6 |
From axioms gix1_def,gix2_def it follows that fix = gix |
1274 | 7 |
*) |
8 |
||
17291 | 9 |
theory Fix2 |
10 |
imports HOLCF |
|
11 |
begin |
|
1274 | 12 |
|
13 |
consts |
|
17291 | 14 |
gix :: "('a->'a)->'a" |
1274 | 15 |
|
17291 | 16 |
axioms |
17 |
gix1_def: "F$(gix$F) = gix$F" |
|
18 |
gix2_def: "F$y=y ==> gix$F << y" |
|
1274 | 19 |
|
17291 | 20 |
ML {* use_legacy_bindings (the_context ()) *} |
1274 | 21 |
|
22 |
end |