src/ZF/AC/DC.thy
author paulson
Wed, 09 Oct 1996 13:32:33 +0200
changeset 2073 fb0655539d05
parent 1478 2b8c2a7547ab
child 2469 b50b8c0eec01
permissions -rw-r--r--
New unified treatment of sequent calculi by Sara Kalvala combines the old LK and Modal with the new ILL (Int. Linear Logic)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      ZF/AC/DC.thy
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     2
    ID:         $Id$
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     3
    Author:     Krzysztof Grabczewski
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     4
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     5
Theory file for the proofs concernind the Axiom of Dependent Choice
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     6
*)
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     7
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     8
DC  =  AC_Equiv + Hartog + first + Cardinal_aux + "DC_lemmas" + 
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     9
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    10
consts
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    11
1401
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1203
diff changeset
    12
  DC  :: i => o
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1203
diff changeset
    13
  DC0 :: o
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1203
diff changeset
    14
  ff  :: [i, i, i, i] => i
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    15
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    16
rules
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    17
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    18
  DC_def  "DC(a) == ALL X R. R<=Pow(X)*X &
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    19
           (ALL Y:Pow(X). Y lesspoll a --> (EX x:X. <Y,x> : R)) 
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    20
           --> (EX f:a->X. ALL b<a. <f``b,f`b> : R)"
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    21
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    22
  DC0_def "DC0 == ALL A B R. R <= A*B & R~=0 & range(R) <= domain(R) 
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    23
           --> (EX f:nat->domain(R). ALL n:nat. <f`n,f`succ(n)>:R)"
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    24
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    25
  ff_def  "ff(b, X, Q, R) == transrec(b, %c r. 
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    26
                             THE x. first(x, {x:X. <r``c, x> : R}, Q))"
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    27
  
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    28
end