:: Many-Argument Relations
:: by Edmund Woronowicz
::
:: Received June 1, 1990
:: Copyright (c) 1990-2021 Association of Mizar Users


definition
let IT be FinSequence-membered set ;
redefine attr IT is with_common_domain means :Def1: :: MARGREL1:def 1
for a, b being FinSequence st a in IT & b in IT holds
len a = len b;
compatibility
( IT is with_common_domain iff for a, b being FinSequence st a in IT & b in IT holds
len a = len b )
proof end;
end;

:: 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 );

registration
cluster FinSequence-membered with_common_domain for set ;
existence
ex b1 being set st
( b1 is FinSequence-membered & b1 is with_common_domain )
proof end;
end;

definition
mode relation is FinSequence-membered with_common_domain set ;
end;

theorem :: MARGREL1:1
for X being set
for p being relation st X c= p holds
X is relation ;

theorem :: MARGREL1:2
for a being FinSequence holds {a} is relation
proof end;

scheme :: MARGREL1:sch 1
relexist{ F1() -> set , P1[ FinSequence] } :
ex r being relation st
for a being FinSequence holds
( a in r iff ( a in F1() & P1[a] ) )
provided
A1: for a, b being FinSequence st P1[a] & P1[b] holds
len a = len b
proof end;

definition
let p, r be relation;
redefine pred p = r means :: MARGREL1:def 2
for a being FinSequence holds
( a in p iff a in r );
compatibility
( p = r iff for a being FinSequence holds
( a in p iff a in r ) )
proof end;
end;

:: 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 ) );

registration
cluster empty -> with_common_domain for set ;
coherence
for b1 being set st b1 is empty holds
b1 is with_common_domain
;
end;

theorem :: MARGREL1:3
for p being relation st ( for a being FinSequence holds not a in p ) holds
p = {} ;

definition
let p be relation;
assume A1: p <> {} ;
func the_arity_of p -> Element of NAT means :: MARGREL1:def 3
for a being FinSequence st a in p holds
it = len a;
existence
ex b1 being Element of NAT st
for a being FinSequence st a in p holds
b1 = len a
proof end;
uniqueness
for b1, b2 being Element of NAT st ( for a being FinSequence st a in p holds
b1 = len a ) & ( for a being FinSequence st a in p holds
b2 = len a ) holds
b1 = b2
proof end;
end;

:: deftheorem defines the_arity_of MARGREL1:def 3 :
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 );

definition
let k be Element of NAT ;
mode relation_length of k -> relation means :: MARGREL1:def 4
for a being FinSequence st a in it holds
len a = k;
existence
ex b1 being relation st
for a being FinSequence st a in b1 holds
len a = k
proof end;
end;

:: deftheorem defines relation_length MARGREL1:def 4 :
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 );

definition
let X be set ;
mode relation of X -> relation means :: MARGREL1:def 5
for a being FinSequence st a in it holds
rng a c= X;
existence
ex b1 being relation st
for a being FinSequence st a in b1 holds
rng a c= X
proof end;
end;

:: deftheorem defines relation MARGREL1:def 5 :
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 Th4: :: MARGREL1:4
for X being set holds {} is relation of X
proof end;

theorem Th5: :: MARGREL1:5
for k being Element of NAT holds {} is relation_length of k
proof end;

definition
let X be set ;
let k be Element of NAT ;
mode relation of X,k -> relation means :: MARGREL1:def 6
( it is relation of X & it is relation_length of k );
existence
ex b1 being relation st
( b1 is relation of X & b1 is relation_length of k )
proof end;
end;

:: deftheorem defines relation MARGREL1:def 6 :
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 ) );

definition
let D be non empty set ;
func relations_on D -> set means :Def7: :: MARGREL1:def 7
for X being set holds
( X in it iff ( X c= D * & ( for a, b being FinSequence of D st a in X & b in X holds
len a = len b ) ) );
existence
ex b1 being set st
for X being set holds
( X in b1 iff ( X c= D * & ( for a, b being FinSequence of D st a in X & b in X holds
len a = len b ) ) )
proof end;
uniqueness
for b1, b2 being set st ( for X being set holds
( X in b1 iff ( X c= D * & ( for a, b being FinSequence of D st a in X & b in X holds
len a = len b ) ) ) ) & ( 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 ) ) ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def7 defines relations_on MARGREL1:def 7 :
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 ) ) ) );

registration
let D be non empty set ;
cluster relations_on D -> non empty ;
coherence
not relations_on D is empty
proof end;
end;

definition
let D be non empty set ;
mode relation of D is Element of relations_on D;
end;

theorem :: MARGREL1:6
for D being non empty set
for X being set
for r being Element of relations_on D st X c= r holds
X is Element of relations_on D
proof end;

theorem :: MARGREL1:7
for D being non empty set
for a being FinSequence of D holds {a} is Element of relations_on D
proof end;

theorem :: MARGREL1:8
for D being non empty set
for x, y being Element of D holds {<*x,y*>} is Element of relations_on D
proof end;

definition
let D be non empty set ;
let p, r be Element of relations_on D;
:: original: =
redefine pred p = r means :: MARGREL1:def 8
for a being FinSequence of D holds
( a in p iff a in r );
compatibility
( p = r iff for a being FinSequence of D holds
( a in p iff a in r ) )
proof end;
end;

:: deftheorem defines = MARGREL1:def 8 :
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 ) );

scheme :: MARGREL1:sch 2
relDexist{ F1() -> non empty set , P1[ FinSequence of F1()] } :
ex r being Element of relations_on F1() st
for a being FinSequence of F1() holds
( a in r iff P1[a] )
provided
A1: for a, b being FinSequence of F1() st P1[a] & P1[b] holds
len a = len b
proof end;

definition
let D be non empty set ;
func empty_rel D -> Element of relations_on D means :Def9: :: MARGREL1:def 9
for a being FinSequence of D holds not a in it;
existence
ex b1 being Element of relations_on D st
for a being FinSequence of D holds not a in b1
proof end;
uniqueness
for b1, b2 being Element of relations_on D st ( for a being FinSequence of D holds not a in b1 ) & ( for a being FinSequence of D holds not a in b2 ) holds
b1 = b2
;
end;

:: deftheorem Def9 defines empty_rel MARGREL1:def 9 :
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 :: MARGREL1:9
for D being non empty set holds empty_rel D = {}
proof end;

definition
let D be non empty set ;
let p be Element of relations_on D;
assume A1: p <> empty_rel D ;
func the_arity_of p -> Element of NAT means :: MARGREL1:def 10
for a being FinSequence of D st a in p holds
it = len a;
existence
ex b1 being Element of NAT st
for a being FinSequence of D st a in p holds
b1 = len a
proof end;
uniqueness
for b1, b2 being Element of NAT st ( for a being FinSequence of D st a in p holds
b1 = len a ) & ( for a being FinSequence of D st a in p holds
b2 = len a ) holds
b1 = b2
proof end;
end;

:: deftheorem defines the_arity_of MARGREL1:def 10 :
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 );

scheme :: MARGREL1:sch 3
relDexist2{ F1() -> non empty set , F2() -> Element of NAT , P1[ FinSequence of F1()] } :
ex r being Element of relations_on F1() st
for a being FinSequence of F1() st len a = F2() holds
( a in r iff P1[a] )
proof end;

definition
func BOOLEAN -> set equals :: MARGREL1:def 11
{0,1};
coherence
{0,1} is set
;
end;

:: deftheorem defines BOOLEAN MARGREL1:def 11 :
BOOLEAN = {0,1};

registration
cluster BOOLEAN -> non empty ;
coherence
not BOOLEAN is empty
;
end;

definition
:: original: FALSE
redefine func FALSE -> Element of BOOLEAN ;
coherence
FALSE is Element of BOOLEAN
by TARSKI:def 2;
:: original: TRUE
redefine func TRUE -> Element of BOOLEAN ;
coherence
TRUE is Element of BOOLEAN
by TARSKI:def 2;
end;

definition
let x be object ;
redefine attr x is boolean means :Def12: :: MARGREL1:def 12
x in BOOLEAN ;
compatibility
( x is boolean iff x in BOOLEAN )
by TARSKI:def 2;
end;

:: deftheorem Def12 defines boolean MARGREL1:def 12 :
for x being object holds
( x is boolean iff x in BOOLEAN );

registration
cluster -> boolean for Element of BOOLEAN ;
coherence
for b1 being Element of BOOLEAN holds b1 is boolean
;
end;

definition
let v be boolean object ;
redefine func 'not' v equals :: MARGREL1:def 13
TRUE if v = FALSE
otherwise FALSE ;
compatibility
for b1 being object holds
( ( v = FALSE implies ( b1 = 'not' v iff b1 = TRUE ) ) & ( not v = FALSE implies ( b1 = 'not' v iff b1 = FALSE ) ) )
proof end;
consistency
for b1 being object holds verum
;
let w be boolean object ;
redefine func v '&' w equals :: MARGREL1:def 14
TRUE if ( v = TRUE & w = TRUE )
otherwise FALSE ;
compatibility
for b1 being object holds
( ( v = TRUE & w = TRUE implies ( b1 = v '&' w iff b1 = TRUE ) ) & ( ( not v = TRUE or not w = TRUE ) implies ( b1 = v '&' w iff b1 = FALSE ) ) )
proof end;
consistency
for b1 being object holds verum
;
end;

:: deftheorem defines 'not' MARGREL1:def 13 :
for v being boolean object holds
( ( v = FALSE implies 'not' v = TRUE ) & ( not v = FALSE implies 'not' v = FALSE ) );

:: deftheorem defines '&' MARGREL1:def 14 :
for v, w being boolean object holds
( ( v = TRUE & w = TRUE implies v '&' w = TRUE ) & ( ( not v = TRUE or not w = TRUE ) implies v '&' w = FALSE ) );

definition
let v be Element of BOOLEAN ;
:: original: 'not'
redefine func 'not' v -> Element of BOOLEAN ;
correctness
coherence
'not' v is Element of BOOLEAN
;
by Def12;
let w be Element of BOOLEAN ;
:: original: '&'
redefine func v '&' w -> Element of BOOLEAN ;
correctness
coherence
v '&' w is Element of BOOLEAN
;
by Def12;
end;

theorem :: MARGREL1:10
canceled;

::$CT
theorem :: MARGREL1:11
for v being boolean object holds
( ( v = FALSE implies 'not' v = TRUE ) & ( 'not' v = TRUE implies v = FALSE ) & ( v = TRUE implies 'not' v = FALSE ) & ( 'not' v = FALSE implies v = TRUE ) ) ;

theorem :: MARGREL1:12
for v, w being boolean object holds
( ( v '&' w = TRUE implies ( v = TRUE & w = TRUE ) ) & ( v = TRUE & w = TRUE implies v '&' w = TRUE ) & ( not v '&' w = FALSE or v = FALSE or w = FALSE ) & ( ( v = FALSE or w = FALSE ) implies v '&' w = FALSE ) ) by XBOOLEAN:101, XBOOLEAN:140;

theorem :: MARGREL1:13
for v being boolean object holds FALSE '&' v = FALSE ;

theorem :: MARGREL1:14
for v being boolean object holds TRUE '&' v = v ;

theorem :: MARGREL1:15
for v being boolean object st v '&' v = FALSE holds
v = FALSE ;

theorem :: MARGREL1:16
for u, v, w being boolean object holds v '&' (w '&' u) = (v '&' w) '&' u ;

definition
let X be set ;
func ALL X -> set equals :Def15: :: MARGREL1:def 15
TRUE if not FALSE in X
otherwise FALSE ;
correctness
coherence
( ( not FALSE in X implies TRUE is set ) & ( FALSE in X implies FALSE is set ) )
;
consistency
for b1 being set holds verum
;
;
end;

:: deftheorem Def15 defines ALL MARGREL1:def 15 :
for X being set holds
( ( not FALSE in X implies ALL X = TRUE ) & ( FALSE in X implies ALL X = FALSE ) );

registration
let X be set ;
cluster ALL X -> boolean ;
correctness
coherence
ALL X is boolean
;
by Def15;
end;

definition
let X be set ;
:: original: ALL
redefine func ALL X -> Element of BOOLEAN ;
correctness
coherence
ALL X is Element of BOOLEAN
;
by Def12;
end;

theorem :: MARGREL1:17
for X being set holds
( ( not FALSE in X implies ALL X = TRUE ) & ( ALL X = TRUE implies not FALSE in X ) & ( FALSE in X implies ALL X = FALSE ) & ( ALL X = FALSE implies FALSE in X ) ) by Def15;

:: from VALUAT_1, 2007.03.15, A.T.
definition
let f be Relation;
attr f is boolean-valued means :Def16: :: MARGREL1:def 16
rng f c= BOOLEAN ;
end;

:: deftheorem Def16 defines boolean-valued MARGREL1:def 16 :
for f being Relation holds
( f is boolean-valued iff rng f c= BOOLEAN );

registration
cluster Relation-like Function-like boolean-valued for set ;
existence
ex b1 being Function st b1 is boolean-valued
proof end;
end;

registration
let f be boolean-valued Function;
let x be object ;
cluster f . x -> boolean ;
coherence
f . x is boolean
proof end;
end;

definition
let p be boolean-valued Function;
func 'not' p -> boolean-valued Function means :Def17: :: MARGREL1:def 17
( dom it = dom p & ( for x being object st x in dom p holds
it . x = 'not' (p . x) ) );
existence
ex b1 being boolean-valued Function st
( dom b1 = dom p & ( for x being object st x in dom p holds
b1 . x = 'not' (p . x) ) )
proof end;
uniqueness
for b1, b2 being boolean-valued Function st dom b1 = dom p & ( for x being object st x in dom p holds
b1 . x = 'not' (p . x) ) & dom b2 = dom p & ( for x being object st x in dom p holds
b2 . x = 'not' (p . x) ) holds
b1 = b2
proof end;
involutiveness
for b1, b2 being boolean-valued Function st dom b1 = dom b2 & ( for x being object st x in dom b2 holds
b1 . x = 'not' (b2 . x) ) holds
( dom b2 = dom b1 & ( for x being object st x in dom b1 holds
b2 . x = 'not' (b1 . x) ) )
proof end;
let q be boolean-valued Function;
func p '&' q -> boolean-valued Function means :Def18: :: MARGREL1:def 18
( dom it = (dom p) /\ (dom q) & ( for x being object st x in dom it holds
it . x = (p . x) '&' (q . x) ) );
existence
ex b1 being boolean-valued Function st
( dom b1 = (dom p) /\ (dom q) & ( for x being object st x in dom b1 holds
b1 . x = (p . x) '&' (q . x) ) )
proof end;
uniqueness
for b1, b2 being boolean-valued Function st dom b1 = (dom p) /\ (dom q) & ( for x being object st x in dom b1 holds
b1 . x = (p . x) '&' (q . x) ) & dom b2 = (dom p) /\ (dom q) & ( for x being object st x in dom b2 holds
b2 . x = (p . x) '&' (q . x) ) holds
b1 = b2
proof end;
commutativity
for b1, p, q being boolean-valued Function st dom b1 = (dom p) /\ (dom q) & ( for x being object st x in dom b1 holds
b1 . x = (p . x) '&' (q . x) ) holds
( dom b1 = (dom q) /\ (dom p) & ( for x being object st x in dom b1 holds
b1 . x = (q . x) '&' (p . x) ) )
;
idempotence
for p being boolean-valued Function holds
( dom p = (dom p) /\ (dom p) & ( for x being object st x in dom p holds
p . x = (p . x) '&' (p . x) ) )
;
end;

:: deftheorem Def17 defines 'not' MARGREL1:def 17 :
for p, b2 being boolean-valued Function holds
( b2 = 'not' p iff ( dom b2 = dom p & ( for x being object st x in dom p holds
b2 . x = 'not' (p . x) ) ) );

:: deftheorem Def18 defines '&' MARGREL1:def 18 :
for p, q, b3 being boolean-valued Function holds
( b3 = p '&' q iff ( dom b3 = (dom p) /\ (dom q) & ( for x being object st x in dom b3 holds
b3 . x = (p . x) '&' (q . x) ) ) );

registration
let A be set ;
cluster Function-like quasi_total -> boolean-valued for Element of bool [:A,BOOLEAN:];
coherence
for b1 being Function of A,BOOLEAN holds b1 is boolean-valued
by RELAT_1:def 19;
end;

definition
let A be non empty set ;
let p be Function of A,BOOLEAN;
:: original: 'not'
redefine func 'not' p -> Function of A,BOOLEAN means :: MARGREL1:def 19
for x being Element of A holds it . x = 'not' (p . x);
coherence
'not' p is Function of A,BOOLEAN
proof end;
compatibility
for b1 being Function of A,BOOLEAN holds
( b1 = 'not' p iff for x being Element of A holds b1 . x = 'not' (p . x) )
proof end;
let q be Function of A,BOOLEAN;
:: original: '&'
redefine func p '&' q -> Function of A,BOOLEAN means :: MARGREL1:def 20
for x being Element of A holds it . x = (p . x) '&' (q . x);
coherence
p '&' q is Function of A,BOOLEAN
proof end;
compatibility
for b1 being Function of A,BOOLEAN holds
( b1 = p '&' q iff for x being Element of A holds b1 . x = (p . x) '&' (q . x) )
proof end;
end;

:: deftheorem defines 'not' MARGREL1:def 19 :
for A being non empty set
for p, b3 being Function of A,BOOLEAN holds
( b3 = 'not' p iff for x being Element of A holds b3 . x = 'not' (p . x) );

:: deftheorem defines '&' MARGREL1:def 20 :
for A being non empty set
for p, q, b4 being Function of A,BOOLEAN holds
( b4 = p '&' q iff for x being Element of A holds b4 . x = (p . x) '&' (q . x) );

definition
let IT be Relation;
attr IT is homogeneous means :Def21: :: MARGREL1:def 21
dom IT is with_common_domain ;
end;

:: deftheorem Def21 defines homogeneous MARGREL1:def 21 :
for IT being Relation holds
( IT is homogeneous iff dom IT is with_common_domain );

definition
let A be set ;
let IT be PartFunc of (A *),A;
attr IT is quasi_total means :Def22: :: MARGREL1:def 22
for x, y being FinSequence of A st len x = len y & x in dom IT holds
y in dom IT;
end;

:: deftheorem Def22 defines quasi_total MARGREL1:def 22 :
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 );

registration
let f be Relation;
let X be with_common_domain set ;
cluster f | X -> homogeneous ;
coherence
f | X is homogeneous
proof end;
end;

registration
let A be non empty set ;
let f be PartFunc of (A *),A;
cluster dom f -> FinSequence-membered ;
coherence
dom f is FinSequence-membered
proof end;
end;

registration
let A be non empty set ;
cluster Relation-like A * -defined A -valued Function-like non empty homogeneous quasi_total for Element of bool [:(A *),A:];
existence
ex b1 being PartFunc of (A *),A st
( b1 is homogeneous & b1 is quasi_total & not b1 is empty )
proof end;
end;

registration
cluster Relation-like Function-like non empty homogeneous for set ;
existence
ex b1 being Function st
( b1 is homogeneous & not b1 is empty )
proof end;
end;

registration
let R be homogeneous Relation;
cluster dom R -> with_common_domain ;
coherence
dom R is with_common_domain
by Def21;
end;

theorem Th17: :: MARGREL1:18
for A being non empty set
for a being Element of A holds (<*> A) .--> a is non empty homogeneous quasi_total PartFunc of (A *),A
proof end;

theorem :: MARGREL1:19
for A being non empty set
for a being Element of A holds (<*> A) .--> a is Element of PFuncs ((A *),A)
proof end;

definition
let A be set ;
mode PFuncFinSequence of A is FinSequence of PFuncs ((A *),A);
end;

definition
let A be set ;
let IT be PFuncFinSequence of A;
attr IT is homogeneous means :Def23: :: MARGREL1:def 23
for n being Nat
for h being PartFunc of (A *),A st n in dom IT & h = IT . n holds
h is homogeneous ;
end;

:: deftheorem Def23 defines homogeneous MARGREL1:def 23 :
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 );

definition
let A be set ;
let IT be PFuncFinSequence of A;
attr IT is quasi_total means :: MARGREL1:def 24
for n being Nat
for h being PartFunc of (A *),A st n in dom IT & h = IT . n holds
h is quasi_total ;
end;

:: deftheorem defines quasi_total MARGREL1:def 24 :
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 );

definition
let A be non empty set ;
let x be Element of PFuncs ((A *),A);
:: original: <*
redefine func <*x*> -> PFuncFinSequence of A;
coherence
<*x*> is PFuncFinSequence of A
proof end;
end;

registration
let A be non empty set ;
cluster Relation-like non-empty NAT -defined PFuncs ((A *),A) -valued Function-like Function-yielding V22() V34() FinSequence-like FinSubsequence-like countable homogeneous quasi_total for FinSequence of PFuncs ((A *),A);
existence
ex b1 being PFuncFinSequence of A st
( b1 is homogeneous & b1 is quasi_total & b1 is non-empty )
proof end;
end;

registration
let A be non empty set ;
let f be homogeneous PFuncFinSequence of A;
let i be set ;
cluster f . i -> homogeneous ;
coherence
f . i is homogeneous
proof end;
end;

theorem :: MARGREL1:20
for A being non empty set
for a being Element of A
for x being Element of PFuncs ((A *),A) st x = (<*> A) .--> a holds
( <*x*> is homogeneous & <*x*> is quasi_total & <*x*> is non-empty )
proof end;

definition
let f be homogeneous Relation;
func arity f -> Nat means :Def25: :: MARGREL1:def 25
for x being FinSequence st x in dom f holds
it = len x if ex x being FinSequence st x in dom f
otherwise it = 0 ;
consistency
for b1 being Nat holds verum
;
existence
( ( ex x being FinSequence st x in dom f implies ex b1 being Nat st
for x being FinSequence st x in dom f holds
b1 = len x ) & ( ( for x being FinSequence holds not x in dom f ) implies ex b1 being Nat st b1 = 0 ) )
proof end;
uniqueness
for b1, b2 being Nat holds
( ( ex x being FinSequence st x in dom f & ( for x being FinSequence st x in dom f holds
b1 = len x ) & ( for x being FinSequence st x in dom f holds
b2 = len x ) implies b1 = b2 ) & ( ( for x being FinSequence holds not x in dom f ) & b1 = 0 & b2 = 0 implies b1 = b2 ) )
proof end;
end;

:: deftheorem Def25 defines arity MARGREL1:def 25 :
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 ) ) );

definition
let f be homogeneous Function;
:: original: arity
redefine func arity f -> Element of NAT ;
coherence
arity f is Element of NAT
by ORDINAL1:def 12;
end;

theorem :: MARGREL1:21
for n being Nat
for D being non empty set
for D1 being non empty Subset of D holds (n -tuples_on D) /\ (n -tuples_on D1) = n -tuples_on D1
proof end;

theorem :: MARGREL1:22
for D being non empty set
for h being non empty homogeneous quasi_total PartFunc of (D *),D holds dom h = (arity h) -tuples_on D
proof end;

definition
let D be non empty set ;
mode PFuncsDomHQN of D -> non empty set means :Def26: :: MARGREL1:def 26
for x being Element of it holds x is non empty homogeneous quasi_total PartFunc of (D *),D;
existence
ex b1 being non empty set st
for x being Element of b1 holds x is non empty homogeneous quasi_total PartFunc of (D *),D
proof end;
end;

:: deftheorem Def26 defines PFuncsDomHQN MARGREL1:def 26 :
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 );

definition
let D be non empty set ;
let P be PFuncsDomHQN of D;
:: original: Element
redefine mode Element of P -> non empty homogeneous quasi_total PartFunc of (D *),D;
coherence
for b1 being Element of P holds b1 is non empty homogeneous quasi_total PartFunc of (D *),D
by Def26;
end;