| author | paulson | 
| Wed, 03 Oct 2001 11:45:24 +0200 | |
| changeset 11652 | b2d27a80b0d0 | 
| parent 11317 | 7f9e4c389318 | 
| child 12776 | 249600a63ba9 | 
| permissions | -rw-r--r-- | 
| 1478 | 1 | (* Title: ZF/AC/HH.thy | 
| 1123 | 2 | ID: $Id$ | 
| 1478 | 3 | Author: Krzysztof Grabczewski | 
| 1123 | 4 | |
| 5 | The theory file for the proofs of | |
| 6 | AC17 ==> AC1 | |
| 7 | AC1 ==> WO2 | |
| 8 | AC15 ==> WO6 | |
| 9 | *) | |
| 10 | ||
| 1196 | 11 | HH = AC_Equiv + Hartog + WO_AC + Let + | 
| 1123 | 12 | |
| 13 | consts | |
| 14 | ||
| 1401 | 15 | HH :: [i, i, i] => i | 
| 1123 | 16 | |
| 17 | defs | |
| 18 | ||
| 11317 | 19 | HH_def "HH(f,x,a) == transrec(a, %b r. let z = x - (\\<Union>c \\<in> b. r`c) | 
| 20 |                         in  if(f`z \\<in> Pow(z)-{0}, f`z, {x}))"
 | |
| 1123 | 21 | |
| 22 | end | |
| 23 |