| author | wenzelm | 
| Sun, 27 Dec 2020 13:52:55 +0100 | |
| changeset 73008 | dacf2598bb27 | 
| parent 63549 | b0d31c7def86 | 
| permissions | -rw-r--r-- | 
| 42151 | 1 | (* Title: HOL/HOLCF/ex/Fix2.thy | 
| 1479 | 2 | Author: Franz Regensburger | 
| 1274 | 3 | |
| 17291 | 4 | Show that fix is the unique least fixed-point operator. | 
| 12036 | 5 | From axioms gix1_def,gix2_def it follows that fix = gix | 
| 1274 | 6 | *) | 
| 7 | ||
| 17291 | 8 | theory Fix2 | 
| 9 | imports HOLCF | |
| 10 | begin | |
| 1274 | 11 | |
| 25135 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19742diff
changeset | 12 | axiomatization | 
| 63549 | 13 |   gix :: "('a \<rightarrow> 'a) \<rightarrow>'a" where
 | 
| 14 | gix1_def: "F\<cdot>(gix\<cdot>F) = gix\<cdot>F" and | |
| 15 | gix2_def: "F\<cdot>y = y \<Longrightarrow> gix\<cdot>F << y" | |
| 1274 | 16 | |
| 19742 | 17 | |
| 18 | lemma lemma1: "fix = gix" | |
| 40002 
c5b5f7a3a3b1
new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
 huffman parents: 
35174diff
changeset | 19 | apply (rule cfun_eqI) | 
| 40431 | 20 | apply (rule below_antisym) | 
| 19742 | 21 | apply (rule fix_least) | 
| 22 | apply (rule gix1_def) | |
| 23 | apply (rule gix2_def) | |
| 24 | apply (rule fix_eq [symmetric]) | |
| 25 | done | |
| 26 | ||
| 63549 | 27 | lemma lemma2: "gix\<cdot>F = lub (range (\<lambda>i. iterate i\<cdot>F\<cdot>UU))" | 
| 19742 | 28 | apply (rule lemma1 [THEN subst]) | 
| 29 | apply (rule fix_def2) | |
| 30 | done | |
| 1274 | 31 | |
| 32 | end |