author | paulson |
Wed, 16 Jan 2002 17:53:22 +0100 | |
changeset 12777 | 70b2651af635 |
parent 12036 | 49f6c49454c2 |
child 14981 | e73f8140af78 |
permissions | -rw-r--r-- |
1479 | 1 |
(* Title: HOLCF/ex/Fix2.thy |
1274 | 2 |
ID: $Id$ |
1479 | 3 |
Author: Franz Regensburger |
12036 | 4 |
License: GPL (GNU GENERAL PUBLIC LICENSE) |
1274 | 5 |
|
12036 | 6 |
Show that fix is the unique least fixed-point operator. |
7 |
From axioms gix1_def,gix2_def it follows that fix = gix |
|
1274 | 8 |
*) |
9 |
||
10 |
Fix2 = Fix + |
|
11 |
||
12 |
consts |
|
13 |
||
14 |
gix :: "('a->'a)->'a" |
|
15 |
||
16 |
rules |
|
17 |
||
10835 | 18 |
gix1_def "F$(gix$F) = gix$F" |
19 |
gix2_def "F$y=y ==> gix$F << y" |
|
1274 | 20 |
|
21 |
end |