:: RELAT_1 semantic presentation

definition
let IT be set ;
attr a1 is Relation-like means :Def1: :: RELAT_1:def 1
for x being set st x in IT holds
ex y, z being set st x = [y,z];
end;

:: deftheorem Def1 defines Relation-like RELAT_1:def 1 :
for IT being set holds
( IT is Relation-like iff for x being set st x in IT holds
ex y, z being set st x = [y,z] );

registration
cluster empty Relation-like set ;
existence
ex b1 being set st
( b1 is Relation-like & b1 is empty )
proof end;
end;

definition
mode Relation is Relation-like set ;
end;

theorem Th1: :: RELAT_1:1
canceled;

theorem Th2: :: RELAT_1:2
canceled;

theorem Th3: :: RELAT_1:3
for A being set
for R being Relation st A c= R holds
A is Relation-like
proof end;

theorem Th4: :: RELAT_1:4
for x, y being set holds {[x,y]} is Relation-like
proof end;

theorem Th5: :: RELAT_1:5
for a, b, c, d being set holds {[a,b],[c,d]} is Relation-like
proof end;

theorem Th6: :: RELAT_1:6
for X, Y being set holds [:X,Y:] is Relation-like
proof end;

scheme :: RELAT_1:sch 33
s33{ F1() -> set , F2() -> set , P1[ set , set ] } :
ex R being Relation st
for x, y being set holds
( [x,y] in R iff ( x in F1() & y in F2() & P1[x,y] ) )
proof end;

Lemma43: for x, y being set
for R being Relation st [x,y] in R holds
( x in union (union R) & y in union (union R) )
proof end;

definition
let P be Relation;
let R be Relation;
redefine pred c1 = c2 means :Def2: :: RELAT_1:def 2
for a, b being set holds
( [a,b] in P iff [a,b] in R );
compatibility
( P = R iff for a, b being set holds
( [a,b] in P iff [a,b] in R ) )
proof end;
end;

:: deftheorem Def2 defines = RELAT_1:def 2 :
for P, R being Relation holds
( P = R iff for a, b being set holds
( [a,b] in P iff [a,b] in R ) );

registration
let P be Relation;
let R be Relation;
cluster a1 /\ a2 -> Relation-like ;
coherence
P /\ R is Relation-like
proof end;
cluster a1 \/ a2 -> Relation-like ;
coherence
P \/ R is Relation-like
proof end;
cluster a1 \ a2 -> Relation-like ;
coherence
P \ R is Relation-like
proof end;
end;

definition
let P be Relation;
let R be Relation;
redefine pred c1 c= c2 means :Def3: :: RELAT_1:def 3
for a, b being set st [a,b] in P holds
[a,b] in R;
compatibility
( P c= R iff for a, b being set st [a,b] in P holds
[a,b] in R )
proof end;
end;

:: deftheorem Def3 defines c= RELAT_1:def 3 :
for P, R being Relation holds
( P c= R iff for a, b being set st [a,b] in P holds
[a,b] in R );

theorem Th7: :: RELAT_1:7
canceled;

theorem Th8: :: RELAT_1:8
canceled;

theorem Th9: :: RELAT_1:9
for X being set
for R being Relation holds X /\ R is Relation
proof end;

theorem Th10: :: RELAT_1:10
for X being set
for R being Relation holds R \ X is Relation by ;

definition
let R be Relation;
func dom c1 -> set means :Def4: :: RELAT_1:def 4
for x being set holds
( x in it iff ex y being set st [x,y] in R );
existence
ex b1 being set st
for x being set holds
( x in b1 iff ex y being set st [x,y] in R )
proof end;
uniqueness
for b1, b2 being set st ( for x being set holds
( x in b1 iff ex y being set st [x,y] in R ) ) & ( for x being set holds
( x in b2 iff ex y being set st [x,y] in R ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def4 defines dom RELAT_1:def 4 :
for R being Relation
for b2 being set holds
( b2 = dom R iff for x being set holds
( x in b2 iff ex y being set st [x,y] in R ) );

theorem Th11: :: RELAT_1:11
canceled;

theorem Th12: :: RELAT_1:12
canceled;

theorem Th13: :: RELAT_1:13
for P, R being Relation holds dom (P \/ R) = (dom P) \/ (dom R)
proof end;

theorem Th14: :: RELAT_1:14
for P, R being Relation holds dom (P /\ R) c= (dom P) /\ (dom R)
proof end;

theorem Th15: :: RELAT_1:15
for P, R being Relation holds (dom P) \ (dom R) c= dom (P \ R)
proof end;

definition
let R be Relation;
func rng c1 -> set means :Def5: :: RELAT_1:def 5
for y being set holds
( y in it iff ex x being set st [x,y] in R );
existence
ex b1 being set st
for y being set holds
( y in b1 iff ex x being set st [x,y] in R )
proof end;
uniqueness
for b1, b2 being set st ( for y being set holds
( y in b1 iff ex x being set st [x,y] in R ) ) & ( for y being set holds
( y in b2 iff ex x being set st [x,y] in R ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def5 defines rng RELAT_1:def 5 :
for R being Relation
for b2 being set holds
( b2 = rng R iff for y being set holds
( y in b2 iff ex x being set st [x,y] in R ) );

theorem Th16: :: RELAT_1:16
canceled;

theorem Th17: :: RELAT_1:17
canceled;

theorem Th18: :: RELAT_1:18
for x being set
for R being Relation st x in dom R holds
ex y being set st y in rng R
proof end;

theorem Th19: :: RELAT_1:19
for y being set
for R being Relation st y in rng R holds
ex x being set st x in dom R
proof end;

theorem Th20: :: RELAT_1:20
for x, y being set
for R being Relation st [x,y] in R holds
( x in dom R & y in rng R ) by , ;

theorem Th21: :: RELAT_1:21
for R being Relation holds R c= [:(dom R),(rng R):]
proof end;

theorem Th22: :: RELAT_1:22
for R being Relation holds R /\ [:(dom R),(rng R):] = R
proof end;

theorem Th23: :: RELAT_1:23
for x, y being set
for R being Relation st R = {[x,y]} holds
( dom R = {x} & rng R = {y} )
proof end;

theorem Th24: :: RELAT_1:24
for a, b, x, y being set
for R being Relation st R = {[a,b],[x,y]} holds
( dom R = {a,x} & rng R = {b,y} )
proof end;

theorem Th25: :: RELAT_1:25
for P, R being Relation st P c= R holds
( dom P c= dom R & rng P c= rng R )
proof end;

theorem Th26: :: RELAT_1:26
for P, R being Relation holds rng (P \/ R) = (rng P) \/ (rng R)
proof end;

theorem Th27: :: RELAT_1:27
for P, R being Relation holds rng (P /\ R) c= (rng P) /\ (rng R)
proof end;

theorem Th28: :: RELAT_1:28
for P, R being Relation holds (rng P) \ (rng R) c= rng (P \ R)
proof end;

definition
let R be Relation;
func field c1 -> set equals :: RELAT_1:def 6
(dom R) \/ (rng R);
correctness
coherence
(dom R) \/ (rng R) is set
;
;
end;

:: deftheorem Def6 defines field RELAT_1:def 6 :
for R being Relation holds field R = (dom R) \/ (rng R);

theorem Th29: :: RELAT_1:29
for R being Relation holds
( dom R c= field R & rng R c= field R ) by XBOOLE_1:7;

theorem Th30: :: RELAT_1:30
for a, b being set
for R being Relation st [a,b] in R holds
( a in field R & b in field R )
proof end;

theorem Th31: :: RELAT_1:31
for P, R being Relation st P c= R holds
field P c= field R
proof end;

theorem Th32: :: RELAT_1:32
for x, y being set
for R being Relation st R = {[x,y]} holds
field R = {x,y}
proof end;

theorem Th33: :: RELAT_1:33
for P, R being Relation holds field (P \/ R) = (field P) \/ (field R)
proof end;

theorem Th34: :: RELAT_1:34
for P, R being Relation holds field (P /\ R) c= (field P) /\ (field R)
proof end;

definition
let R be Relation;
func c1 ~ -> Relation means :Def7: :: RELAT_1:def 7
for x, y being set holds
( [x,y] in it iff [y,x] in R );
existence
ex b1 being Relation st
for x, y being set holds
( [x,y] in b1 iff [y,x] in R )
proof end;
uniqueness
for b1, b2 being Relation st ( for x, y being set holds
( [x,y] in b1 iff [y,x] in R ) ) & ( for x, y being set holds
( [x,y] in b2 iff [y,x] in R ) ) holds
b1 = b2
proof end;
involutiveness
for b1, b2 being Relation st ( for x, y being set holds
( [x,y] in b1 iff [y,x] in b2 ) ) holds
for x, y being set holds
( [x,y] in b2 iff [y,x] in b1 )
;
end;

:: deftheorem Def7 defines ~ RELAT_1:def 7 :
for R, b2 being Relation holds
( b2 = R ~ iff for x, y being set holds
( [x,y] in b2 iff [y,x] in R ) );

theorem Th35: :: RELAT_1:35
canceled;

theorem Th36: :: RELAT_1:36
canceled;

theorem Th37: :: RELAT_1:37
for R being Relation holds
( rng R = dom (R ~ ) & dom R = rng (R ~ ) )
proof end;

theorem Th38: :: RELAT_1:38
for R being Relation holds field R = field (R ~ )
proof end;

theorem Th39: :: RELAT_1:39
for P, R being Relation holds (P /\ R) ~ = (P ~ ) /\ (R ~ )
proof end;

theorem Th40: :: RELAT_1:40
for P, R being Relation holds (P \/ R) ~ = (P ~ ) \/ (R ~ )
proof end;

theorem Th41: :: RELAT_1:41
for P, R being Relation holds (P \ R) ~ = (P ~ ) \ (R ~ )
proof end;

definition
let P be Relation;
let R be Relation;
func c1 * c2 -> Relation means :Def8: :: RELAT_1:def 8
for x, y being set holds
( [x,y] in it iff ex z being set st
( [x,z] in P & [z,y] in R ) );
existence
ex b1 being Relation st
for x, y being set holds
( [x,y] in b1 iff ex z being set st
( [x,z] in P & [z,y] in R ) )
proof end;
uniqueness
for b1, b2 being Relation st ( for x, y being set holds
( [x,y] in b1 iff ex z being set st
( [x,z] in P & [z,y] in R ) ) ) & ( for x, y being set holds
( [x,y] in b2 iff ex z being set st
( [x,z] in P & [z,y] in R ) ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def8 defines * RELAT_1:def 8 :
for P, R, b3 being Relation holds
( b3 = P * R iff for x, y being set holds
( [x,y] in b3 iff ex z being set st
( [x,z] in P & [z,y] in R ) ) );

theorem Th42: :: RELAT_1:42
canceled;

theorem Th43: :: RELAT_1:43
canceled;

theorem Th44: :: RELAT_1:44
for P, R being Relation holds dom (P * R) c= dom P
proof end;

theorem Th45: :: RELAT_1:45
for P, R being Relation holds rng (P * R) c= rng R
proof end;

theorem Th46: :: RELAT_1:46
for R, P being Relation st rng R c= dom P holds
dom (R * P) = dom R
proof end;

theorem Th47: :: RELAT_1:47
for P, R being Relation st dom P c= rng R holds
rng (R * P) = rng P
proof end;

theorem Th48: :: RELAT_1:48
for P, R, Q being Relation st P c= R holds
Q * P c= Q * R
proof end;

theorem Th49: :: RELAT_1:49
for P, Q, R being Relation st P c= Q holds
P * R c= Q * R
proof end;

theorem Th50: :: RELAT_1:50
for P, R, Q, S being Relation st P c= R & Q c= S holds
P * Q c= R * S
proof end;

theorem Th51: :: RELAT_1:51
for P, R, Q being Relation holds P * (R \/ Q) = (P * R) \/ (P * Q)
proof end;

theorem Th52: :: RELAT_1:52
for P, R, Q being Relation holds P * (R /\ Q) c= (P * R) /\ (P * Q)
proof end;

theorem Th53: :: RELAT_1:53
for P, R, Q being Relation holds (P * R) \ (P * Q) c= P * (R \ Q)
proof end;

theorem Th54: :: RELAT_1:54
for P, R being Relation holds (P * R) ~ = (R ~ ) * (P ~ )
proof end;

theorem Th55: :: RELAT_1:55
for P, R, Q being Relation holds (P * R) * Q = P * (R * Q)
proof end;

registration
cluster empty -> Relation-like set ;
coherence
for b1 being set st b1 is empty holds
b1 is Relation-like
proof end;
end;

registration
cluster {} -> Relation-like ;
coherence
{} is Relation-like
;
end;

registration
cluster non empty set ;
existence
not for b1 being Relation holds b1 is empty
proof end;
end;

registration
let f be non empty Relation;
cluster dom a1 -> non empty ;
coherence
not dom f is empty
proof end;
cluster rng a1 -> non empty ;
coherence
not rng f is empty
proof end;
end;

theorem Th56: :: RELAT_1:56
for R being Relation st ( for x, y being set holds not [x,y] in R ) holds
R = {}
proof end;

theorem Th57: :: RELAT_1:57
canceled;

theorem Th58: :: RELAT_1:58
canceled;

theorem Th59: :: RELAT_1:59
canceled;

theorem Th60: :: RELAT_1:60
( dom {} = {} & rng {} = {} )
proof end;

theorem Th61: :: RELAT_1:61
canceled;

theorem Th62: :: RELAT_1:62
for R being Relation holds
( {} * R = {} & R * {} = {} )
proof end;

registration
let X be empty set ;
cluster dom a1 -> empty Relation-like ;
coherence
dom X is empty
by ;
cluster rng a1 -> empty Relation-like ;
coherence
rng X is empty
by ;
let R be Relation;
cluster a1 * a2 -> empty ;
coherence
X * R is empty
by ;
cluster a2 * a1 -> empty ;
coherence
R * X is empty
by ;
end;

theorem Th63: :: RELAT_1:63
field {} = {} ;

theorem Th64: :: RELAT_1:64
for R being Relation st ( dom R = {} or rng R = {} ) holds
R = {}
proof end;

theorem Th65: :: RELAT_1:65
for R being Relation holds
( dom R = {} iff rng R = {} ) by , ;

theorem Th66: :: RELAT_1:66
{} ~ = {}
proof end;

registration
let X be empty set ;
cluster a1 ~ -> empty ;
coherence
X ~ is empty
by ;
end;

theorem Th67: :: RELAT_1:67
for R, P being Relation st rng R misses dom P holds
R * P = {}
proof end;

definition
let R be Relation;
attr a1 is non-empty means :Def9: :: RELAT_1:def 9
not {} in rng R;
end;

:: deftheorem Def9 defines non-empty RELAT_1:def 9 :
for R being Relation holds
( R is non-empty iff not {} in rng R );

registration
cluster non-empty set ;
existence
ex b1 being Relation st b1 is non-empty
proof end;
end;

registration
let R be Relation;
let S be non-empty Relation;
cluster a1 * a2 -> non-empty ;
coherence
R * S is non-empty
proof end;
end;

definition
let X be set ;
func id c1 -> Relation means :Def10: :: RELAT_1:def 10
for x, y being set holds
( [x,y] in it iff ( x in X & x = y ) );
existence
ex b1 being Relation st
for x, y being set holds
( [x,y] in b1 iff ( x in X & x = y ) )
proof end;
uniqueness
for b1, b2 being Relation st ( for x, y being set holds
( [x,y] in b1 iff ( x in X & x = y ) ) ) & ( for x, y being set holds
( [x,y] in b2 iff ( x in X & x = y ) ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def10 defines id RELAT_1:def 10 :
for X being set
for b2 being Relation holds
( b2 = id X iff for x, y being set holds
( [x,y] in b2 iff ( x in X & x = y ) ) );

theorem Th68: :: RELAT_1:68
canceled;

theorem Th69: :: RELAT_1:69
canceled;

theorem Th70: :: RELAT_1:70
canceled;

theorem Th71: :: RELAT_1:71
for X being set holds
( dom (id X) = X & rng (id X) = X )
proof end;

theorem Th72: :: RELAT_1:72
for X being set holds (id X) ~ = id X
proof end;

theorem Th73: :: RELAT_1:73
for X being set
for R being Relation st ( for x being set st x in X holds
[x,x] in R ) holds
id X c= R
proof end;

theorem Th74: :: RELAT_1:74
for x, y, X being set
for R being Relation holds
( [x,y] in (id X) * R iff ( x in X & [x,y] in R ) )
proof end;

theorem Th75: :: RELAT_1:75
for x, y, Y being set
for R being Relation holds
( [x,y] in R * (id Y) iff ( y in Y & [x,y] in R ) )
proof end;

theorem Th76: :: RELAT_1:76
for X being set
for R being Relation holds
( R * (id X) c= R & (id X) * R c= R )
proof end;

theorem Th77: :: RELAT_1:77
for X being set
for R being Relation st dom R c= X holds
(id X) * R = R
proof end;

theorem Th78: :: RELAT_1:78
for R being Relation holds (id (dom R)) * R = R by ;

theorem Th79: :: RELAT_1:79
for Y being set
for R being Relation st rng R c= Y holds
R * (id Y) = R
proof end;

theorem Th80: :: RELAT_1:80
for R being Relation holds R * (id (rng R)) = R by ;

theorem Th81: :: RELAT_1:81
id {} = {}
proof end;

theorem Th82: :: RELAT_1:82
for X being set
for R, P2, P1 being Relation st dom R = X & rng P2 c= X & P2 * R = id (dom P1) & R * P1 = id X holds
P1 = P2
proof end;

definition
let R be Relation;
let X be set ;
func c1 | c2 -> Relation means :Def11: :: RELAT_1:def 11
for x, y being set holds
( [x,y] in it iff ( x in X & [x,y] in R ) );
existence
ex b1 being Relation st
for x, y being set holds
( [x,y] in b1 iff ( x in X & [x,y] in R ) )
proof end;
uniqueness
for b1, b2 being Relation st ( for x, y being set holds
( [x,y] in b1 iff ( x in X & [x,y] in R ) ) ) & ( for x, y being set holds
( [x,y] in b2 iff ( x in X & [x,y] in R ) ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def11 defines | RELAT_1:def 11 :
for R being Relation
for X being set
for b3 being Relation holds
( b3 = R | X iff for x, y being set holds
( [x,y] in b3 iff ( x in X & [x,y] in R ) ) );

theorem Th83: :: RELAT_1:83
canceled;

theorem Th84: :: RELAT_1:84
canceled;

theorem Th85: :: RELAT_1:85
canceled;

theorem Th86: :: RELAT_1:86
for x, X being set
for R being Relation holds
( x in dom (R | X) iff ( x in X & x in dom R ) )
proof end;

theorem Th87: :: RELAT_1:87
for X being set
for R being Relation holds dom (R | X) c= X
proof end;

theorem Th88: :: RELAT_1:88
for X being set
for R being Relation holds R | X c= R
proof end;

theorem Th89: :: RELAT_1:89
for X being set
for R being Relation holds dom (R | X) c= dom R
proof end;

theorem Th90: :: RELAT_1:90
for X being set
for R being Relation holds dom (R | X) = (dom R) /\ X
proof end;

theorem Th91: :: RELAT_1:91
for X being set
for R being Relation st X c= dom R holds
dom (R | X) = X
proof end;

theorem Th92: :: RELAT_1:92
for X being set
for R, P being Relation holds (R | X) * P c= R * P
proof end;

theorem Th93: :: RELAT_1:93
for X being set
for P, R being Relation holds P * (R | X) c= P * R
proof end;

theorem Th94: :: RELAT_1:94
for X being set
for R being Relation holds R | X = (id X) * R
proof end;

theorem Th95: :: RELAT_1:95
for X being set
for R being Relation holds
( R | X = {} iff dom R misses X )
proof end;

theorem Th96: :: RELAT_1:96
for X being set
for R being Relation holds R | X = R /\ [:X,(rng R):]
proof end;

theorem Th97: :: RELAT_1:97
for X being set
for R being Relation st dom R c= X holds
R | X = R
proof end;

theorem Th98: :: RELAT_1:98
for R being Relation holds R | (dom R) = R by Lemma43;

theorem Th99: :: RELAT_1:99
for X being set
for R being Relation holds rng (R | X) c= rng R
proof end;

theorem Th100: :: RELAT_1:100
for X, Y being set
for R being Relation holds (R | X) | Y = R | (X /\ Y)
proof end;

theorem Th101: :: RELAT_1:101
for X being set
for R being Relation holds (R | X) | X = R | X
proof end;

theorem Th102: :: RELAT_1:102
for X, Y being set
for R being Relation st X c= Y holds
(R | X) | Y = R | X
proof end;

theorem Th103: :: RELAT_1:103
for Y, X being set
for R being Relation st Y c= X holds
(R | X) | Y = R | Y
proof end;

theorem Th104: :: RELAT_1:104
for X, Y being set
for R being Relation st X c= Y holds
R | X c= R | Y
proof end;

theorem Th105: :: RELAT_1:105
for X being set
for P, R being Relation st P c= R holds
P | X c= R | X
proof end;

theorem Th106: :: RELAT_1:106
for X, Y being set
for P, R being Relation st P c= R & X c= Y holds
P | X c= R | Y
proof end;

theorem Th107: :: RELAT_1:107
for X, Y being set
for R being Relation holds R | (X \/ Y) = (R | X) \/ (R | Y)
proof end;

theorem Th108: :: RELAT_1:108
for X, Y being set
for R being Relation holds R | (X /\ Y) = (R | X) /\ (R | Y)
proof end;

theorem Th109: :: RELAT_1:109
for X, Y being set
for R being Relation holds R | (X \ Y) = (R | X) \ (R | Y)
proof end;

theorem Th110: :: RELAT_1:110
for R being Relation holds R | {} = {}
proof end;

theorem Th111: :: RELAT_1:111
for X being set holds {} | X = {}
proof end;

theorem Th112: :: RELAT_1:112
for X being set
for P, R being Relation holds (P * R) | X = (P | X) * R
proof end;

definition
let Y be set ;
let R be Relation;
func c1 | c2 -> Relation means :Def12: :: RELAT_1:def 12
for x, y being set holds
( [x,y] in it iff ( y in Y & [x,y] in R ) );
existence
ex b1 being Relation st
for x, y being set holds
( [x,y] in b1 iff ( y in Y & [x,y] in R ) )
proof end;
uniqueness
for b1, b2 being Relation st ( for x, y being set holds
( [x,y] in b1 iff ( y in Y & [x,y] in R ) ) ) & ( for x, y being set holds
( [x,y] in b2 iff ( y in Y & [x,y] in R ) ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def12 defines | RELAT_1:def 12 :
for Y being set
for R, b3 being Relation holds
( b3 = Y | R iff for x, y being set holds
( [x,y] in b3 iff ( y in Y & [x,y] in R ) ) );

theorem Th113: :: RELAT_1:113
canceled;

theorem Th114: :: RELAT_1:114
canceled;

theorem Th115: :: RELAT_1:115
for y, Y being set
for R being Relation holds
( y in rng (Y | R) iff ( y in Y & y in rng R ) )
proof end;

theorem Th116: :: RELAT_1:116
for Y being set
for R being Relation holds rng (Y | R) c= Y
proof end;

theorem Th117: :: RELAT_1:117
for Y being set
for R being Relation holds Y | R c= R
proof end;

theorem Th118: :: RELAT_1:118
for Y being set
for R being Relation holds rng (Y | R) c= rng R
proof end;

theorem Th119: :: RELAT_1:119
for Y being set
for R being Relation holds rng (Y | R) = (rng R) /\ Y
proof end;

theorem Th120: :: RELAT_1:120
for Y being set
for R being Relation st Y c= rng R holds
rng (Y | R) = Y
proof end;

theorem Th121: :: RELAT_1:121
for Y being set
for R, P being Relation holds (Y | R) * P c= R * P
proof end;

theorem Th122: :: RELAT_1:122
for Y being set
for P, R being Relation holds P * (Y | R) c= P * R
proof end;

theorem Th123: :: RELAT_1:123
for Y being set
for R being Relation holds Y | R = R * (id Y)
proof end;

theorem Th124: :: RELAT_1:124
for Y being set
for R being Relation holds Y | R = R /\ [:(dom R),Y:]
proof end;

theorem Th125: :: RELAT_1:125
for Y being set
for R being Relation st rng R c= Y holds
Y | R = R
proof end;

theorem Th126: :: RELAT_1:126
for R being Relation holds (rng R) | R = R by Th37;

theorem Th127: :: RELAT_1:127
for Y, X being set
for R being Relation holds Y | (X | R) = (Y /\ X) | R
proof end;

theorem Th128: :: RELAT_1:128
for Y being set
for R being Relation holds Y | (Y | R) = Y | R
proof end;

theorem Th129: :: RELAT_1:129
for X, Y being set
for R being Relation st X c= Y holds
Y | (X | R) = X | R
proof end;

theorem Th130: :: RELAT_1:130
for Y, X being set
for R being Relation st Y c= X holds
Y | (X | R) = Y | R
proof end;

theorem Th131: :: RELAT_1:131
for X, Y being set
for R being Relation st X c= Y holds
X | R c= Y | R
proof end;

theorem Th132: :: RELAT_1:132
for Y being set
for P1, P2 being Relation st P1 c= P2 holds
Y | P1 c= Y | P2
proof end;

theorem Th133: :: RELAT_1:133
for Y1, Y2 being set
for P1, P2 being Relation st P1 c= P2 & Y1 c= Y2 holds
Y1 | P1 c= Y2 | P2
proof end;

theorem Th134: :: RELAT_1:134
for X, Y being set
for R being Relation holds (X \/ Y) | R = (X | R) \/ (Y | R)
proof end;

theorem Th135: :: RELAT_1:135
for X, Y being set
for R being Relation holds (X /\ Y) | R = (X | R) /\ (Y | R)
proof end;

theorem Th136: :: RELAT_1:136
for X, Y being set
for R being Relation holds (X \ Y) | R = (X | R) \ (Y | R)
proof end;

theorem Th137: :: RELAT_1:137
for R being Relation holds {} | R = {}
proof end;

theorem Th138: :: RELAT_1:138
for Y being set holds Y | {} = {}
proof end;

theorem Th139: :: RELAT_1:139
for Y being set
for P, R being Relation holds Y | (P * R) = P * (Y | R)
proof end;

theorem Th140: :: RELAT_1:140
for Y, X being set
for R being Relation holds (Y | R) | X = Y | (R | X)
proof end;

definition
let R be Relation;
let X be set ;
func c1 .: c2 -> set means :Def13: :: RELAT_1:def 13
for y being set holds
( y in it iff ex x being set st
( [x,y] in R & x in X ) );
existence
ex b1 being set st
for y being set holds
( y in b1 iff ex x being set st
( [x,y] in R & x in X ) )
proof end;
uniqueness
for b1, b2 being set st ( for y being set holds
( y in b1 iff ex x being set st
( [x,y] in R & x in X ) ) ) & ( for y being set holds
( y in b2 iff ex x being set st
( [x,y] in R & x in X ) ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def13 defines .: RELAT_1:def 13 :
for R being Relation
for X, b3 being set holds
( b3 = R .: X iff for y being set holds
( y in b3 iff ex x being set st
( [x,y] in R & x in X ) ) );

theorem Th141: :: RELAT_1:141
canceled;

theorem Th142: :: RELAT_1:142
canceled;

theorem Th143: :: RELAT_1:143
for y, X being set
for R being Relation holds
( y in R .: X iff ex x being set st
( x in dom R & [x,y] in R & x in X ) )
proof end;

theorem Th144: :: RELAT_1:144
for X being set
for R being Relation holds R .: X c= rng R
proof end;

theorem Th145: :: RELAT_1:145
for X being set
for R being Relation holds R .: X = R .: ((dom R) /\ X)
proof end;

theorem Th146: :: RELAT_1:146
for R being Relation holds R .: (dom R) = rng R
proof end;

theorem Th147: :: RELAT_1:147
for X being set
for R being Relation holds R .: X c= R .: (dom R)
proof end;

theorem Th148: :: RELAT_1:148
for X being set
for R being Relation holds rng (R | X) = R .: X
proof end;

theorem Th149: :: RELAT_1:149
for R being Relation holds R .: {} = {}
proof end;

theorem Th150: :: RELAT_1:150
for X being set holds {} .: X = {}
proof end;

theorem Th151: :: RELAT_1:151
for X being set
for R being Relation holds
( R .: X = {} iff dom R misses X )
proof end;

theorem Th152: :: RELAT_1:152
for X being set
for R being Relation st X <> {} & X c= dom R holds
R .: X <> {}
proof end;

theorem Th153: :: RELAT_1:153
for X, Y being set
for R being Relation holds R .: (X \/ Y) = (R .: X) \/ (R .: Y)
proof end;

theorem Th154: :: RELAT_1:154
for X, Y being set
for R being Relation holds R .: (X /\ Y) c= (R .: X) /\ (R .: Y)
proof end;

theorem Th155: :: RELAT_1:155
for X, Y being set
for R being Relation holds (R .: X) \ (R .: Y) c= R .: (X \ Y)
proof end;

theorem Th156: :: RELAT_1:156
for X, Y being set
for R being Relation st X c= Y holds
R .: X c= R .: Y
proof end;

theorem Th157: :: RELAT_1:157
for X being set
for P, R being Relation st P c= R holds
P .: X c= R .: X
proof end;

theorem Th158: :: RELAT_1:158
for X, Y being set
for P, R being Relation st P c= R & X c= Y holds
P .: X c= R .: Y
proof end;

theorem Th159: :: RELAT_1:159
for X being set
for P, R being Relation holds (P * R) .: X = R .: (P .: X)
proof end;

theorem Th160: :: RELAT_1:160
for P, R being Relation holds rng (P * R) = R .: (rng P)
proof end;

theorem Th161: :: RELAT_1:161
for X, Y being set
for R being Relation holds (R | X) .: Y c= R .: Y
proof end;

theorem Th162: :: RELAT_1:162
for R being Relation
for X, Y being set st X c= Y holds
(R | Y) .: X = R .: X
proof end;

theorem Th163: :: RELAT_1:163
for X being set
for R being Relation holds (dom R) /\ X c= (R ~ ) .: (R .: X)
proof end;

definition
let R be Relation;
let Y be set ;
func c1 " c2 -> set means :Def14: :: RELAT_1:def 14
for x being set holds
( x in it iff ex y being set st
( [x,y] in R & y in Y ) );
existence
ex b1 being set st
for x being set holds
( x in b1 iff ex y being set st
( [x,y] in R & y in Y ) )
proof end;
uniqueness
for b1, b2 being set st ( for x being set holds
( x in b1 iff ex y being set st
( [x,y] in R & y in Y ) ) ) & ( for x being set holds
( x in b2 iff ex y being set st
( [x,y] in R & y in Y ) ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def14 defines " RELAT_1:def 14 :
for R being Relation
for Y, b3 being set holds
( b3 = R " Y iff for x being set holds
( x in b3 iff ex y being set st
( [x,y] in R & y in Y ) ) );

theorem Th164: :: RELAT_1:164
canceled;

theorem Th165: :: RELAT_1:165
canceled;

theorem Th166: :: RELAT_1:166
for x, Y being set
for R being Relation holds
( x in R " Y iff ex y being set st
( y in rng R & [x,y] in R & y in Y ) )
proof end;

theorem Th167: :: RELAT_1:167
for Y being set
for R being Relation holds R " Y c= dom R
proof end;

theorem Th168: :: RELAT_1:168
for Y being set
for R being Relation holds R " Y = R " ((rng R) /\ Y)
proof end;

theorem Th169: :: RELAT_1:169
for R being Relation holds R " (rng R) = dom R
proof end;

theorem Th170: :: RELAT_1:170
for Y being set
for R being Relation holds R " Y c= R " (rng R)
proof end;

theorem Th171: :: RELAT_1:171
for R being Relation holds R " {} = {}
proof end;

theorem Th172: :: RELAT_1:172
for Y being set holds {} " Y = {}
proof end;

theorem Th173: :: RELAT_1:173
for Y being set
for R being Relation holds
( R " Y = {} iff rng R misses Y )
proof end;

theorem Th174: :: RELAT_1:174
for Y being set
for R being Relation st Y <> {} & Y c= rng R holds
R " Y <> {}
proof end;

theorem Th175: :: RELAT_1:175
for X, Y being set
for R being Relation holds R " (X \/ Y) = (R " X) \/ (R " Y)
proof end;

theorem Th176: :: RELAT_1:176
for X, Y being set
for R being Relation holds R " (X /\ Y) c= (R " X) /\ (R " Y)
proof end;

theorem Th177: :: RELAT_1:177
for X, Y being set
for R being Relation holds (R " X) \ (R " Y) c= R " (X \ Y)
proof end;

theorem Th178: :: RELAT_1:178
for X, Y being set
for R being Relation st X c= Y holds
R " X c= R " Y
proof end;

theorem Th179: :: RELAT_1:179
for Y being set
for P, R being Relation st P c= R holds
P " Y c= R " Y
proof end;

theorem Th180: :: RELAT_1:180
for X, Y being set
for P, R being Relation st P c= R & X c= Y holds
P " X c= R " Y
proof end;

theorem Th181: :: RELAT_1:181
for Y being set
for P, R being Relation holds (P * R) " Y = P " (R " Y)
proof end;

theorem Th182: :: RELAT_1:182
for P, R being Relation holds dom (P * R) = P " (dom R)
proof end;

theorem Th183: :: RELAT_1:183
for Y being set
for R being Relation holds (rng R) /\ Y c= (R ~ ) " (R " Y)
proof end;

definition
let R be Relation;
attr a1 is empty-yielding means :Def15: :: RELAT_1:def 15
rng R c= {{} };
end;

:: deftheorem Def15 defines empty-yielding RELAT_1:def 15 :
for R being Relation holds
( R is empty-yielding iff rng R c= {{} } );

theorem Th184: :: RELAT_1:184
for R being Relation holds
( R is empty-yielding iff for X being set st X in rng R holds
X = {} )
proof end;

theorem Th185: :: RELAT_1:185
for f, g being Relation
for A, B being set st f | A = g | A & f | B = g | B holds
f | (A \/ B) = g | (A \/ B)
proof end;

theorem Th186: :: RELAT_1:186
for X being set
for f, g being Relation st dom g c= X & g c= f holds
g c= f | X
proof end;

theorem Th187: :: RELAT_1:187
for f being Relation
for X being set st X misses dom f holds
f | X = {}
proof end;

theorem Th188: :: RELAT_1:188
for f, g being Relation
for A, B being set st A c= B & f | B = g | B holds
f | A = g | A
proof end;

theorem Th189: :: RELAT_1:189
for R, S being Relation holds R | (dom S) = R | (dom (S | (dom R)))
proof end;

registration
cluster {} -> Relation-like empty-yielding ;
coherence
{} is empty-yielding
proof end;
end;

registration
cluster empty-yielding set ;
existence
ex b1 being Relation st b1 is empty-yielding
proof end;
end;

registration
let R be empty-yielding Relation;
let X be set ;
cluster a1 | a2 -> empty-yielding ;
coherence
R | X is empty-yielding
proof end;
end;

theorem Th190: :: RELAT_1:190
for X being set
for R being Relation st not R | X is empty-yielding holds
not R is empty-yielding
proof end;