| author | berghofe |
| Tue, 07 Aug 2001 19:26:42 +0200 | |
| changeset 11470 | d3a3be6660f9 |
| parent 10835 | f4745d77e620 |
| child 12036 | 49f6c49454c2 |
| permissions | -rw-r--r-- |
| 1479 | 1 |
(* Title: HOLCF/ex/Fix2.thy |
| 1274 | 2 |
ID: $Id$ |
| 1479 | 3 |
Author: Franz Regensburger |
4 |
Copyright 1995 Technische Universitaet Muenchen |
|
| 1274 | 5 |
|
6 |
Show that fix is the unique least fixed-point operator. |
|
7 |
From axioms gix1_def,gix2_def it follows that fix = gix |
|
8 |
||
9 |
*) |
|
10 |
||
11 |
Fix2 = Fix + |
|
12 |
||
13 |
consts |
|
14 |
||
15 |
gix :: "('a->'a)->'a"
|
|
16 |
||
17 |
rules |
|
18 |
||
| 10835 | 19 |
gix1_def "F$(gix$F) = gix$F" |
20 |
gix2_def "F$y=y ==> gix$F << y" |
|
| 1274 | 21 |
|
22 |
end |