:: FUNCT_4 semantic presentation

Lemma33: for x, y, Z being set st [x,y] in Z holds
( x in union (union Z) & y in union (union Z) )
proof end;

Lemma37: for x, x', y, y', x1, x1', y1, y1' being set st [[x,x'],[y,y']] = [[x1,x1'],[y1,y1']] holds
( x = x1 & y = y1 & x' = x1' & y' = y1' )
proof end;

theorem Th1: :: FUNCT_4:1
for Z being set st ( for z being set st z in Z holds
ex x, y being set st z = [x,y] ) holds
ex X, Y being set st Z c= [:X,Y:]
proof end;

theorem Th2: :: FUNCT_4:2
for g, f being Function holds g * f = (g | (rng f)) * f
proof end;

theorem Th3: :: FUNCT_4:3
canceled;

theorem Th4: :: FUNCT_4:4
for X, Y being set holds
( id X c= id Y iff X c= Y )
proof end;

theorem Th5: :: FUNCT_4:5
for X, Y, a being set st X c= Y holds
X --> a c= Y --> a
proof end;

theorem Th6: :: FUNCT_4:6
for X, a, Y, b being set st X --> a c= Y --> b holds
X c= Y
proof end;

theorem Th7: :: FUNCT_4:7
for X, a, Y, b being set st X <> {} & X --> a c= Y --> b holds
a = b
proof end;

theorem Th8: :: FUNCT_4:8
for x being set
for f being Function st x in dom f holds
{x} --> (f . x) c= f
proof end;

notation
let f be Function;
let g be Function;
synonym c1 <= c2 for c1 c= c2;
end;

theorem Th9: :: FUNCT_4:9
for Y, X being set
for f being Function holds (Y | f) | X <= f
proof end;

theorem Th10: :: FUNCT_4:10
for Y, X being set
for f, g being Function st f <= g holds
(Y | f) | X <= (Y | g) | X
proof end;

definition
let f be Function;
let g be Function;
func c1 +* c2 -> Function means :Def1: :: FUNCT_4:def 1
( dom it = (dom f) \/ (dom g) & ( for x being set st x in (dom f) \/ (dom g) holds
( ( x in dom g implies it . x = g . x ) & ( not x in dom g implies it . x = f . x ) ) ) );
existence
ex b1 being Function st
( dom b1 = (dom f) \/ (dom g) & ( for x being set st x in (dom f) \/ (dom g) holds
( ( x in dom g implies b1 . x = g . x ) & ( not x in dom g implies b1 . x = f . x ) ) ) )
proof end;
uniqueness
for b1, b2 being Function st dom b1 = (dom f) \/ (dom g) & ( for x being set st x in (dom f) \/ (dom g) holds
( ( x in dom g implies b1 . x = g . x ) & ( not x in dom g implies b1 . x = f . x ) ) ) & dom b2 = (dom f) \/ (dom g) & ( for x being set st x in (dom f) \/ (dom g) holds
( ( x in dom g implies b2 . x = g . x ) & ( not x in dom g implies b2 . x = f . x ) ) ) holds
b1 = b2
proof end;
idempotence
for f being Function holds
( dom f = (dom f) \/ (dom f) & ( for x being set st x in (dom f) \/ (dom f) holds
( ( x in dom f implies f . x = f . x ) & ( not x in dom f implies f . x = f . x ) ) ) )
;
end;

:: deftheorem Def1 defines +* FUNCT_4:def 1 :
for f, g, b3 being Function holds
( b3 = f +* g iff ( dom b3 = (dom f) \/ (dom g) & ( for x being set st x in (dom f) \/ (dom g) holds
( ( x in dom g implies b3 . x = g . x ) & ( not x in dom g implies b3 . x = f . x ) ) ) ) );

theorem Th11: :: FUNCT_4:11
for f, g being Function holds
( dom f c= dom (f +* g) & dom g c= dom (f +* g) )
proof end;

theorem Th12: :: FUNCT_4:12
for x being set
for g, f being Function st not x in dom g holds
(f +* g) . x = f . x
proof end;

theorem Th13: :: FUNCT_4:13
for x being set
for f, g being Function holds
( x in dom (f +* g) iff ( x in dom f or x in dom g ) )
proof end;

theorem Th14: :: FUNCT_4:14
for x being set
for g, f being Function st x in dom g holds
(f +* g) . x = g . x
proof end;

theorem Th15: :: FUNCT_4:15
for f, g, h being Function holds (f +* g) +* h = f +* (g +* h)
proof end;

theorem Th16: :: FUNCT_4:16
for x being set
for f, g being Function st f tolerates g & x in dom f holds
(f +* g) . x = f . x
proof end;

theorem Th17: :: FUNCT_4:17
for x being set
for f, g being Function st dom f misses dom g & x in dom f holds
(f +* g) . x = f . x
proof end;

theorem Th18: :: FUNCT_4:18
for f, g being Function holds rng (f +* g) c= (rng f) \/ (rng g)
proof end;

theorem Th19: :: FUNCT_4:19
for g, f being Function holds rng g c= rng (f +* g)
proof end;

theorem Th20: :: FUNCT_4:20
for f, g being Function st dom f c= dom g holds
f +* g = g
proof end;

theorem Th21: :: FUNCT_4:21
for f being Function holds {} +* f = f
proof end;

theorem Th22: :: FUNCT_4:22
for f being Function holds f +* {} = f
proof end;

theorem Th23: :: FUNCT_4:23
for X, Y being set holds (id X) +* (id Y) = id (X \/ Y)
proof end;

theorem Th24: :: FUNCT_4:24
for f, g being Function holds (f +* g) | (dom g) = g
proof end;

theorem Th25: :: FUNCT_4:25
for f, g being Function holds (f +* g) | ((dom f) \ (dom g)) c= f
proof end;

theorem Th26: :: FUNCT_4:26
for g, f being Function holds g c= f +* g
proof end;

theorem Th27: :: FUNCT_4:27
for f, g, h being Function st f tolerates g +* h holds
f | ((dom f) \ (dom h)) tolerates g
proof end;

theorem Th28: :: FUNCT_4:28
for f, g, h being Function st f tolerates g +* h holds
f tolerates h
proof end;

theorem Th29: :: FUNCT_4:29
for f, g being Function holds
( f tolerates g iff f c= f +* g )
proof end;

theorem Th30: :: FUNCT_4:30
for f, g being Function holds f +* g c= f \/ g
proof end;

theorem Th31: :: FUNCT_4:31
for f, g being Function holds
( f tolerates g iff f \/ g = f +* g )
proof end;

theorem Th32: :: FUNCT_4:32
for f, g being Function st dom f misses dom g holds
f \/ g = f +* g
proof end;

theorem Th33: :: FUNCT_4:33
for f, g being Function st dom f misses dom g holds
f c= f +* g
proof end;

theorem Th34: :: FUNCT_4:34
for f, g being Function st dom f misses dom g holds
(f +* g) | (dom f) = f
proof end;

theorem Th35: :: FUNCT_4:35
for f, g being Function holds
( f tolerates g iff f +* g = g +* f )
proof end;

theorem Th36: :: FUNCT_4:36
for f, g being Function st dom f misses dom g holds
f +* g = g +* f
proof end;

theorem Th37: :: FUNCT_4:37
for X, Y being set
for f, g being PartFunc of X,Y st g is total holds
f +* g = g
proof end;

theorem Th38: :: FUNCT_4:38
for X, Y being set
for f, g being Function of X,Y st ( Y = {} implies X = {} ) holds
f +* g = g
proof end;

theorem Th39: :: FUNCT_4:39
for X being set
for f, g being Function of X,X holds f +* g = g
proof end;

theorem Th40: :: FUNCT_4:40
for X being set
for D being non empty set
for f, g being Function of X,D holds f +* g = g by ;

theorem Th41: :: FUNCT_4:41
for X, Y being set
for f, g being PartFunc of X,Y holds f +* g is PartFunc of X,Y
proof end;

definition
let f be Function;
func ~ c1 -> Function means :Def2: :: FUNCT_4:def 2
( ( for x being set holds
( x in dom it iff ex y, z being set st
( x = [z,y] & [y,z] in dom f ) ) ) & ( for y, z being set st [y,z] in dom f holds
it . [z,y] = f . [y,z] ) );
existence
ex b1 being Function st
( ( for x being set holds
( x in dom b1 iff ex y, z being set st
( x = [z,y] & [y,z] in dom f ) ) ) & ( for y, z being set st [y,z] in dom f holds
b1 . [z,y] = f . [y,z] ) )
proof end;
uniqueness
for b1, b2 being Function st ( for x being set holds
( x in dom b1 iff ex y, z being set st
( x = [z,y] & [y,z] in dom f ) ) ) & ( for y, z being set st [y,z] in dom f holds
b1 . [z,y] = f . [y,z] ) & ( for x being set holds
( x in dom b2 iff ex y, z being set st
( x = [z,y] & [y,z] in dom f ) ) ) & ( for y, z being set st [y,z] in dom f holds
b2 . [z,y] = f . [y,z] ) holds
b1 = b2
proof end;
end;

:: deftheorem Def2 defines ~ FUNCT_4:def 2 :
for f, b2 being Function holds
( b2 = ~ f iff ( ( for x being set holds
( x in dom b2 iff ex y, z being set st
( x = [z,y] & [y,z] in dom f ) ) ) & ( for y, z being set st [y,z] in dom f holds
b2 . [z,y] = f . [y,z] ) ) );

theorem Th42: :: FUNCT_4:42
for f being Function holds rng (~ f) c= rng f
proof end;

theorem Th43: :: FUNCT_4:43
for x, y being set
for f being Function holds
( [x,y] in dom f iff [y,x] in dom (~ f) )
proof end;

theorem Th44: :: FUNCT_4:44
for y, x being set
for f being Function st [y,x] in dom (~ f) holds
(~ f) . [y,x] = f . [x,y]
proof end;

theorem Th45: :: FUNCT_4:45
for f being Function ex X, Y being set st dom (~ f) c= [:X,Y:]
proof end;

theorem Th46: :: FUNCT_4:46
for X, Y being set
for f being Function st dom f c= [:X,Y:] holds
dom (~ f) c= [:Y,X:]
proof end;

theorem Th47: :: FUNCT_4:47
for X, Y being set
for f being Function st dom f = [:X,Y:] holds
dom (~ f) = [:Y,X:]
proof end;

theorem Th48: :: FUNCT_4:48
for X, Y being set
for f being Function st dom f c= [:X,Y:] holds
rng (~ f) = rng f
proof end;

theorem Th49: :: FUNCT_4:49
for X, Y, Z being set
for f being PartFunc of [:X,Y:],Z holds ~ f is PartFunc of [:Y,X:],Z
proof end;

theorem Th50: :: FUNCT_4:50
for X, Y, Z being set
for f being Function of [:X,Y:],Z st Z <> {} holds
~ f is Function of [:Y,X:],Z
proof end;

theorem Th51: :: FUNCT_4:51
for X, Y being set
for D being non empty set
for f being Function of [:X,Y:],D holds ~ f is Function of [:Y,X:],D by ;

theorem Th52: :: FUNCT_4:52
for f being Function holds ~ (~ f) c= f
proof end;

theorem Th53: :: FUNCT_4:53
for X, Y being set
for f being Function st dom f c= [:X,Y:] holds
~ (~ f) = f
proof end;

theorem Th54: :: FUNCT_4:54
for X, Y, Z being set
for f being PartFunc of [:X,Y:],Z holds ~ (~ f) = f
proof end;

theorem Th55: :: FUNCT_4:55
canceled;

theorem Th56: :: FUNCT_4:56
canceled;

definition
let f be Function;
let g be Function;
func |:c1,c2:| -> Function means :Def3: :: FUNCT_4:def 3
( ( for z being set holds
( z in dom it iff ex x, y, x', y' being set st
( z = [[x,x'],[y,y']] & [x,y] in dom f & [x',y'] in dom g ) ) ) & ( for x, y, x', y' being set st [x,y] in dom f & [x',y'] in dom g holds
it . [[x,x'],[y,y']] = [(f . [x,y]),(g . [x',y'])] ) );
existence
ex b1 being Function st
( ( for z being set holds
( z in dom b1 iff ex x, y, x', y' being set st
( z = [[x,x'],[y,y']] & [x,y] in dom f & [x',y'] in dom g ) ) ) & ( for x, y, x', y' being set st [x,y] in dom f & [x',y'] in dom g holds
b1 . [[x,x'],[y,y']] = [(f . [x,y]),(g . [x',y'])] ) )
proof end;
uniqueness
for b1, b2 being Function st ( for z being set holds
( z in dom b1 iff ex x, y, x', y' being set st
( z = [[x,x'],[y,y']] & [x,y] in dom f & [x',y'] in dom g ) ) ) & ( for x, y, x', y' being set st [x,y] in dom f & [x',y'] in dom g holds
b1 . [[x,x'],[y,y']] = [(f . [x,y]),(g . [x',y'])] ) & ( for z being set holds
( z in dom b2 iff ex x, y, x', y' being set st
( z = [[x,x'],[y,y']] & [x,y] in dom f & [x',y'] in dom g ) ) ) & ( for x, y, x', y' being set st [x,y] in dom f & [x',y'] in dom g holds
b2 . [[x,x'],[y,y']] = [(f . [x,y]),(g . [x',y'])] ) holds
b1 = b2
proof end;
end;

:: deftheorem Def3 defines |: FUNCT_4:def 3 :
for f, g, b3 being Function holds
( b3 = |:f,g:| iff ( ( for z being set holds
( z in dom b3 iff ex x, y, x', y' being set st
( z = [[x,x'],[y,y']] & [x,y] in dom f & [x',y'] in dom g ) ) ) & ( for x, y, x', y' being set st [x,y] in dom f & [x',y'] in dom g holds
b3 . [[x,x'],[y,y']] = [(f . [x,y]),(g . [x',y'])] ) ) );

theorem Th57: :: FUNCT_4:57
for x, x', y, y' being set
for f, g being Function holds
( [[x,x'],[y,y']] in dom |:f,g:| iff ( [x,y] in dom f & [x',y'] in dom g ) )
proof end;

theorem Th58: :: FUNCT_4:58
for x, x', y, y' being set
for f, g being Function st [[x,x'],[y,y']] in dom |:f,g:| holds
|:f,g:| . [[x,x'],[y,y']] = [(f . [x,y]),(g . [x',y'])]
proof end;

theorem Th59: :: FUNCT_4:59
for f, g being Function holds rng |:f,g:| c= [:(rng f),(rng g):]
proof end;

theorem Th60: :: FUNCT_4:60
for X, Y, X', Y' being set
for f, g being Function st dom f c= [:X,Y:] & dom g c= [:X',Y':] holds
dom |:f,g:| c= [:[:X,X':],[:Y,Y':]:]
proof end;

theorem Th61: :: FUNCT_4:61
for X, Y, X', Y' being set
for f, g being Function st dom f = [:X,Y:] & dom g = [:X',Y':] holds
dom |:f,g:| = [:[:X,X':],[:Y,Y':]:]
proof end;

theorem Th62: :: FUNCT_4:62
for X, Y, Z, X', Y', Z' being set
for f being PartFunc of [:X,Y:],Z
for g being PartFunc of [:X',Y':],Z' holds |:f,g:| is PartFunc of [:[:X,X':],[:Y,Y':]:],[:Z,Z':]
proof end;

theorem Th63: :: FUNCT_4:63
for X, Y, Z, X', Y', Z' being set
for f being Function of [:X,Y:],Z
for g being Function of [:X',Y':],Z' st Z <> {} & Z' <> {} holds
|:f,g:| is Function of [:[:X,X':],[:Y,Y':]:],[:Z,Z':]
proof end;

theorem Th64: :: FUNCT_4:64
for X, Y, X', Y' being set
for D, D' being non empty set
for f being Function of [:X,Y:],D
for g being Function of [:X',Y':],D' holds |:f,g:| is Function of [:[:X,X':],[:Y,Y':]:],[:D,D':] by Th1;

definition
let x be set , y be set , a be set , b be set ;
func c1,c2 --> c3,c4 -> set equals :: FUNCT_4:def 4
({x} --> a) +* ({y} --> b);
correctness
coherence
({x} --> a) +* ({y} --> b) is set
;
;
end;

:: deftheorem Def4 defines --> FUNCT_4:def 4 :
for x, y, a, b being set holds x,y --> a,b = ({x} --> a) +* ({y} --> b);

registration
let x be set , y be set , a be set , b be set ;
cluster a1,a2 --> a3,a4 -> Relation-like Function-like ;
coherence
( x,y --> a,b is Function-like & x,y --> a,b is Relation-like )
;
end;

theorem Th65: :: FUNCT_4:65
for x1, x2, y1, y2 being set holds
( dom (x1,x2 --> y1,y2) = {x1,x2} & rng (x1,x2 --> y1,y2) c= {y1,y2} )
proof end;

theorem Th66: :: FUNCT_4:66
for x1, x2, y1, y2 being set st x1 <> x2 holds
( (x1,x2 --> y1,y2) . x1 = y1 & (x1,x2 --> y1,y2) . x2 = y2 )
proof end;

theorem Th67: :: FUNCT_4:67
for x1, x2, y1, y2 being set st x1 <> x2 holds
rng (x1,x2 --> y1,y2) = {y1,y2}
proof end;

theorem Th68: :: FUNCT_4:68
for x1, x2, y being set holds x1,x2 --> y,y = {x1,x2} --> y
proof end;

definition
let A be non empty set ;
let x1 be set ;
let x2 be set ;
let y1 be Element of A, y2 be Element of A;
redefine func --> as c2,c3 --> c4,c5 -> Function of {a2,a3},a1;
coherence
x1,x2 --> y1,y2 is Function of {x1,x2},A
proof end;
end;

theorem Th69: :: FUNCT_4:69
for a, b, c, d being set
for g being Function st dom g = {a,b} & g . a = c & g . b = d holds
g = a,b --> c,d
proof end;

theorem Th70: :: FUNCT_4:70
for x, y being set holds {x} --> y = {[x,y]}
proof end;

theorem Th71: :: FUNCT_4:71
for a, b, c, d being set st a <> c holds
a,c --> b,d = {[a,b],[c,d]}
proof end;

theorem Th72: :: FUNCT_4:72
for a, b, x, y, x', y' being set st a <> b & a,b --> x,y = a,b --> x',y' holds
( x = x' & y = y' )
proof end;

theorem Th73: :: FUNCT_4:73
for f1, f2, g1, g2 being Function st rng g1 c= dom f1 & rng g2 c= dom f2 & f1 tolerates f2 holds
(f1 +* f2) * (g1 +* g2) = (f1 * g1) +* (f2 * g2)
proof end;

theorem Th74: :: FUNCT_4:74
for f being Function
for A, B being set st dom f c= A \/ B holds
(f | A) +* (f | B) = f
proof end;

theorem Th75: :: FUNCT_4:75
for p, q being Function
for A being set holds (p +* q) | A = (p | A) +* (q | A)
proof end;

theorem Th76: :: FUNCT_4:76
for f, g being Function
for A being set st A misses dom g holds
(f +* g) | A = f | A
proof end;

theorem Th77: :: FUNCT_4:77
for f, g being Function
for A being set st dom f misses A holds
(f +* g) | A = g | A
proof end;

theorem Th78: :: FUNCT_4:78
for f, g, h being Function st dom g = dom h holds
(f +* g) +* h = f +* h
proof end;

theorem Th79: :: FUNCT_4:79
for f, g being Function st f c= g holds
( f +* g = g & g +* f = g )
proof end;

theorem Th80: :: FUNCT_4:80
for f being Function
for A being set holds f +* (f | A) = f
proof end;

theorem Th81: :: FUNCT_4:81
for f, g being Function
for B, C being set st dom f c= B & dom g c= C & B misses C holds
( (f +* g) | B = f & (f +* g) | C = g )
proof end;

theorem Th82: :: FUNCT_4:82
for p, q being Function
for A being set st dom p c= A & dom q misses A holds
(p +* q) | A = p
proof end;

theorem Th83: :: FUNCT_4:83
for f being Function
for A, B being set holds f | (A \/ B) = (f | A) +* (f | B)
proof end;