begin
:: deftheorem defines . BINOP_1:def 1 :
for f being Function
for a, b being set holds f . (a,b) = f . [a,b];
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
theorem
scheme
FuncEx2{
F1()
-> set ,
F2()
-> set ,
F3()
-> set ,
P1[
set ,
set ,
set ] } :
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] )
scheme
Lambda2{
F1()
-> set ,
F2()
-> set ,
F3()
-> set ,
F4(
set ,
set )
-> set } :
ex
f being
Function of
[:F1(),F2():],
F3() st
for
x,
y being
set st
x in F1() &
y in F2() holds
f . (
x,
y)
= F4(
x,
y)
provided
A1:
for
x,
y being
set st
x in F1() &
y in F2() holds
F4(
x,
y)
in F3()
definition
let A be
set ;
let o be
BinOp of
A;
attr o is
commutative means :
Def2:
for
a,
b being
Element of
A holds
o . (
a,
b)
= o . (
b,
a);
attr o is
associative means :
Def3:
for
a,
b,
c being
Element of
A holds
o . (
a,
(o . (b,c)))
= o . (
(o . (a,b)),
c);
attr o is
idempotent means :
Def4:
for
a being
Element of
A holds
o . (
a,
a)
= a;
end;
:: deftheorem Def2 defines commutative BINOP_1:def 2 :
for A being set
for o being BinOp of A holds
( o is commutative iff for a, b being Element of A holds o . (a,b) = o . (b,a) );
:: deftheorem Def3 defines associative BINOP_1:def 3 :
for A being set
for o being BinOp of A holds
( o is associative iff for a, b, c being Element of A holds o . (a,(o . (b,c))) = o . ((o . (a,b)),c) );
:: deftheorem Def4 defines idempotent BINOP_1:def 4 :
for A being set
for o being BinOp of A holds
( o is idempotent iff for a being Element of A holds o . (a,a) = a );
:: deftheorem Def5 defines is_a_left_unity_wrt BINOP_1:def 5 :
for A being set
for e being Element of A
for o being BinOp of A holds
( e is_a_left_unity_wrt o iff for a being Element of A holds o . (e,a) = a );
:: deftheorem Def6 defines is_a_right_unity_wrt BINOP_1:def 6 :
for A being set
for e being Element of A
for o being BinOp of A holds
( e is_a_right_unity_wrt o iff for a being Element of A holds o . (a,e) = a );
:: deftheorem defines is_a_unity_wrt BINOP_1:def 7 :
for A being set
for e being Element of A
for o being BinOp of A holds
( e is_a_unity_wrt o iff ( e is_a_left_unity_wrt o & e is_a_right_unity_wrt o ) );
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th11:
theorem Th12:
theorem Th13:
theorem Th14:
theorem Th15:
theorem
theorem Th17:
theorem Th18:
:: deftheorem defines the_unity_wrt BINOP_1:def 8 :
for A being set
for o being BinOp of A st ex e being Element of A st e is_a_unity_wrt o holds
for b3 being Element of A holds
( b3 = the_unity_wrt o iff b3 is_a_unity_wrt o );
definition
let A be
set ;
let o9,
o be
BinOp of
A;
pred o9 is_left_distributive_wrt o means :
Def9:
for
a,
b,
c being
Element of
A holds
o9 . (
a,
(o . (b,c)))
= o . (
(o9 . (a,b)),
(o9 . (a,c)));
pred o9 is_right_distributive_wrt o means :
Def10:
for
a,
b,
c being
Element of
A holds
o9 . (
(o . (a,b)),
c)
= o . (
(o9 . (a,c)),
(o9 . (b,c)));
end;
:: deftheorem Def9 defines is_left_distributive_wrt BINOP_1:def 9 :
for A being set
for o9, o being BinOp of A holds
( 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))) );
:: deftheorem Def10 defines is_right_distributive_wrt BINOP_1:def 10 :
for A being set
for o9, o being BinOp of A holds
( 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))) );
:: deftheorem defines is_distributive_wrt BINOP_1:def 11 :
for A being set
for o9, o being BinOp of A holds
( o9 is_distributive_wrt o iff ( o9 is_left_distributive_wrt o & o9 is_right_distributive_wrt o ) );
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th23:
for
A being
set for
o9,
o 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))) ) )
theorem Th24:
theorem Th25:
theorem Th26:
theorem Th27:
theorem
:: deftheorem Def12 defines is_distributive_wrt BINOP_1:def 12 :
for A being set
for u being UnOp of A
for o being BinOp of A holds
( u is_distributive_wrt o iff for a, b being Element of A holds u . (o . (a,b)) = o . ((u . a),(u . b)) );
definition
let A be non
empty set ;
let o be
BinOp of
A;
redefine attr o is
commutative means
for
a,
b being
Element of
A holds
o . (
a,
b)
= o . (
b,
a);
correctness
compatibility
( o is commutative iff for a, b being Element of A holds o . (a,b) = o . (b,a) );
by Def2;
redefine attr o is
associative means
for
a,
b,
c being
Element of
A holds
o . (
a,
(o . (b,c)))
= o . (
(o . (a,b)),
c);
correctness
compatibility
( o is associative iff for a, b, c being Element of A holds o . (a,(o . (b,c))) = o . ((o . (a,b)),c) );
by Def3;
redefine attr o is
idempotent means
for
a being
Element of
A holds
o . (
a,
a)
= a;
correctness
compatibility
( o is idempotent iff for a being Element of A holds o . (a,a) = a );
by Def4;
end;
:: deftheorem defines commutative BINOP_1:def 13 :
for A being non empty set
for o being BinOp of A holds
( o is commutative iff for a, b being Element of A holds o . (a,b) = o . (b,a) );
:: deftheorem defines associative BINOP_1:def 14 :
for A being non empty set
for o being BinOp of A holds
( o is associative iff for a, b, c being Element of A holds o . (a,(o . (b,c))) = o . ((o . (a,b)),c) );
:: deftheorem defines idempotent BINOP_1:def 15 :
for A being non empty set
for o being BinOp of A holds
( o is idempotent iff for a being Element of A holds o . (a,a) = a );
:: deftheorem defines is_a_left_unity_wrt BINOP_1:def 16 :
for A being non empty set
for e being Element of A
for o being BinOp of A holds
( e is_a_left_unity_wrt o iff for a being Element of A holds o . (e,a) = a );
:: deftheorem defines is_a_right_unity_wrt BINOP_1:def 17 :
for A being non empty set
for e being Element of A
for o being BinOp of A holds
( e is_a_right_unity_wrt o iff for a being Element of A holds o . (a,e) = a );
definition
let A be non
empty set ;
let o9,
o be
BinOp of
A;
redefine pred o9 is_left_distributive_wrt o means
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_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))) );
by Def9;
redefine pred o9 is_right_distributive_wrt o means
for
a,
b,
c being
Element of
A holds
o9 . (
(o . (a,b)),
c)
= o . (
(o9 . (a,c)),
(o9 . (b,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))) );
by Def10;
end;
:: deftheorem defines is_left_distributive_wrt BINOP_1:def 18 :
for A being non empty set
for o9, o being BinOp of A holds
( 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))) );
:: deftheorem defines is_right_distributive_wrt BINOP_1:def 19 :
for A being non empty set
for o9, o being BinOp of A holds
( 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))) );
:: deftheorem defines is_distributive_wrt BINOP_1:def 20 :
for A being non empty set
for u being UnOp of A
for o being BinOp of A holds
( u is_distributive_wrt o iff for a, b being Element of A holds u . (o . (a,b)) = o . ((u . a),(u . b)) );
theorem
theorem
theorem
theorem
scheme
PartFuncEx2{
F1()
-> set ,
F2()
-> set ,
F3()
-> set ,
P1[
set ,
set ,
set ] } :
ex
f being
PartFunc of
[:F1(),F2():],
F3() st
( ( for
x,
y being
set holds
(
[x,y] in dom f iff (
x in F1() &
y in F2() & ex
z being
set st
P1[
x,
y,
z] ) ) ) & ( for
x,
y being
set st
[x,y] in dom f holds
P1[
x,
y,
f . (
x,
y)] ) )
provided
A1:
for
x,
y,
z being
set st
x in F1() &
y in F2() &
P1[
x,
y,
z] holds
z in F3()
and A2:
for
x,
y,
z1,
z2 being
set st
x in F1() &
y in F2() &
P1[
x,
y,
z1] &
P1[
x,
y,
z2] holds
z1 = z2
scheme
LambdaR2{
F1()
-> set ,
F2()
-> set ,
F3()
-> set ,
F4(
set ,
set )
-> set ,
P1[
set ,
set ] } :
ex
f being
PartFunc of
[:F1(),F2():],
F3() st
( ( for
x,
y being
set holds
(
[x,y] in dom f iff (
x in F1() &
y in F2() &
P1[
x,
y] ) ) ) & ( for
x,
y being
set st
[x,y] in dom f holds
f . (
x,
y)
= F4(
x,
y) ) )
provided
A1:
for
x,
y being
set st
P1[
x,
y] holds
F4(
x,
y)
in F3()
scheme
PartLambda2{
F1()
-> set ,
F2()
-> set ,
F3()
-> set ,
F4(
set ,
set )
-> set ,
P1[
set ,
set ] } :
ex
f being
PartFunc of
[:F1(),F2():],
F3() st
( ( for
x,
y being
set holds
(
[x,y] in dom f iff (
x in F1() &
y in F2() &
P1[
x,
y] ) ) ) & ( for
x,
y being
set st
[x,y] in dom f holds
f . (
x,
y)
= F4(
x,
y) ) )
provided
A1:
for
x,
y being
set st
x in F1() &
y in F2() &
P1[
x,
y] holds
F4(
x,
y)
in F3()
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;
:: deftheorem defines = BINOP_1:def 21 :
for X, Y, Z being set
for f1, f2 being Function of [:X,Y:],Z holds
( f1 = f2 iff for x, y being set st x in X & y in Y holds
f1 . (x,y) = f2 . (x,y) );