author | paulson |
Wed, 27 Mar 1996 18:46:42 +0100 | |
changeset 1619 | cb62d89b7adb |
parent 1478 | 2b8c2a7547ab |
child 11317 | 7f9e4c389318 |
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 |
||
1196 | 19 |
HH_def "HH(f,x,a) == transrec(a, %b r. let z = x - (UN c:b. r`c) |
20 |
in if(f`z:Pow(z)-{0}, f`z, {x}))" |
|
1123 | 21 |
|
22 |
end |
|
23 |