begin
theorem
canceled;
theorem Th2:
theorem Th3:
theorem Th4:
:: deftheorem Def1 defines lower WELLFND1:def 1 :
for R being RelStr
for X being Subset of R holds
( X is lower iff for x, y being set st x in X & [y,x] in the InternalRel of R holds
y in X );
theorem Th5:
theorem Th6:
begin
:: deftheorem Def2 defines well_founded WELLFND1:def 2 :
for R being RelStr holds
( R is well_founded iff the InternalRel of R is_well_founded_in the carrier of R );
:: deftheorem Def3 defines well_founded WELLFND1:def 3 :
for R being RelStr
for X being Subset of R holds
( X is well_founded iff the InternalRel of R is_well_founded_in X );
:: deftheorem Def4 defines well_founded-Part WELLFND1:def 4 :
for R being RelStr
for b2 being Subset of R holds
( b2 = well_founded-Part R iff b2 = union { S where S is Subset of R : ( S is well_founded & S is lower ) } );
theorem Th7:
theorem Th8:
theorem Th9:
theorem Th10:
theorem Th11:
:: deftheorem Def5 defines is_recursively_expressed_by WELLFND1:def 5 :
for R being non empty RelStr
for V being non empty set
for H being Function of [: the carrier of R,(PFuncs ( the carrier of R,V)):],V
for F being Function holds
( F is_recursively_expressed_by H iff for x being Element of R holds F . x = H . (x,(F | ( the InternalRel of R -Seg x))) );
theorem
theorem
theorem
:: deftheorem WELLFND1:def 6 :
canceled;
:: deftheorem Def7 defines descending WELLFND1:def 7 :
for R being RelStr
for f being sequence of R holds
( f is descending iff for n being Nat holds
( f . (n + 1) <> f . n & [(f . (n + 1)),(f . n)] in the InternalRel of R ) );
theorem