definition
let S1,
S2,
T1,
T2 be non
empty reflexive antisymmetric RelStr ;
let f be
Function of
S1,
S2;
assume A1:
f is
directed-sups-preserving
;
let g be
Function of
T1,
T2;
assume A2:
g is
directed-sups-preserving
;
existence
ex b1 being Function of (UPS (S2,T1)),(UPS (S1,T2)) st
for h being directed-sups-preserving Function of S2,T1 holds b1 . h = (g * h) * f
uniqueness
for b1, b2 being Function of (UPS (S2,T1)),(UPS (S1,T2)) st ( for h being directed-sups-preserving Function of S2,T1 holds b1 . h = (g * h) * f ) & ( for h being directed-sups-preserving Function of S2,T1 holds b2 . h = (g * h) * f ) holds
b1 = b2
end;
::
deftheorem Def5 defines
UPS WAYBEL27:def 5 :
for S1, S2, T1, T2 being non empty reflexive antisymmetric RelStr
for f being Function of S1,S2 st f is directed-sups-preserving holds
for g being Function of T1,T2 st g is directed-sups-preserving holds
for b7 being Function of (UPS (S2,T1)),(UPS (S1,T2)) holds
( b7 = UPS (f,g) iff for h being directed-sups-preserving Function of S2,T1 holds b7 . h = (g * h) * f );
theorem Th28:
for
S1,
S2,
S3,
T1,
T2,
T3 being non
empty Poset for
f1 being
directed-sups-preserving Function of
S2,
S3 for
f2 being
directed-sups-preserving Function of
S1,
S2 for
g1 being
directed-sups-preserving Function of
T1,
T2 for
g2 being
directed-sups-preserving Function of
T2,
T3 holds
(UPS (f2,g2)) * (UPS (f1,g1)) = UPS (
(f1 * f2),
(g2 * g1))
Lm1:
now for M being non empty set
for X, Y being non empty Poset
for f being directed-sups-preserving Function of X,(Y |^ M) holds
( f in Funcs ( the carrier of X,(Funcs (M, the carrier of Y))) & rng (commute f) c= Funcs ( the carrier of X, the carrier of Y) & dom (commute f) = M )
let M be non
empty set ;
for X, Y being non empty Poset
for f being directed-sups-preserving Function of X,(Y |^ M) holds
( f in Funcs ( the carrier of X,(Funcs (M, the carrier of Y))) & rng (commute f) c= Funcs ( the carrier of X, the carrier of Y) & dom (commute f) = M )let X,
Y be non
empty Poset;
for f being directed-sups-preserving Function of X,(Y |^ M) holds
( f in Funcs ( the carrier of X,(Funcs (M, the carrier of Y))) & rng (commute f) c= Funcs ( the carrier of X, the carrier of Y) & dom (commute f) = M )let f be
directed-sups-preserving Function of
X,
(Y |^ M);
( f in Funcs ( the carrier of X,(Funcs (M, the carrier of Y))) & rng (commute f) c= Funcs ( the carrier of X, the carrier of Y) & dom (commute f) = M )
the
carrier of
(Y |^ M) = Funcs (
M, the
carrier of
Y)
by YELLOW_1:28;
then A1:
rng f c= Funcs (
M, the
carrier of
Y)
;
dom f = the
carrier of
X
by FUNCT_2:def 1;
hence
f in Funcs ( the
carrier of
X,
(Funcs (M, the carrier of Y)))
by A1, FUNCT_2:def 2;
( rng (commute f) c= Funcs ( the carrier of X, the carrier of Y) & dom (commute f) = M )then
commute f in Funcs (
M,
(Funcs ( the carrier of X, the carrier of Y)))
by FUNCT_6:55;
then
ex
g being
Function st
(
commute f = g &
dom g = M &
rng g c= Funcs ( the
carrier of
X, the
carrier of
Y) )
by FUNCT_2:def 2;
hence
(
rng (commute f) c= Funcs ( the
carrier of
X, the
carrier of
Y) &
dom (commute f) = M )
;
verum
end;
:: preservation of composition