begin
:: deftheorem Def1 defines is_a_prefix_of TREES_1:def 1 :
for p, q being FinSequence holds
( p is_a_prefix_of q iff ex n being Element of NAT st p = q | (Seg n) );
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th8:
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th15:
theorem Th16:
Lm1:
for A, B being finite set st A c= B & card A = card B holds
A = B
theorem
canceled;
theorem
canceled;
theorem Th19:
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th23:
theorem Th24:
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th30:
theorem Th31:
theorem Th32:
theorem
:: deftheorem TREES_1:def 2 :
canceled;
:: deftheorem TREES_1:def 3 :
canceled;
:: deftheorem Def4 defines ProperPrefixes TREES_1:def 4 :
for p being FinSequence
for b2 being set holds
( b2 = ProperPrefixes p iff for x being set holds
( x in b2 iff ex q being FinSequence st
( x = q & q is_a_proper_prefix_of p ) ) );
theorem
canceled;
theorem Th35:
theorem Th36:
theorem Th37:
theorem
theorem Th39:
theorem Th40:
theorem
theorem Th42:
:: deftheorem Def5 defines Tree-like TREES_1:def 5 :
for X being set holds
( X is Tree-like iff ( X c= NAT * & ( for p being FinSequence of NAT st p in X holds
ProperPrefixes p c= X ) & ( for p being FinSequence of NAT
for k, n being Element of NAT st p ^ <*k*> in X & n <= k holds
p ^ <*n*> in X ) ) );
theorem
canceled;
theorem Th44:
theorem Th45:
theorem Th46:
theorem Th47:
theorem Th48:
theorem Th49:
theorem Th50:
theorem
canceled;
theorem
theorem
:: deftheorem TREES_1:def 6 :
canceled;
:: deftheorem defines elementary_tree TREES_1:def 7 :
for n being Element of NAT holds elementary_tree n = { <*k*> where k is Element of NAT : k < n } \/ {{}};
theorem
canceled;
theorem Th55:
theorem Th56:
theorem
:: deftheorem Def8 defines Leaves TREES_1:def 8 :
for T being Tree
for b2 being Subset of T holds
( b2 = Leaves T iff for p being FinSequence of NAT holds
( p in b2 iff ( p in T & ( for q being FinSequence of NAT holds
( not q in T or not p is_a_proper_prefix_of q ) ) ) ) );
:: deftheorem Def9 defines | TREES_1:def 9 :
for T being Tree
for p being FinSequence of NAT st p in T holds
for b3 being Tree holds
( b3 = T | p iff for q being FinSequence of NAT holds
( q in b3 iff p ^ q in T ) );
theorem
canceled;
theorem
canceled;
theorem
:: deftheorem defines Leaf TREES_1:def 10 :
for T being Tree st Leaves T <> {} holds
for b2 being Element of T holds
( b2 is Leaf of T iff b2 in Leaves T );
:: deftheorem defines Subtree TREES_1:def 11 :
for T, b2 being Tree holds
( b2 is Subtree of T iff ex p being Element of T st b2 = T | p );
:: deftheorem Def12 defines with-replacement TREES_1:def 12 :
for T being Tree
for p being FinSequence of NAT
for T1 being Tree st p in T holds
for b4 being Tree holds
( b4 = T with-replacement (p,T1) iff for q being FinSequence of NAT holds
( q in b4 iff ( ( q in T & not p is_a_proper_prefix_of q ) or ex r being FinSequence of NAT st
( r in T1 & q = p ^ r ) ) ) );
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th64:
theorem
canceled;
theorem
theorem Th67:
theorem
:: deftheorem Def13 defines AntiChain_of_Prefixes-like TREES_1:def 13 :
for IT being set holds
( IT is AntiChain_of_Prefixes-like iff ( ( for x being set st x in IT holds
x is FinSequence ) & ( for p1, p2 being FinSequence st p1 in IT & p2 in IT & p1 <> p2 holds
not p1,p2 are_c=-comparable ) ) );
theorem
canceled;
theorem Th70:
theorem Th71:
:: deftheorem Def14 defines AntiChain_of_Prefixes TREES_1:def 14 :
for T being Tree
for b2 being AntiChain_of_Prefixes holds
( b2 is AntiChain_of_Prefixes of T iff b2 c= T );
theorem
canceled;
theorem Th73:
theorem
theorem
:: deftheorem Def15 defines height TREES_1:def 15 :
for T being finite Tree
for b2 being Element of NAT holds
( b2 = height T iff ( ex p being FinSequence of NAT st
( p in T & len p = b2 ) & ( for p being FinSequence of NAT st p in T holds
len p <= b2 ) ) );
:: deftheorem Def16 defines width TREES_1:def 16 :
for T being finite Tree
for b2 being Element of NAT holds
( b2 = width T iff ex X being AntiChain_of_Prefixes of T st
( b2 = card X & ( for Y being AntiChain_of_Prefixes of T holds card Y <= card X ) ) );
theorem
canceled;
theorem
canceled;
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem Th85:
begin
theorem
theorem
theorem
theorem
theorem
theorem
theorem
:: deftheorem defines TrivialInfiniteTree TREES_1:def 17 :
TrivialInfiniteTree = { (k |-> 0) where k is Element of NAT : verum } ;
theorem Th19: