begin
:: deftheorem Def1 defines with_common_domain MARGREL1:def 1 :
for IT being FinSequence-membered set holds
( IT is with_common_domain iff for a, b being FinSequence st a in IT & b in IT holds
len a = len b );
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
theorem
:: deftheorem defines = MARGREL1:def 2 :
for p, r being relation holds
( p = r iff for a being FinSequence holds
( a in p iff a in r ) );
theorem Th9:
:: deftheorem MARGREL1:def 3 :
canceled;
:: deftheorem defines the_arity_of MARGREL1:def 4 :
for p being relation st p <> {} holds
for b2 being Element of NAT holds
( b2 = the_arity_of p iff for a being FinSequence st a in p holds
b2 = len a );
:: deftheorem defines relation_length MARGREL1:def 5 :
for k being Element of NAT
for b2 being relation holds
( b2 is relation_length of k iff for a being FinSequence st a in b2 holds
len a = k );
:: deftheorem defines relation MARGREL1:def 6 :
for X being set
for b2 being relation holds
( b2 is relation of X iff for a being FinSequence st a in b2 holds
rng a c= X );
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th20:
theorem Th21:
:: deftheorem defines relation MARGREL1:def 7 :
for X being set
for k being Element of NAT
for b3 being relation holds
( b3 is relation of X,k iff ( b3 is relation of X & b3 is relation_length of k ) );
:: deftheorem Def8 defines relations_on MARGREL1:def 8 :
for D being non empty set
for b2 being set holds
( b2 = relations_on D iff for X being set holds
( X in b2 iff ( X c= D * & ( for a, b being FinSequence of D st a in X & b in X holds
len a = len b ) ) ) );
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
theorem
theorem
:: deftheorem Def9 defines = MARGREL1:def 9 :
for D being non empty set
for p, r being Element of relations_on D holds
( p = r iff for a being FinSequence of D holds
( a in p iff a in r ) );
:: deftheorem Def10 defines empty_rel MARGREL1:def 10 :
for D being non empty set
for b2 being Element of relations_on D holds
( b2 = empty_rel D iff for a being FinSequence of D holds not a in b2 );
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
:: deftheorem defines the_arity_of MARGREL1:def 11 :
for D being non empty set
for p being Element of relations_on D st p <> empty_rel D holds
for b3 being Element of NAT holds
( b3 = the_arity_of p iff for a being FinSequence of D st a in p holds
b3 = len a );
:: deftheorem defines BOOLEAN MARGREL1:def 12 :
BOOLEAN = {0,1};
:: deftheorem Def13 defines boolean MARGREL1:def 13 :
for x being set holds
( x is boolean iff x in BOOLEAN );
:: deftheorem defines 'not' MARGREL1:def 14 :
for v being boolean set holds
( ( v = FALSE implies 'not' v = TRUE ) & ( not v = FALSE implies 'not' v = FALSE ) );
:: deftheorem defines '&' MARGREL1:def 15 :
for v, w being boolean set holds
( ( v = TRUE & w = TRUE implies v '&' w = TRUE ) & ( ( not v = TRUE or not w = TRUE ) implies v '&' w = FALSE ) );
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
theorem
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
theorem
theorem
theorem
:: deftheorem Def16 defines ALL MARGREL1:def 16 :
for X being set holds
( ( not FALSE in X implies ALL X = TRUE ) & ( FALSE in X implies ALL X = FALSE ) );
theorem
begin
:: deftheorem Def17 defines boolean-valued MARGREL1:def 17 :
for f being Function holds
( f is boolean-valued iff rng f c= BOOLEAN );
:: deftheorem Def18 defines 'not' MARGREL1:def 18 :
for p, b2 being boolean-valued Function holds
( b2 = 'not' p iff ( dom b2 = dom p & ( for x being set st x in dom p holds
b2 . x = 'not' (p . x) ) ) );
:: deftheorem Def19 defines '&' MARGREL1:def 19 :
for p, q, b3 being boolean-valued Function holds
( b3 = p '&' q iff ( dom b3 = (dom p) /\ (dom q) & ( for x being set st x in dom b3 holds
b3 . x = (p . x) '&' (q . x) ) ) );
:: deftheorem defines 'not' MARGREL1:def 20 :
for A being non empty set
for p, b3 being Element of Funcs (A,BOOLEAN) holds
( b3 = 'not' p iff for x being Element of A holds b3 . x = 'not' (p . x) );
:: deftheorem defines '&' MARGREL1:def 21 :
for A being non empty set
for p, q, b4 being Element of Funcs (A,BOOLEAN) holds
( b4 = p '&' q iff for x being Element of A holds b4 . x = (p . x) '&' (q . x) );
begin
:: deftheorem Def1 defines homogeneous MARGREL1:def 22 :
for IT being Relation holds
( IT is homogeneous iff dom IT is with_common_domain );
:: deftheorem Def23 defines quasi_total MARGREL1:def 23 :
for A being set
for IT being PartFunc of (A *),A holds
( IT is quasi_total iff for x, y being FinSequence of A st len x = len y & x in dom IT holds
y in dom IT );
theorem Th54:
theorem
:: deftheorem Def4 defines homogeneous MARGREL1:def 24 :
for A being set
for IT being PFuncFinSequence of A holds
( IT is homogeneous iff for n being Nat
for h being PartFunc of (A *),A st n in dom IT & h = IT . n holds
h is homogeneous );
:: deftheorem Def5 defines quasi_total MARGREL1:def 25 :
for A being set
for IT being PFuncFinSequence of A holds
( IT is quasi_total iff for n being Nat
for h being PartFunc of (A *),A st n in dom IT & h = IT . n holds
h is quasi_total );
theorem
:: deftheorem Def26 defines arity MARGREL1:def 26 :
for f being homogeneous Relation
for b2 being Nat holds
( ( ex x being FinSequence st x in dom f implies ( b2 = arity f iff for x being FinSequence st x in dom f holds
b2 = len x ) ) & ( ( for x being FinSequence holds not x in dom f ) implies ( b2 = arity f iff b2 = 0 ) ) );
begin
theorem
theorem
:: deftheorem Def1 defines PFuncsDomHQN MARGREL1:def 27 :
for D, b2 being non empty set holds
( b2 is PFuncsDomHQN of D iff for x being Element of b2 holds x is non empty homogeneous quasi_total PartFunc of (D *),D );