:: BINOP_1 semantic presentation
:: deftheorem Def1 defines . BINOP_1:def 1 :
theorem Th1: :: BINOP_1:1
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 Th2: :: BINOP_1:2
scheme :: BINOP_1:sch 37
s37{
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
E27:
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 :: BINOP_1:sch 44
s44{
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
E27:
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 a2 is
commutative means :
Def2:
:: BINOP_1:def 2
for
a,
b being
Element of
A holds
o . a,
b = o . b,
a;
attr a2 is
associative means :
Def3:
:: BINOP_1:def 3
for
a,
b,
c being
Element of
A holds
o . a,
(o . b,c) = o . (o . a,b),
c;
attr a2 is
idempotent means :
Def4:
:: BINOP_1:def 4
for
a being
Element of
A holds
o . a,
a = a;
end;
:: deftheorem Def2 defines commutative BINOP_1:def 2 :
:: deftheorem Def3 defines associative BINOP_1:def 3 :
:: deftheorem Def4 defines idempotent BINOP_1:def 4 :
:: deftheorem Def5 defines is_a_left_unity_wrt BINOP_1:def 5 :
:: deftheorem Def6 defines is_a_right_unity_wrt BINOP_1:def 6 :
:: deftheorem Def7 defines is_a_unity_wrt BINOP_1:def 7 :
theorem Th3: :: BINOP_1:3
canceled;
theorem Th4: :: BINOP_1:4
canceled;
theorem Th5: :: BINOP_1:5
canceled;
theorem Th6: :: BINOP_1:6
canceled;
theorem Th7: :: BINOP_1:7
canceled;
theorem Th8: :: BINOP_1:8
canceled;
theorem Th9: :: BINOP_1:9
canceled;
theorem Th10: :: BINOP_1:10
canceled;
theorem Th11: :: BINOP_1:11
theorem Th12: :: BINOP_1:12
theorem Th13: :: BINOP_1:13
theorem Th14: :: BINOP_1:14
theorem Th15: :: BINOP_1:15
theorem Th16: :: BINOP_1:16
theorem Th17: :: BINOP_1:17
theorem Th18: :: BINOP_1:18
:: deftheorem Def8 defines the_unity_wrt BINOP_1:def 8 :
definition
let A be
set ;
let o' be
BinOp of
A;
let o be
BinOp of
A;
pred c2 is_left_distributive_wrt c3 means :
Def9:
:: BINOP_1:def 9
for
a,
b,
c being
Element of
A holds
o' . a,
(o . b,c) = o . (o' . a,b),
(o' . a,c);
pred c2 is_right_distributive_wrt c3 means :
Def10:
:: BINOP_1:def 10
for
a,
b,
c being
Element of
A holds
o' . (o . a,b),
c = o . (o' . a,c),
(o' . b,c);
end;
:: deftheorem Def9 defines is_left_distributive_wrt BINOP_1:def 9 :
:: deftheorem Def10 defines is_right_distributive_wrt BINOP_1:def 10 :
:: deftheorem Def11 defines is_distributive_wrt BINOP_1:def 11 :
theorem Th19: :: BINOP_1:19
canceled;
theorem Th20: :: BINOP_1:20
canceled;
theorem Th21: :: BINOP_1:21
canceled;
theorem Th22: :: BINOP_1:22
canceled;
theorem Th23: :: BINOP_1:23
for
A being
set for
o',
o being
BinOp of
A holds
(
o' is_distributive_wrt o iff for
a,
b,
c being
Element of
A holds
(
o' . a,
(o . b,c) = o . (o' . a,b),
(o' . a,c) &
o' . (o . a,b),
c = o . (o' . a,c),
(o' . b,c) ) )
theorem Th24: :: BINOP_1:24
theorem Th25: :: BINOP_1:25
theorem Th26: :: BINOP_1:26
theorem Th27: :: BINOP_1:27
theorem Th28: :: BINOP_1:28
:: deftheorem Def12 defines is_distributive_wrt BINOP_1:def 12 :
definition
let A be non
empty set ;
let o be
BinOp of
A;
redefine attr a2 is
commutative means :: BINOP_1:def 13
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 ;
redefine attr a2 is
associative means :: BINOP_1:def 14
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 ;
redefine attr a2 is
idempotent means :: BINOP_1:def 15
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 ;
end;
:: deftheorem Def13 defines commutative BINOP_1:def 13 :
:: deftheorem Def14 defines associative BINOP_1:def 14 :
:: deftheorem Def15 defines idempotent BINOP_1:def 15 :
:: deftheorem Def16 defines is_a_left_unity_wrt BINOP_1:def 16 :
:: deftheorem Def17 defines is_a_right_unity_wrt BINOP_1:def 17 :
definition
let A be non
empty set ;
let o' be
BinOp of
A,
o be
BinOp of
A;
redefine pred c2 is_left_distributive_wrt c3 means :: BINOP_1:def 18
for
a,
b,
c being
Element of
A holds
o' . a,
(o . b,c) = o . (o' . a,b),
(o' . a,c);
correctness
compatibility
( o' is_left_distributive_wrt o iff for a, b, c being Element of A holds o' . a,(o . b,c) = o . (o' . a,b),(o' . a,c) );
by ;
redefine pred c2 is_right_distributive_wrt c3 means :: BINOP_1:def 19
for
a,
b,
c being
Element of
A holds
o' . (o . a,b),
c = o . (o' . a,c),
(o' . b,c);
correctness
compatibility
( o' is_right_distributive_wrt o iff for a, b, c being Element of A holds o' . (o . a,b),c = o . (o' . a,c),(o' . b,c) );
by ;
end;
:: deftheorem Def18 defines is_left_distributive_wrt BINOP_1:def 18 :
:: deftheorem Def19 defines is_right_distributive_wrt BINOP_1:def 19 :
:: deftheorem Def20 defines is_distributive_wrt BINOP_1:def 20 :
theorem Th29: :: BINOP_1:29
theorem Th30: :: BINOP_1:30
theorem Th31: :: BINOP_1:31