src/ZF/AC/AC16_WO4.thy
author wenzelm
Wed, 09 Jul 1997 17:00:34 +0200
changeset 3511 da4dd8b7ced4
parent 1478 2b8c2a7547ab
child 5284 c77e9dd9bafc
permissions -rw-r--r--
removed obsolete init_pps and init_thy_reader;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      ZF/AC/AC16_WO4.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
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     6
AC16_WO4 = OrderType + AC16_lemmas + Cardinal_aux +
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
consts
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
     9
1401
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1203
diff changeset
    10
  ww  :: [i, i] => i
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1203
diff changeset
    11
  s_u :: [i, i, i, i] => i
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1203
diff changeset
    12
  MM  :: [i, i, i] => i
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1203
diff changeset
    13
  LL  :: [i, i, i] => i
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1203
diff changeset
    14
  GG  :: [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
defs
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
  ww_def  "ww(x, k) == {u:Pow(x). u eqpoll k}"
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    19
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    20
  s_u_def "s_u(u, t_n, k, y) == {v:t_n. u:v & k lepoll v Int y}"
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
  MM_def  "MM(t_n, k, y) == {v:t_n. k lepoll v Int y}"
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    23
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    24
  LL_def  "LL(t_n, k, y) == {v Int y. v:MM(t_n, k, y)}"
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    25
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    26
  GG_def  "GG(t_n, k, y) == lam v:LL(t_n, k, y). 
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    27
                            (THE w. w:MM(t_n, k, y) & v <= w) - v"
1196
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    28
d43c1f7a53fe Numerous small improvements by KG and LCP
lcp
parents:
diff changeset
    29
end