definition
existence
ex b1 being BinOp of [.0,1.] st
for a, b being Element of [.0,1.] holds b1 . (a,b) = min (a,b)
uniqueness
for b1, b2 being BinOp of [.0,1.] st ( for a, b being Element of [.0,1.] holds b1 . (a,b) = min (a,b) ) & ( for a, b being Element of [.0,1.] holds b2 . (a,b) = min (a,b) ) holds
b1 = b2
end;
registration
cluster Relation-like [:[.0,1.],[.0,1.]:] -defined [.0,1.] -valued Function-like V30(
[:[.0,1.],[.0,1.]:],
[.0,1.])
commutative associative monotonic with-1-identity for
Element of
bool [:[:[.0,1.],[.0,1.]:],[.0,1.]:];
existence
ex b1 being BinOp of [.0,1.] st
( b1 is commutative & b1 is associative & b1 is monotonic & b1 is with-1-identity )
end;
definition
existence
ex b1 being BinOp of [.0,1.] st
for a, b being Element of [.0,1.] holds b1 . (a,b) = max (a,b)
uniqueness
for b1, b2 being BinOp of [.0,1.] st ( for a, b being Element of [.0,1.] holds b1 . (a,b) = max (a,b) ) & ( for a, b being Element of [.0,1.] holds b2 . (a,b) = max (a,b) ) holds
b1 = b2
end;
registration
cluster Relation-like [:[.0,1.],[.0,1.]:] -defined [.0,1.] -valued Function-like V30(
[:[.0,1.],[.0,1.]:],
[.0,1.])
commutative associative monotonic with-0-identity for
Element of
bool [:[:[.0,1.],[.0,1.]:],[.0,1.]:];
existence
ex b1 being BinOp of [.0,1.] st
( b1 is commutative & b1 is associative & b1 is monotonic & b1 is with-0-identity )
end;
registration
cluster Function-like V30(
[:[.0,1.],[.0,1.]:],
[.0,1.])
commutative monotonic with-1-identity -> commutative monotonic with-1-identity with-0-annihilating for
Element of
bool [:[:[.0,1.],[.0,1.]:],[.0,1.]:];
coherence
for b1 being commutative monotonic with-1-identity BinOp of [.0,1.] holds b1 is with-0-annihilating
by NormIs0;
cluster Function-like V30(
[:[.0,1.],[.0,1.]:],
[.0,1.])
commutative monotonic with-0-identity -> commutative monotonic with-1-annihilating with-0-identity for
Element of
bool [:[:[.0,1.],[.0,1.]:],[.0,1.]:];
coherence
for b1 being commutative monotonic with-0-identity BinOp of [.0,1.] holds b1 is with-1-annihilating
by ConormIs1;
end;
definition
existence
ex b1 being BinOp of [.0,1.] st
for a, b being Element of [.0,1.] holds b1 . (a,b) = a * b
uniqueness
for b1, b2 being BinOp of [.0,1.] st ( for a, b being Element of [.0,1.] holds b1 . (a,b) = a * b ) & ( for a, b being Element of [.0,1.] holds b2 . (a,b) = a * b ) holds
b1 = b2
end;
definition
existence
ex b1 being BinOp of [.0,1.] st
for a, b being Element of [.0,1.] holds b1 . (a,b) = (a + b) - (a * b)
uniqueness
for b1, b2 being BinOp of [.0,1.] st ( for a, b being Element of [.0,1.] holds b1 . (a,b) = (a + b) - (a * b) ) & ( for a, b being Element of [.0,1.] holds b2 . (a,b) = (a + b) - (a * b) ) holds
b1 = b2
end;
definition
existence
ex b1 being BinOp of [.0,1.] st
for a, b being Element of [.0,1.] holds b1 . (a,b) = max (0,((a + b) - 1))
uniqueness
for b1, b2 being BinOp of [.0,1.] st ( for a, b being Element of [.0,1.] holds b1 . (a,b) = max (0,((a + b) - 1)) ) & ( for a, b being Element of [.0,1.] holds b2 . (a,b) = max (0,((a + b) - 1)) ) holds
b1 = b2
end;
definition
existence
ex b1 being BinOp of [.0,1.] st
for a, b being Element of [.0,1.] holds
( ( max (a,b) = 1 implies b1 . (a,b) = min (a,b) ) & ( max (a,b) <> 1 implies b1 . (a,b) = 0 ) )
uniqueness
for b1, b2 being BinOp of [.0,1.] st ( for a, b being Element of [.0,1.] holds
( ( max (a,b) = 1 implies b1 . (a,b) = min (a,b) ) & ( max (a,b) <> 1 implies b1 . (a,b) = 0 ) ) ) & ( for a, b being Element of [.0,1.] holds
( ( max (a,b) = 1 implies b2 . (a,b) = min (a,b) ) & ( max (a,b) <> 1 implies b2 . (a,b) = 0 ) ) ) holds
b1 = b2
end;
definition
existence
ex b1 being BinOp of [.0,1.] st
for a, b being Element of [.0,1.] holds
( ( a + b > 1 implies b1 . (a,b) = min (a,b) ) & ( a + b <= 1 implies b1 . (a,b) = 0 ) )
uniqueness
for b1, b2 being BinOp of [.0,1.] st ( for a, b being Element of [.0,1.] holds
( ( a + b > 1 implies b1 . (a,b) = min (a,b) ) & ( a + b <= 1 implies b1 . (a,b) = 0 ) ) ) & ( for a, b being Element of [.0,1.] holds
( ( a + b > 1 implies b2 . (a,b) = min (a,b) ) & ( a + b <= 1 implies b2 . (a,b) = 0 ) ) ) holds
b1 = b2
end;
definition
existence
ex b1 being BinOp of [.0,1.] st
for a, b being Element of [.0,1.] holds b1 . (a,b) = (a * b) / ((a + b) - (a * b))
uniqueness
for b1, b2 being BinOp of [.0,1.] st ( for a, b being Element of [.0,1.] holds b1 . (a,b) = (a * b) / ((a + b) - (a * b)) ) & ( for a, b being Element of [.0,1.] holds b2 . (a,b) = (a * b) / ((a + b) - (a * b)) ) holds
b1 = b2
end;
definition
existence
ex b1 being BinOp of [.0,1.] st
for a, b being Element of [.0,1.] holds
( ( min (a,b) = 0 implies b1 . (a,b) = max (a,b) ) & ( min (a,b) <> 0 implies b1 . (a,b) = 1 ) )
uniqueness
for b1, b2 being BinOp of [.0,1.] st ( for a, b being Element of [.0,1.] holds
( ( min (a,b) = 0 implies b1 . (a,b) = max (a,b) ) & ( min (a,b) <> 0 implies b1 . (a,b) = 1 ) ) ) & ( for a, b being Element of [.0,1.] holds
( ( min (a,b) = 0 implies b2 . (a,b) = max (a,b) ) & ( min (a,b) <> 0 implies b2 . (a,b) = 1 ) ) ) holds
b1 = b2
end;
definition
let t be
BinOp of
[.0,1.];
existence
ex b1 being BinOp of [.0,1.] st
for a, b being Element of [.0,1.] holds b1 . (a,b) = 1 - (t . ((1 - a),(1 - b)))
uniqueness
for b1, b2 being BinOp of [.0,1.] st ( for a, b being Element of [.0,1.] holds b1 . (a,b) = 1 - (t . ((1 - a),(1 - b))) ) & ( for a, b being Element of [.0,1.] holds b2 . (a,b) = 1 - (t . ((1 - a),(1 - b))) ) holds
b1 = b2
end;
definition
existence
ex b1 being BinOp of [.0,1.] st
for a, b being Element of [.0,1.] holds
( ( a = b & b = 1 implies b1 . (a,b) = 1 ) & ( ( a <> 1 or b <> 1 ) implies b1 . (a,b) = ((a + b) - ((2 * a) * b)) / (1 - (a * b)) ) )
uniqueness
for b1, b2 being BinOp of [.0,1.] st ( for a, b being Element of [.0,1.] holds
( ( a = b & b = 1 implies b1 . (a,b) = 1 ) & ( ( a <> 1 or b <> 1 ) implies b1 . (a,b) = ((a + b) - ((2 * a) * b)) / (1 - (a * b)) ) ) ) & ( for a, b being Element of [.0,1.] holds
( ( a = b & b = 1 implies b2 . (a,b) = 1 ) & ( ( a <> 1 or b <> 1 ) implies b2 . (a,b) = ((a + b) - ((2 * a) * b)) / (1 - (a * b)) ) ) ) holds
b1 = b2
end;
definition
existence
ex b1 being BinOp of [.0,1.] st
for a, b being Element of [.0,1.] holds
( ( a + b < 1 implies b1 . (a,b) = max (a,b) ) & ( a + b >= 1 implies b1 . (a,b) = 1 ) )
uniqueness
for b1, b2 being BinOp of [.0,1.] st ( for a, b being Element of [.0,1.] holds
( ( a + b < 1 implies b1 . (a,b) = max (a,b) ) & ( a + b >= 1 implies b1 . (a,b) = 1 ) ) ) & ( for a, b being Element of [.0,1.] holds
( ( a + b < 1 implies b2 . (a,b) = max (a,b) ) & ( a + b >= 1 implies b2 . (a,b) = 1 ) ) ) holds
b1 = b2
end;
definition
existence
ex b1 being BinOp of [.0,1.] st
for a, b being Element of [.0,1.] holds b1 . (a,b) = min ((a + b),1)
uniqueness
for b1, b2 being BinOp of [.0,1.] st ( for a, b being Element of [.0,1.] holds b1 . (a,b) = min ((a + b),1) ) & ( for a, b being Element of [.0,1.] holds b2 . (a,b) = min ((a + b),1) ) holds
b1 = b2
end;
theorem ConormLukasiewicz: