theorem Th1:
for
X,
Y,
Z being
set for
f1,
f2 being
Function of
[:X,Y:],
Z st ( for
x,
y being
set st
x in X &
y in Y holds
f1 . (
x,
y)
= f2 . (
x,
y) ) holds
f1 = f2
scheme
FuncEx2{
F1()
-> set ,
F2()
-> set ,
F3()
-> set ,
P1[
object ,
object ,
object ] } :
ex
f being
Function of
[:F1(),F2():],
F3() st
for
x,
y being
object st
x in F1() &
y in F2() holds
P1[
x,
y,
f . (
x,
y)]
provided
A1:
for
x,
y being
object st
x in F1() &
y in F2() holds
ex
z being
object st
(
z in F3() &
P1[
x,
y,
z] )
theorem Th11:
for
A being
set for
o,
o9 being
BinOp of
A holds
(
o9 is_distributive_wrt o iff for
a,
b,
c being
Element of
A holds
(
o9 . (
a,
(o . (b,c)))
= o . (
(o9 . (a,b)),
(o9 . (a,c))) &
o9 . (
(o . (a,b)),
c)
= o . (
(o9 . (a,c)),
(o9 . (b,c))) ) )
definition
let A be non
empty set ;
let o9,
o be
BinOp of
A;
correctness
compatibility
( o9 is_left_distributive_wrt o iff for a, b, c being Element of A holds o9 . (a,(o . (b,c))) = o . ((o9 . (a,b)),(o9 . (a,c))) );
;
correctness
compatibility
( o9 is_right_distributive_wrt o iff for a, b, c being Element of A holds o9 . ((o . (a,b)),c) = o . ((o9 . (a,c)),(o9 . (b,c))) );
;
end;
scheme
PartFuncEx2{
F1()
-> set ,
F2()
-> set ,
F3()
-> set ,
P1[
object ,
object ,
object ] } :
ex
f being
PartFunc of
[:F1(),F2():],
F3() st
( ( for
x,
y being
object holds
(
[x,y] in dom f iff (
x in F1() &
y in F2() & ex
z being
object st
P1[
x,
y,
z] ) ) ) & ( for
x,
y being
object st
[x,y] in dom f holds
P1[
x,
y,
f . (
x,
y)] ) )
provided
A1:
for
x,
y,
z being
object st
x in F1() &
y in F2() &
P1[
x,
y,
z] holds
z in F3()
and A2:
for
x,
y,
z1,
z2 being
object st
x in F1() &
y in F2() &
P1[
x,
y,
z1] &
P1[
x,
y,
z2] holds
z1 = z2
definition
let X,
Y,
Z be
set ;
let f1,
f2 be
Function of
[:X,Y:],
Z;
=redefine pred f1 = f2 means
for
x,
y being
set st
x in X &
y in Y holds
f1 . (
x,
y)
= f2 . (
x,
y);
compatibility
( f1 = f2 iff for x, y being set st x in X & y in Y holds
f1 . (x,y) = f2 . (x,y) )
by Th1;
end;
scheme
Sch9{
F1()
-> set ,
F2()
-> set ,
F3()
-> set ,
P1[
object ,
object ,
object ] } :
ex
f being
Function of
[:F1(),F2():],
F3() st
for
x,
y being
set st
x in F1() &
y in F2() holds
P1[
x,
y,
f . (
x,
y)]
provided
A1:
for
x,
y being
set st
x in F1() &
y in F2() holds
ex
z being
set st
(
z in F3() &
P1[
x,
y,
z] )