:: FUNCT_2 semantic presentation

definition
let c1, c2 be set ;
let c3 be Relation of c1,c2;
attr a3 is quasi_total means :Def1: :: FUNCT_2:def 1
a1 = dom a3 if ( a2 = {} implies a1 = {} )
otherwise a3 = {} ;
consistency
verum
;
end;

:: deftheorem Def1 defines quasi_total FUNCT_2:def 1 :
for b1, b2 being set
for b3 being Relation of b1,b2 holds
( ( ( b2 = {} implies b1 = {} ) implies ( b3 is quasi_total iff b1 = dom b3 ) ) & ( b2 = {} implies b1 = {} ( b3 is quasi_total iff b3 = {} ) ) );

registration
let c1, c2 be set ;
cluster Function-like quasi_total Relation of a1,a2;
existence
ex b1 being Relation of c1,c2 st
( b1 is quasi_total & b1 is Function-like )
proof end;
end;

registration
let c1, c2 be set ;
cluster total -> quasi_total Relation of a1,a2;
coherence
for b1 being PartFunc of c1,c2 holds
( b1 is total implies b1 is quasi_total )
proof end;
end;

definition
let c1, c2 be set ;
mode Function is Function-like quasi_total Relation of a1,a2;
end;

theorem Th1: :: FUNCT_2:1
canceled;

theorem Th2: :: FUNCT_2:2
canceled;

theorem Th3: :: FUNCT_2:3
for b1 being Function holds
b1 is Function of dom b1, rng b1
proof end;

theorem Th4: :: FUNCT_2:4
for b1 being set
for b2 being Function holds
( rng b2 c= b1 implies b2 is Function of dom b2,b1 )
proof end;

theorem Th5: :: FUNCT_2:5
for b1, b2 being set
for b3 being Function holds
( dom b3 = b1 & ( for b4 being set holds
( b4 in b1 implies b3 . b4 in b2 ) ) implies b3 is Function of b1,b2 )
proof end;

theorem Th6: :: FUNCT_2:6
for b1, b2, b3 being set
for b4 being Function of b1,b2 holds
( b2 <> {} & b3 in b1 implies b4 . b3 in rng b4 )
proof end;

theorem Th7: :: FUNCT_2:7
for b1, b2, b3 being set
for b4 being Function of b1,b2 holds
( b2 <> {} & b3 in b1 implies b4 . b3 in b2 )
proof end;

theorem Th8: :: FUNCT_2:8
for b1, b2, b3 being set
for b4 being Function of b1,b2 holds
( ( b2 = {} implies b1 = {} ) & rng b4 c= b3 implies b4 is Function of b1,b3 )
proof end;

theorem Th9: :: FUNCT_2:9
for b1, b2, b3 being set
for b4 being Function of b1,b2 holds
( ( b2 = {} implies b1 = {} ) & b2 c= b3 implies b4 is Function of b1,b3 )
proof end;

scheme :: FUNCT_2:sch 1
s1{ F1() -> set , F2() -> set , P1[ set , set ] } :
ex b1 being Function of F1(),F2() st
for b2 being set holds
( b2 in F1() implies P1[b2,b1 . b2] )
provided
E5: for b1 being set holds
not ( b1 in F1() & ( for b2 being set holds
not ( b2 in F2() & P1[b1,b2] ) ) )
proof end;

scheme :: FUNCT_2:sch 2
s2{ F1() -> set , F2() -> set , F3( set ) -> set } :
ex b1 being Function of F1(),F2() st
for b2 being set holds
( b2 in F1() implies b1 . b2 = F3(b2) )
provided
E5: for b1 being set holds
( b1 in F1() implies F3(b1) in F2() )
proof end;

definition
let c1, c2 be set ;
func Funcs c1,c2 -> set means :Def2: :: FUNCT_2:def 2
for b1 being set holds
( b1 in a3 iff ex b2 being Function st
( b1 = b2 & dom b2 = a1 & rng b2 c= a2 ) );
existence
ex b1 being set st
for b2 being set holds
( b2 in b1 iff ex b3 being Function st
( b2 = b3 & dom b3 = c1 & rng b3 c= c2 ) )
proof end;
uniqueness
for b1, b2 being set holds
( ( for b3 being set holds
( b3 in b1 iff ex b4 being Function st
( b3 = b4 & dom b4 = c1 & rng b4 c= c2 ) ) ) & ( for b3 being set holds
( b3 in b2 iff ex b4 being Function st
( b3 = b4 & dom b4 = c1 & rng b4 c= c2 ) ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def2 defines Funcs FUNCT_2:def 2 :
for b1, b2, b3 being set holds
( b3 = Funcs b1,b2 iff for b4 being set holds
( b4 in b3 iff ex b5 being Function st
( b4 = b5 & dom b5 = b1 & rng b5 c= b2 ) ) );

theorem Th10: :: FUNCT_2:10
canceled;

theorem Th11: :: FUNCT_2:11
for b1, b2 being set
for b3 being Function of b1,b2 holds
( ( b2 = {} implies b1 = {} ) implies b3 in Funcs b1,b2 )
proof end;

theorem Th12: :: FUNCT_2:12
for b1 being set
for b2 being Function of b1,b1 holds b2 in Funcs b1,b1
proof end;

registration
let c1 be set ;
let c2 be non empty set ;
cluster Funcs a1,a2 -> non empty ;
coherence
not Funcs c1,c2 is empty
by Th11;
end;

registration
let c1 be set ;
cluster Funcs a1,a1 -> non empty ;
coherence
not Funcs c1,c1 is empty
by Th12;
end;

theorem Th13: :: FUNCT_2:13
canceled;

theorem Th14: :: FUNCT_2:14
for b1 being set holds
( b1 <> {} implies Funcs b1,{} = {} )
proof end;

theorem Th15: :: FUNCT_2:15
canceled;

theorem Th16: :: FUNCT_2:16
for b1, b2 being set
for b3 being Function of b1,b2 holds
( b2 <> {} & ( for b4 being set holds
not ( b4 in b2 & ( for b5 being set holds
not ( b5 in b1 & b4 = b3 . b5 ) ) ) ) implies rng b3 = b2 )
proof end;

theorem Th17: :: FUNCT_2:17
for b1, b2, b3 being set
for b4 being Function of b1,b2 holds
not ( b3 in rng b4 & ( for b5 being set holds
not ( b5 in b1 & b4 . b5 = b3 ) ) )
proof end;

theorem Th18: :: FUNCT_2:18
for b1, b2 being set
for b3, b4 being Function of b1,b2 holds
( ( for b5 being set holds
( b5 in b1 implies b3 . b5 = b4 . b5 ) ) implies b3 = b4 )
proof end;

theorem Th19: :: FUNCT_2:19
for b1, b2, b3 being set
for b4 being quasi_total Relation of b1,b2
for b5 being quasi_total Relation of b2,b3 holds
( not ( b2 = {} & not b3 = {} & not b1 = {} ) implies b4 * b5 is quasi_total )
proof end;

theorem Th20: :: FUNCT_2:20
for b1, b2, b3 being set
for b4 being Function of b1,b2
for b5 being Function of b2,b3 holds
( b2 <> {} & b3 <> {} & rng b4 = b2 & rng b5 = b3 implies rng (b5 * b4) = b3 )
proof end;

theorem Th21: :: FUNCT_2:21
for b1, b2, b3 being set
for b4 being Function of b1,b2
for b5 being Function holds
( b2 <> {} & b3 in b1 implies (b5 * b4) . b3 = b5 . (b4 . b3) )
proof end;

theorem Th22: :: FUNCT_2:22
for b1, b2 being set
for b3 being Function of b1,b2 holds
( b2 <> {} implies ( rng b3 = b2 iff for b4 being set holds
( b4 <> {} implies for b5, b6 being Function of b2,b4 holds
( b5 * b3 = b6 * b3 implies b5 = b6 ) ) ) )
proof end;

theorem Th23: :: FUNCT_2:23
for b1, b2 being set
for b3 being Relation of b1,b2 holds
( (id b1) * b3 = b3 & b3 * (id b2) = b3 )
proof end;

theorem Th24: :: FUNCT_2:24
for b1, b2 being set
for b3 being Function of b1,b2
for b4 being Function of b2,b1 holds
( b3 * b4 = id b2 implies rng b3 = b2 )
proof end;

theorem Th25: :: FUNCT_2:25
for b1, b2 being set
for b3 being Function of b1,b2 holds
( ( b2 = {} implies b1 = {} ) implies ( b3 is one-to-one iff for b4, b5 being set holds
( b4 in b1 & b5 in b1 & b3 . b4 = b3 . b5 implies b4 = b5 ) ) )
proof end;

theorem Th26: :: FUNCT_2:26
for b1, b2, b3 being set
for b4 being Function of b1,b2
for b5 being Function of b2,b3 holds
( ( b3 = {} implies b2 = {} ) & ( b2 = {} implies b1 = {} ) & b5 * b4 is one-to-one implies b4 is one-to-one )
proof end;

theorem Th27: :: FUNCT_2:27
for b1, b2 being set
for b3 being Function of b1,b2 holds
( b1 <> {} & b2 <> {} implies ( b3 is one-to-one iff for b4 being set
for b5, b6 being Function of b4,b1 holds
( b3 * b5 = b3 * b6 implies b5 = b6 ) ) )
proof end;

theorem Th28: :: FUNCT_2:28
for b1, b2, b3 being set
for b4 being Function of b1,b2
for b5 being Function of b2,b3 holds
( b3 <> {} & rng (b5 * b4) = b3 & b5 is one-to-one implies rng b4 = b2 )
proof end;

theorem Th29: :: FUNCT_2:29
for b1, b2 being set
for b3 being Function of b1,b2
for b4 being Function of b2,b1 holds
( b2 <> {} & b4 * b3 = id b1 implies ( b3 is one-to-one & rng b4 = b1 ) )
proof end;

theorem Th30: :: FUNCT_2:30
for b1, b2, b3 being set
for b4 being Function of b1,b2
for b5 being Function of b2,b3 holds
( ( b3 = {} implies b2 = {} ) & b5 * b4 is one-to-one & rng b4 = b2 implies ( b4 is one-to-one & b5 is one-to-one ) )
proof end;

theorem Th31: :: FUNCT_2:31
for b1, b2 being set
for b3 being Function of b1,b2 holds
( b3 is one-to-one & rng b3 = b2 implies b3 " is Function of b2,b1 )
proof end;

theorem Th32: :: FUNCT_2:32
for b1, b2, b3 being set
for b4 being Function of b1,b2 holds
( b2 <> {} & b4 is one-to-one & b3 in b1 implies (b4 " ) . (b4 . b3) = b3 )
proof end;

theorem Th33: :: FUNCT_2:33
canceled;

theorem Th34: :: FUNCT_2:34
for b1, b2 being set
for b3 being Function of b1,b2
for b4 being Function of b2,b1 holds
( b1 <> {} & b2 <> {} & rng b3 = b2 & b3 is one-to-one & ( for b5, b6 being set holds
( ( b5 in b2 & b4 . b5 = b6 implies ( b6 in b1 & b3 . b6 = b5 ) ) & ( b6 in b1 & b3 . b6 = b5 implies ( b5 in b2 & b4 . b5 = b6 ) ) ) ) implies b4 = b3 " )
proof end;

theorem Th35: :: FUNCT_2:35
for b1, b2 being set
for b3 being Function of b1,b2 holds
( b2 <> {} & rng b3 = b2 & b3 is one-to-one implies ( (b3 " ) * b3 = id b1 & b3 * (b3 " ) = id b2 ) )
proof end;

theorem Th36: :: FUNCT_2:36
for b1, b2 being set
for b3 being Function of b1,b2
for b4 being Function of b2,b1 holds
( b1 <> {} & b2 <> {} & rng b3 = b2 & b4 * b3 = id b1 & b3 is one-to-one implies b4 = b3 " )
proof end;

theorem Th37: :: FUNCT_2:37
for b1, b2 being set
for b3 being Function of b1,b2 holds
( b2 <> {} & ex b4 being Function of b2,b1 st b4 * b3 = id b1 implies b3 is one-to-one )
proof end;

theorem Th38: :: FUNCT_2:38
for b1, b2, b3 being set
for b4 being Function of b1,b2 holds
( ( b2 = {} implies b1 = {} ) & b3 c= b1 implies b4 | b3 is Function of b3,b2 )
proof end;

theorem Th39: :: FUNCT_2:39
canceled;

theorem Th40: :: FUNCT_2:40
for b1, b2, b3 being set
for b4 being Function of b1,b2 holds
( b1 c= b3 implies b4 | b3 = b4 )
proof end;

theorem Th41: :: FUNCT_2:41
for b1, b2, b3, b4 being set
for b5 being Function of b1,b2 holds
( b2 <> {} & b3 in b1 & b5 . b3 in b4 implies (b4 | b5) . b3 = b5 . b3 )
proof end;

theorem Th42: :: FUNCT_2:42
canceled;

theorem Th43: :: FUNCT_2:43
for b1, b2, b3 being set
for b4 being Function of b1,b2 holds
( b2 <> {} implies for b5 being set holds
( ex b6 being set st
( b6 in b1 & b6 in b3 & b5 = b4 . b6 ) implies b5 in b4 .: b3 ) )
proof end;

theorem Th44: :: FUNCT_2:44
for b1, b2, b3 being set
for b4 being Function of b1,b2 holds b4 .: b3 c= b2 ;

theorem Th45: :: FUNCT_2:45
for b1, b2 being set
for b3 being Function of b1,b2 holds
( ( b2 = {} implies b1 = {} ) implies b3 .: b1 = rng b3 )
proof end;

theorem Th46: :: FUNCT_2:46
for b1, b2, b3 being set
for b4 being Function of b1,b2 holds
( b2 <> {} implies for b5 being set holds
( b5 in b4 " b3 iff ( b5 in b1 & b4 . b5 in b3 ) ) )
proof end;

theorem Th47: :: FUNCT_2:47
for b1, b2, b3 being set
for b4 being PartFunc of b1,b2 holds b4 " b3 c= b1 ;

theorem Th48: :: FUNCT_2:48
for b1, b2 being set
for b3 being Function of b1,b2 holds
( ( b2 = {} implies b1 = {} ) implies b3 " b2 = b1 )
proof end;

theorem Th49: :: FUNCT_2:49
for b1, b2 being set
for b3 being Function of b1,b2 holds
( ( for b4 being set holds
not ( b4 in b2 & not b3 " {b4} <> {} ) ) iff rng b3 = b2 )
proof end;

theorem Th50: :: FUNCT_2:50
for b1, b2, b3 being set
for b4 being Function of b1,b2 holds
( ( b2 = {} implies b1 = {} ) & b3 c= b1 implies b3 c= b4 " (b4 .: b3) )
proof end;

theorem Th51: :: FUNCT_2:51
for b1, b2 being set
for b3 being Function of b1,b2 holds
( ( b2 = {} implies b1 = {} ) implies b3 " (b3 .: b1) = b1 )
proof end;

theorem Th52: :: FUNCT_2:52
canceled;

theorem Th53: :: FUNCT_2:53
for b1, b2, b3, b4 being set
for b5 being Function of b1,b2
for b6 being Function of b2,b3 holds
( ( b3 = {} implies b2 = {} ) & ( b2 = {} implies b1 = {} ) implies b5 " b4 c= (b6 * b5) " (b6 .: b4) )
proof end;

theorem Th54: :: FUNCT_2:54
canceled;

theorem Th55: :: FUNCT_2:55
for b1 being set
for b2 being Function holds
( dom b2 = {} implies b2 is Function of {} ,b1 )
proof end;

theorem Th56: :: FUNCT_2:56
canceled;

theorem Th57: :: FUNCT_2:57
canceled;

theorem Th58: :: FUNCT_2:58
canceled;

theorem Th59: :: FUNCT_2:59
for b1, b2 being set
for b3 being Function of {} ,b1 holds b3 .: b2 = {}
proof end;

theorem Th60: :: FUNCT_2:60
for b1, b2 being set
for b3 being Function of {} ,b1 holds b3 " b2 = {}
proof end;

theorem Th61: :: FUNCT_2:61
for b1, b2 being set
for b3 being Function of {b1},b2 holds
( b2 <> {} implies b3 . b1 in b2 )
proof end;

theorem Th62: :: FUNCT_2:62
for b1, b2 being set
for b3 being Function of {b1},b2 holds
( b2 <> {} implies rng b3 = {(b3 . b1)} )
proof end;

theorem Th63: :: FUNCT_2:63
canceled;

theorem Th64: :: FUNCT_2:64
for b1, b2, b3 being set
for b4 being Function of {b1},b2 holds
( b2 <> {} implies b4 .: b3 c= {(b4 . b1)} )
proof end;

theorem Th65: :: FUNCT_2:65
for b1, b2, b3 being set
for b4 being Function of b1,{b2} holds
( b3 in b1 implies b4 . b3 = b2 )
proof end;

theorem Th66: :: FUNCT_2:66
for b1, b2 being set
for b3, b4 being Function of b1,{b2} holds b3 = b4
proof end;

registration
let c1 be set ;
let c2, c3 be quasi_total PartFunc of c1,c1;
cluster a2 * a3 -> quasi_total ;
coherence
c3 * c2 is quasi_total PartFunc of c1,c1
proof end;
end;

theorem Th67: :: FUNCT_2:67
for b1 being set
for b2 being Function of b1,b1 holds dom b2 = b1
proof end;

theorem Th68: :: FUNCT_2:68
canceled;

theorem Th69: :: FUNCT_2:69
canceled;

theorem Th70: :: FUNCT_2:70
for b1, b2 being set
for b3 being Function of b1,b1
for b4 being Function holds
( b2 in b1 implies (b4 * b3) . b2 = b4 . (b3 . b2) )
proof end;

theorem Th71: :: FUNCT_2:71
canceled;

theorem Th72: :: FUNCT_2:72
canceled;

theorem Th73: :: FUNCT_2:73
for b1 being set
for b2, b3 being Relation of b1,b1 holds
( rng b2 = b1 & rng b3 = b1 implies rng (b3 * b2) = b1 )
proof end;

theorem Th74: :: FUNCT_2:74
canceled;

theorem Th75: :: FUNCT_2:75
for b1 being set
for b2, b3 being Function of b1,b1 holds
( b3 * b2 = b2 & rng b2 = b1 implies b3 = id b1 )
proof end;

theorem Th76: :: FUNCT_2:76
for b1 being set
for b2, b3 being Function of b1,b1 holds
( b2 * b3 = b2 & b2 is one-to-one implies b3 = id b1 )
proof end;

theorem Th77: :: FUNCT_2:77
for b1 being set
for b2 being Function of b1,b1 holds
( b2 is one-to-one iff for b3, b4 being set holds
( b3 in b1 & b4 in b1 & b2 . b3 = b2 . b4 implies b3 = b4 ) )
proof end;

theorem Th78: :: FUNCT_2:78
canceled;

theorem Th79: :: FUNCT_2:79
for b1 being set
for b2 being Function of b1,b1 holds b2 .: b1 = rng b2
proof end;

theorem Th80: :: FUNCT_2:80
canceled;

theorem Th81: :: FUNCT_2:81
canceled;

theorem Th82: :: FUNCT_2:82
for b1 being set
for b2 being Function of b1,b1 holds b2 " (b2 .: b1) = b1
proof end;

definition
let c1, c2 be set ;
let c3 be Relation of c1,c2;
attr a3 is onto means :Def3: :: FUNCT_2:def 3
rng a3 = a2;
end;

:: deftheorem Def3 defines onto FUNCT_2:def 3 :
for b1, b2 being set
for b3 being Relation of b1,b2 holds
( b3 is onto iff rng b3 = b2 );

definition
let c1, c2 be set ;
let c3 be Function of c1,c2;
attr a3 is bijective means :Def4: :: FUNCT_2:def 4
( a3 is one-to-one & a3 is onto );
end;

:: deftheorem Def4 defines bijective FUNCT_2:def 4 :
for b1, b2 being set
for b3 being Function of b1,b2 holds
( b3 is bijective iff ( b3 is one-to-one & b3 is onto ) );

registration
let c1, c2 be set ;
cluster bijective -> one-to-one onto Relation of a1,a2;
coherence
for b1 being Function of c1,c2 holds
( b1 is bijective implies ( b1 is one-to-one & b1 is onto ) )
by Def4;
cluster one-to-one onto -> bijective Relation of a1,a2;
coherence
for b1 being Function of c1,c2 holds
( b1 is one-to-one & b1 is onto implies b1 is bijective )
by Def4;
end;

registration
let c1 be set ;
cluster one-to-one onto bijective Relation of a1,a1;
existence
ex b1 being Function of c1,c1 st b1 is bijective
proof end;
end;

definition
let c1 be set ;
mode Permutation is bijective Function of a1,a1;
end;

theorem Th83: :: FUNCT_2:83
for b1 being set
for b2 being Function of b1,b1 holds
( b2 is one-to-one & rng b2 = b1 implies b2 is Permutation of b1 )
proof end;

theorem Th84: :: FUNCT_2:84
canceled;

theorem Th85: :: FUNCT_2:85
for b1 being set
for b2 being Function of b1,b1 holds
( b2 is one-to-one implies for b3, b4 being set holds
( b3 in b1 & b4 in b1 & b2 . b3 = b2 . b4 implies b3 = b4 ) ) by Th77;

registration
let c1 be set ;
let c2, c3 be onto PartFunc of c1,c1;
cluster a3 * a2 -> onto ;
coherence
c2 * c3 is onto PartFunc of c1,c1
proof end;
end;

registration
let c1, c2 be one-to-one Function;
cluster a1 * a2 -> one-to-one ;
coherence
c2 * c1 is one-to-one Function
by FUNCT_1:46;
end;

registration
let c1 be set ;
let c2, c3 be bijective Function of c1,c1;
cluster a2 * a3 -> one-to-one onto bijective ;
coherence
c3 * c2 is bijective Function of c1,c1
;
end;

registration
let c1 be set ;
cluster reflexive total -> one-to-one onto bijective Relation of a1,a1;
coherence
for b1 being Function of c1,c1 holds
( b1 is reflexive & b1 is total implies b1 is bijective )
proof end;
end;

definition
let c1 be set ;
let c2 be Permutation of c1;
redefine func " as c2 " -> Permutation of a1;
coherence
c2 " is Permutation of c1
proof end;
end;

theorem Th86: :: FUNCT_2:86
for b1 being set
for b2, b3 being Permutation of b1 holds
( b3 * b2 = b3 implies b2 = id b1 )
proof end;

theorem Th87: :: FUNCT_2:87
for b1 being set
for b2, b3 being Permutation of b1 holds
( b3 * b2 = id b1 implies b3 = b2 " )
proof end;

theorem Th88: :: FUNCT_2:88
for b1 being set
for b2 being Permutation of b1 holds
( (b2 " ) * b2 = id b1 & b2 * (b2 " ) = id b1 )
proof end;

theorem Th89: :: FUNCT_2:89
canceled;

theorem Th90: :: FUNCT_2:90
canceled;

theorem Th91: :: FUNCT_2:91
canceled;

theorem Th92: :: FUNCT_2:92
for b1, b2 being set
for b3 being Permutation of b1 holds
( b2 c= b1 implies ( b3 .: (b3 " b2) = b2 & b3 " (b3 .: b2) = b2 ) )
proof end;

registration
let c1 be set ;
let c2 be non empty set ;
cluster quasi_total -> total quasi_total Relation of a1,a2;
coherence
for b1 being PartFunc of c1,c2 holds
( b1 is quasi_total implies b1 is total )
proof end;
end;

registration
let c1 be set ;
let c2 be non empty set ;
let c3 be set ;
let c4 be Function of c1,c2;
let c5 be Function of c2,c3;
cluster a4 * a5 -> quasi_total ;
coherence
c5 * c4 is quasi_total PartFunc of c1,c3
by Th19;
end;

definition
let c1 be non empty set ;
let c2 be set ;
let c3 be Function of c1,c2;
let c4 be Element of c1;
redefine func . as c3 . c4 -> Element of a2;
coherence
c3 . c4 is Element of c2
proof end;
end;

scheme :: FUNCT_2:sch 3
s3{ F1() -> non empty set , F2() -> non empty set , P1[ set , set ] } :
ex b1 being Function of F1(),F2() st
for b2 being Element of F1() holds P1[b2,b1 . b2]
provided
E24: for b1 being Element of F1() holds
ex b2 being Element of F2() st P1[b1,b2]
proof end;

scheme :: FUNCT_2:sch 4
s4{ F1() -> non empty set , F2() -> non empty set , F3( Element of F1()) -> Element of F2() } :
ex b1 being Function of F1(),F2() st
for b2 being Element of F1() holds b1 . b2 = F3(b2)
proof end;

theorem Th93: :: FUNCT_2:93
canceled;

theorem Th94: :: FUNCT_2:94
canceled;

theorem Th95: :: FUNCT_2:95
canceled;

theorem Th96: :: FUNCT_2:96
canceled;

theorem Th97: :: FUNCT_2:97
canceled;

theorem Th98: :: FUNCT_2:98
canceled;

theorem Th99: :: FUNCT_2:99
canceled;

theorem Th100: :: FUNCT_2:100
canceled;

theorem Th101: :: FUNCT_2:101
canceled;

theorem Th102: :: FUNCT_2:102
canceled;

theorem Th103: :: FUNCT_2:103
canceled;

theorem Th104: :: FUNCT_2:104
canceled;

theorem Th105: :: FUNCT_2:105
canceled;

theorem Th106: :: FUNCT_2:106
canceled;

theorem Th107: :: FUNCT_2:107
canceled;

theorem Th108: :: FUNCT_2:108
canceled;

theorem Th109: :: FUNCT_2:109
canceled;

theorem Th110: :: FUNCT_2:110
canceled;

theorem Th111: :: FUNCT_2:111
canceled;

theorem Th112: :: FUNCT_2:112
canceled;

theorem Th113: :: FUNCT_2:113
for b1, b2 being set
for b3, b4 being Function of b1,b2 holds
( ( for b5 being Element of b1 holds b3 . b5 = b4 . b5 ) implies b3 = b4 )
proof end;

theorem Th114: :: FUNCT_2:114
canceled;

theorem Th115: :: FUNCT_2:115
for b1, b2, b3 being set
for b4 being Function of b1,b2
for b5 being set holds
not ( b5 in b4 .: b3 & ( for b6 being set holds
not ( b6 in b1 & b6 in b3 & b5 = b4 . b6 ) ) )
proof end;

theorem Th116: :: FUNCT_2:116
for b1, b2, b3 being set
for b4 being Function of b1,b2
for b5 being set holds
not ( b5 in b4 .: b3 & ( for b6 being Element of b1 holds
not ( b6 in b3 & b5 = b4 . b6 ) ) )
proof end;

theorem Th117: :: FUNCT_2:117
canceled;

theorem Th118: :: FUNCT_2:118
canceled;

theorem Th119: :: FUNCT_2:119
canceled;

theorem Th120: :: FUNCT_2:120
canceled;

theorem Th121: :: FUNCT_2:121
for b1, b2, b3 being set holds
( b3 in Funcs b1,b2 implies b3 is Function of b1,b2 )
proof end;

scheme :: FUNCT_2:sch 5
s5{ F1() -> set , F2() -> set , P1[ set ], F3( set ) -> set , F4( set ) -> set } :
ex b1 being Function of F1(),F2() st
for b2 being set holds
( b2 in F1() implies ( ( P1[b2] implies b1 . b2 = F3(b2) ) & ( not P1[b2] implies b1 . b2 = F4(b2) ) ) )
provided
E27: for b1 being set holds
( b1 in F1() implies ( ( P1[b1] implies F3(b1) in F2() ) & ( not P1[b1] implies F4(b1) in F2() ) ) )
proof end;

theorem Th122: :: FUNCT_2:122
canceled;

theorem Th123: :: FUNCT_2:123
canceled;

theorem Th124: :: FUNCT_2:124
canceled;

theorem Th125: :: FUNCT_2:125
canceled;

theorem Th126: :: FUNCT_2:126
canceled;

theorem Th127: :: FUNCT_2:127
canceled;

theorem Th128: :: FUNCT_2:128
canceled;

theorem Th129: :: FUNCT_2:129
canceled;

theorem Th130: :: FUNCT_2:130
for b1, b2 being set
for b3 being PartFunc of b1,b2 holds
( dom b3 = b1 implies b3 is Function of b1,b2 )
proof end;

theorem Th131: :: FUNCT_2:131
for b1, b2 being set
for b3 being PartFunc of b1,b2 holds
( b3 is total implies b3 is Function of b1,b2 )
proof end;

theorem Th132: :: FUNCT_2:132
for b1, b2 being set
for b3 being PartFunc of b1,b2 holds
( ( b2 = {} implies b1 = {} ) & b3 is Function of b1,b2 implies b3 is total )
proof end;

theorem Th133: :: FUNCT_2:133
for b1, b2 being set
for b3 being Function of b1,b2 holds
( ( b2 = {} implies b1 = {} ) implies <:b3,b1,b2:> is total )
proof end;

theorem Th134: :: FUNCT_2:134
for b1 being set
for b2 being Function of b1,b1 holds <:b2,b1,b1:> is total
proof end;

theorem Th135: :: FUNCT_2:135
canceled;

theorem Th136: :: FUNCT_2:136
for b1, b2 being set
for b3 being PartFunc of b1,b2 holds
not ( ( b2 = {} implies b1 = {} ) & ( for b4 being Function of b1,b2 holds
ex b5 being set st
( b5 in dom b3 & not b4 . b5 = b3 . b5 ) ) )
proof end;

theorem Th137: :: FUNCT_2:137
canceled;

theorem Th138: :: FUNCT_2:138
canceled;

theorem Th139: :: FUNCT_2:139
canceled;

theorem Th140: :: FUNCT_2:140
canceled;

theorem Th141: :: FUNCT_2:141
for b1, b2 being set holds Funcs b1,b2 c= PFuncs b1,b2
proof end;

theorem Th142: :: FUNCT_2:142
for b1, b2 being set
for b3, b4 being Function of b1,b2 holds
( ( b2 = {} implies b1 = {} ) & b3 tolerates b4 implies b3 = b4 )
proof end;

theorem Th143: :: FUNCT_2:143
for b1 being set
for b2, b3 being Function of b1,b1 holds
( b2 tolerates b3 implies b2 = b3 )
proof end;

theorem Th144: :: FUNCT_2:144
canceled;

theorem Th145: :: FUNCT_2:145
for b1, b2 being set
for b3 being PartFunc of b1,b2
for b4 being Function of b1,b2 holds
( ( b2 = {} implies b1 = {} ) implies ( b3 tolerates b4 iff for b5 being set holds
( b5 in dom b3 implies b3 . b5 = b4 . b5 ) ) )
proof end;

theorem Th146: :: FUNCT_2:146
for b1 being set
for b2 being PartFunc of b1,b1
for b3 being Function of b1,b1 holds
( b2 tolerates b3 iff for b4 being set holds
( b4 in dom b2 implies b2 . b4 = b3 . b4 ) )
proof end;

theorem Th147: :: FUNCT_2:147
canceled;

theorem Th148: :: FUNCT_2:148
for b1, b2 being set
for b3 being PartFunc of b1,b2 holds
not ( ( b2 = {} implies b1 = {} ) & ( for b4 being Function of b1,b2 holds
not b3 tolerates b4 ) )
proof end;

theorem Th149: :: FUNCT_2:149
for b1 being set
for b2 being PartFunc of b1,b1 holds
ex b3 being Function of b1,b1 st b2 tolerates b3
proof end;

theorem Th150: :: FUNCT_2:150
canceled;

theorem Th151: :: FUNCT_2:151
for b1, b2 being set
for b3, b4 being PartFunc of b1,b2
for b5 being Function of b1,b2 holds
( ( b2 = {} implies b1 = {} ) & b3 tolerates b5 & b4 tolerates b5 implies b3 tolerates b4 )
proof end;

theorem Th152: :: FUNCT_2:152
for b1 being set
for b2, b3 being PartFunc of b1,b1
for b4 being Function of b1,b1 holds
( b2 tolerates b4 & b3 tolerates b4 implies b2 tolerates b3 )
proof end;

theorem Th153: :: FUNCT_2:153
canceled;

theorem Th154: :: FUNCT_2:154
for b1, b2 being set
for b3, b4 being PartFunc of b1,b2 holds
not ( ( b2 = {} implies b1 = {} ) & b3 tolerates b4 & ( for b5 being Function of b1,b2 holds
not ( b3 tolerates b5 & b4 tolerates b5 ) ) )
proof end;

theorem Th155: :: FUNCT_2:155
for b1, b2 being set
for b3 being PartFunc of b1,b2
for b4 being Function of b1,b2 holds
( ( b2 = {} implies b1 = {} ) & b3 tolerates b4 implies b4 in TotFuncs b3 )
proof end;

theorem Th156: :: FUNCT_2:156
for b1 being set
for b2 being PartFunc of b1,b1
for b3 being Function of b1,b1 holds
( b2 tolerates b3 implies b3 in TotFuncs b2 )
proof end;

theorem Th157: :: FUNCT_2:157
canceled;

theorem Th158: :: FUNCT_2:158
for b1, b2 being set
for b3 being PartFunc of b1,b2
for b4 being set holds
( b4 in TotFuncs b3 implies b4 is Function of b1,b2 )
proof end;

theorem Th159: :: FUNCT_2:159
for b1, b2 being set
for b3 being PartFunc of b1,b2 holds TotFuncs b3 c= Funcs b1,b2
proof end;

theorem Th160: :: FUNCT_2:160
for b1, b2 being set holds TotFuncs <:{} ,b1,b2:> = Funcs b1,b2
proof end;

theorem Th161: :: FUNCT_2:161
for b1, b2 being set
for b3 being Function of b1,b2 holds
( ( b2 = {} implies b1 = {} ) implies TotFuncs <:b3,b1,b2:> = {b3} )
proof end;

theorem Th162: :: FUNCT_2:162
for b1 being set
for b2 being Function of b1,b1 holds TotFuncs <:b2,b1,b1:> = {b2}
proof end;

theorem Th163: :: FUNCT_2:163
canceled;

theorem Th164: :: FUNCT_2:164
for b1, b2 being set
for b3 being PartFunc of b1,{b2}
for b4 being Function of b1,{b2} holds TotFuncs b3 = {b4}
proof end;

theorem Th165: :: FUNCT_2:165
for b1, b2 being set
for b3, b4 being PartFunc of b1,b2 holds
( b4 c= b3 implies TotFuncs b3 c= TotFuncs b4 )
proof end;

theorem Th166: :: FUNCT_2:166
for b1, b2 being set
for b3, b4 being PartFunc of b1,b2 holds
( dom b4 c= dom b3 & TotFuncs b3 c= TotFuncs b4 implies b4 c= b3 )
proof end;

theorem Th167: :: FUNCT_2:167
for b1, b2 being set
for b3, b4 being PartFunc of b1,b2 holds
( TotFuncs b3 c= TotFuncs b4 & ( for b5 being set holds
b2 <> {b5} ) implies b4 c= b3 )
proof end;

theorem Th168: :: FUNCT_2:168
for b1, b2 being set
for b3, b4 being PartFunc of b1,b2 holds
( ( for b5 being set holds
b2 <> {b5} ) & TotFuncs b3 = TotFuncs b4 implies b3 = b4 )
proof end;

registration
let c1, c2 be non empty set ;
cluster -> non empty total Relation of a1,a2;
coherence
for b1 being Function of c1,c2 holds
not b1 is empty
by Def1, RELAT_1:60;
end;

definition
let c1, c2, c3 be set ;
func c1,c2 :-> c3 -> Function of [:{a1},{a2}:],{a3} means :: FUNCT_2:def 5
verum;
existence
ex b1 being Function of [:{c1},{c2}:],{c3} st
verum
;
uniqueness
for b1, b2 being Function of [:{c1},{c2}:],{c3} holds b1 = b2
by Th66;
end;

:: deftheorem Def5 defines :-> FUNCT_2:def 5 :
for b1, b2, b3 being set
for b4 being Function of [:{b1},{b2}:],{b3} holds
( b4 = b1,b2 :-> b3 iff verum );

scheme :: FUNCT_2:sch 6
s6{ F1() -> non empty set , F2() -> non empty set , F3() -> Element of F1(), F4() -> Element of F2(), F5( set ) -> Element of F2() } :
ex b1 being Function of F1(),F2() st
( b1 . F3() = F4() & ( for b2 being Element of F1() holds
( b2 <> F3() implies b1 . b2 = F5(b2) ) ) )
proof end;

scheme :: FUNCT_2:sch 7
s7{ F1() -> non empty set , F2() -> non empty set , F3() -> Element of F1(), F4() -> Element of F1(), F5() -> Element of F2(), F6() -> Element of F2(), F7( set ) -> Element of F2() } :
ex b1 being Function of F1(),F2() st
( b1 . F3() = F5() & b1 . F4() = F6() & ( for b2 being Element of F1() holds
( b2 <> F3() & b2 <> F4() implies b1 . b2 = F7(b2) ) ) )
provided
E40: F3() <> F4()
proof end;

theorem Th169: :: FUNCT_2:169
for b1, b2 being set
for b3 being Function holds
( b3 in Funcs b1,b2 implies ( dom b3 = b1 & rng b3 c= b2 ) )
proof end;

scheme :: FUNCT_2:sch 8
s8{ F1() -> non empty set , F2() -> set , F3( set ) -> set } :
ex b1 being Function of F1(),F2() st
for b2 being Element of F1() holds b1 . b2 = F3(b2)
provided
E40: for b1 being Element of F1() holds F3(b1) in F2()
proof end;

scheme :: FUNCT_2:sch 9
s9{ F1() -> non empty set , F2() -> non empty set , F3( set ) -> set } :
ex b1 being Function of F1(),F2() st
for b2 being Element of F1() holds b1 . b2 = F3(b2)
provided
E40: for b1 being Element of F1() holds
F3(b1) is Element of F2()
proof end;

scheme :: FUNCT_2:sch 10
s10{ F1() -> non empty set , F2() -> non empty set , P1[ set , set ] } :
ex b1 being Function of F1(),F2() st
for b2 being Element of F1() holds P1[b2,b1 . b2]
provided
E40: for b1 being Element of F1() holds
ex b2 being Element of F2() st P1[b1,b2]
proof end;

definition
let c1, c2, c3 be non empty set ;
let c4 be Function of c1,[:c2,c3:];
redefine func pr1 as pr1 c4 -> Function of a1,a2 means :: FUNCT_2:def 6
for b1 being Element of a1 holds a5 . b1 = (a4 . b1) `1 ;
coherence
pr1 c4 is Function of c1,c2
proof end;
compatibility
for b1 being Function of c1,c2 holds
( b1 = pr1 c4 iff for b2 being Element of c1 holds b1 . b2 = (c4 . b2) `1 )
proof end;
redefine func pr2 as pr2 c4 -> Function of a1,a3 means :: FUNCT_2:def 7
for b1 being Element of a1 holds a5 . b1 = (a4 . b1) `2 ;
coherence
pr2 c4 is Function of c1,c3
proof end;
compatibility
for b1 being Function of c1,c3 holds
( b1 = pr2 c4 iff for b2 being Element of c1 holds b1 . b2 = (c4 . b2) `2 )
proof end;
end;

:: deftheorem Def6 defines pr1 FUNCT_2:def 6 :
for b1, b2, b3 being non empty set
for b4 being Function of b1,[:b2,b3:]
for b5 being Function of b1,b2 holds
( b5 = pr1 b4 iff for b6 being Element of b1 holds b5 . b6 = (b4 . b6) `1 );

:: deftheorem Def7 defines pr2 FUNCT_2:def 7 :
for b1, b2, b3 being non empty set
for b4 being Function of b1,[:b2,b3:]
for b5 being Function of b1,b3 holds
( b5 = pr2 b4 iff for b6 being Element of b1 holds b5 . b6 = (b4 . b6) `2 );

definition
let c1 be set ;
let c2 be non empty set ;
let c3 be set ;
let c4 be non empty set ;
let c5 be Function of c1,c2;
let c6 be Function of c3,c4;
redefine pred = as c5 = c6 means :: FUNCT_2:def 8
( a1 = a3 & ( for b1 being Element of a1 holds a5 . b1 = a6 . b1 ) );
compatibility
( c5 = c6 iff ( c1 = c3 & ( for b1 being Element of c1 holds c5 . b1 = c6 . b1 ) ) )
proof end;
end;

:: deftheorem Def8 defines = FUNCT_2:def 8 :
for b1 being set
for b2 being non empty set
for b3 being set
for b4 being non empty set
for b5 being Function of b1,b2
for b6 being Function of b3,b4 holds
( b5 = b6 iff ( b1 = b3 & ( for b7 being Element of b1 holds b5 . b7 = b6 . b7 ) ) );

definition
let c1 be set ;
let c2 be non empty set ;
let c3, c4 be Function of c1,c2;
redefine pred = as c3 = c4 means :: FUNCT_2:def 9
for b1 being Element of a1 holds a3 . b1 = a4 . b1;
compatibility
( c3 = c4 iff for b1 being Element of c1 holds c3 . b1 = c4 . b1 )
by Th113;
end;

:: deftheorem Def9 defines = FUNCT_2:def 9 :
for b1 being set
for b2 being non empty set
for b3, b4 being Function of b1,b2 holds
( b3 = b4 iff for b5 being Element of b1 holds b3 . b5 = b4 . b5 );