begin
:: deftheorem Def1 defines finite FINSET_1:def 1 :
for IT being set holds
( IT is finite iff ex p being Function st
( rng p = IT & dom p in omega ) );
Lm1:
for x being set holds {x} is finite
Lm2:
for A, B being set st A is finite & B is finite holds
A \/ B is finite
registration
let x1,
x2,
x3,
x4,
x5 be
set ;
cluster {x1,x2,x3,x4,x5} -> finite ;
coherence
{x1,x2,x3,x4,x5} is finite
end;
registration
let x1,
x2,
x3,
x4,
x5,
x6 be
set ;
cluster {x1,x2,x3,x4,x5,x6} -> finite ;
coherence
{x1,x2,x3,x4,x5,x6} is finite
end;
registration
let x1,
x2,
x3,
x4,
x5,
x6,
x7 be
set ;
cluster {x1,x2,x3,x4,x5,x6,x7} -> finite ;
coherence
{x1,x2,x3,x4,x5,x6,x7} is finite
end;
registration
let x1,
x2,
x3,
x4,
x5,
x6,
x7,
x8 be
set ;
cluster {x1,x2,x3,x4,x5,x6,x7,x8} -> finite ;
coherence
{x1,x2,x3,x4,x5,x6,x7,x8} is finite
end;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
theorem
theorem
theorem
theorem
theorem Th18:
scheme
Finite{
F1()
-> set ,
P1[
set ] } :
provided
A1:
F1() is
finite
and A2:
P1[
{} ]
and A3:
for
x,
B being
set st
x in F1() &
B c= F1() &
P1[
B] holds
P1[
B \/ {x}]
Lm3:
for A being set st A is finite & ( for X being set st X in A holds
X is finite ) holds
union A is finite
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
theorem Th26:
theorem
begin
theorem
canceled;
theorem Th29:
theorem
theorem
:: deftheorem Def2 defines finite-yielding FINSET_1:def 2 :
for R being Relation holds
( R is finite-yielding iff for x being set st x in rng R holds
x is finite );
deffunc H1( set ) -> set = $1 `1 ;
theorem
theorem
:: deftheorem defines centered FINSET_1:def 3 :
for F being set holds
( F is centered iff ( F <> {} & ( for G being set st G <> {} & G c= F & G is finite holds
meet G <> {} ) ) );
:: deftheorem defines finite-yielding FINSET_1:def 4 :
for I being set
for IT being b1 -defined Function holds
( IT is finite-yielding iff for i being set st i in I holds
IT . i is finite );
theorem