begin
theorem Th1:
theorem
canceled;
theorem
canceled;
theorem
theorem Th5:
theorem Th6:
theorem
theorem Th8:
theorem
theorem Th10:
theorem Th11:
theorem Th12:
theorem Th13:
theorem
canceled;
theorem Th15:
theorem
for
x,
y,
z,
t being
set holds
(
proj1 {[x,y],[z,t]} = {x,z} &
proj2 {[x,y],[z,t]} = {y,t} )
theorem Th17:
theorem
theorem
theorem Th20:
theorem
definition
canceled;canceled;let f be
Function;
func curry f -> Function means :
Def3:
(
dom it = proj1 (dom f) & ( for
x being
set st
x in proj1 (dom f) holds
ex
g being
Function st
(
it . x = g &
dom g = proj2 ((dom f) /\ [:{x},(proj2 (dom f)):]) & ( for
y being
set st
y in dom g holds
g . y = f . (
x,
y) ) ) ) );
existence
ex b1 being Function st
( dom b1 = proj1 (dom f) & ( for x being set st x in proj1 (dom f) holds
ex g being Function st
( b1 . x = g & dom g = proj2 ((dom f) /\ [:{x},(proj2 (dom f)):]) & ( for y being set st y in dom g holds
g . y = f . (x,y) ) ) ) )
uniqueness
for b1, b2 being Function st dom b1 = proj1 (dom f) & ( for x being set st x in proj1 (dom f) holds
ex g being Function st
( b1 . x = g & dom g = proj2 ((dom f) /\ [:{x},(proj2 (dom f)):]) & ( for y being set st y in dom g holds
g . y = f . (x,y) ) ) ) & dom b2 = proj1 (dom f) & ( for x being set st x in proj1 (dom f) holds
ex g being Function st
( b2 . x = g & dom g = proj2 ((dom f) /\ [:{x},(proj2 (dom f)):]) & ( for y being set st y in dom g holds
g . y = f . (x,y) ) ) ) holds
b1 = b2
func uncurry f -> Function means :
Def4:
( ( for
t being
set holds
(
t in dom it iff ex
x being
set ex
g being
Function ex
y being
set st
(
t = [x,y] &
x in dom f &
g = f . x &
y in dom g ) ) ) & ( for
x being
set for
g being
Function st
x in dom it &
g = f . (x `1) holds
it . x = g . (x `2) ) );
existence
ex b1 being Function st
( ( for t being set holds
( t in dom b1 iff ex x being set ex g being Function ex y being set st
( t = [x,y] & x in dom f & g = f . x & y in dom g ) ) ) & ( for x being set
for g being Function st x in dom b1 & g = f . (x `1) holds
b1 . x = g . (x `2) ) )
uniqueness
for b1, b2 being Function st ( for t being set holds
( t in dom b1 iff ex x being set ex g being Function ex y being set st
( t = [x,y] & x in dom f & g = f . x & y in dom g ) ) ) & ( for x being set
for g being Function st x in dom b1 & g = f . (x `1) holds
b1 . x = g . (x `2) ) & ( for t being set holds
( t in dom b2 iff ex x being set ex g being Function ex y being set st
( t = [x,y] & x in dom f & g = f . x & y in dom g ) ) ) & ( for x being set
for g being Function st x in dom b2 & g = f . (x `1) holds
b2 . x = g . (x `2) ) holds
b1 = b2
end;
:: deftheorem FUNCT_5:def 1 :
canceled;
:: deftheorem FUNCT_5:def 2 :
canceled;
:: deftheorem Def3 defines curry FUNCT_5:def 3 :
for f, b2 being Function holds
( b2 = curry f iff ( dom b2 = proj1 (dom f) & ( for x being set st x in proj1 (dom f) holds
ex g being Function st
( b2 . x = g & dom g = proj2 ((dom f) /\ [:{x},(proj2 (dom f)):]) & ( for y being set st y in dom g holds
g . y = f . (x,y) ) ) ) ) );
:: deftheorem Def4 defines uncurry FUNCT_5:def 4 :
for f, b2 being Function holds
( b2 = uncurry f iff ( ( for t being set holds
( t in dom b2 iff ex x being set ex g being Function ex y being set st
( t = [x,y] & x in dom f & g = f . x & y in dom g ) ) ) & ( for x being set
for g being Function st x in dom b2 & g = f . (x `1) holds
b2 . x = g . (x `2) ) ) );
:: deftheorem defines curry' FUNCT_5:def 5 :
for f being Function holds curry' f = curry (~ f);
:: deftheorem defines uncurry' FUNCT_5:def 6 :
for f being Function holds uncurry' f = ~ (uncurry f);
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th26:
theorem Th27:
theorem
theorem
theorem
theorem Th31:
theorem Th32:
theorem Th33:
theorem
theorem
theorem Th36:
theorem Th37:
theorem Th38:
theorem Th39:
theorem
theorem
theorem Th42:
theorem
theorem Th44:
theorem Th45:
theorem
theorem Th47:
theorem Th48:
theorem Th49:
theorem Th50:
theorem Th51:
theorem Th52:
theorem Th53:
theorem
theorem Th55:
theorem
theorem Th57:
theorem Th58:
theorem
theorem
theorem
theorem
theorem Th63:
theorem Th64:
theorem
theorem Th66:
theorem Th67:
for
X1,
Y1,
X2,
Y2 being
set st
X1,
Y1 are_equipotent &
X2,
Y2 are_equipotent holds
(
Funcs (
X1,
X2),
Funcs (
Y1,
Y2)
are_equipotent &
card (Funcs (X1,X2)) = card (Funcs (Y1,Y2)) )
theorem
theorem
for
X1,
X2,
X being
set st
X1 misses X2 holds
(
Funcs (
(X1 \/ X2),
X),
[:(Funcs (X1,X)),(Funcs (X2,X)):] are_equipotent &
card (Funcs ((X1 \/ X2),X)) = card [:(Funcs (X1,X)),(Funcs (X2,X)):] )
theorem
for
X,
Y,
Z being
set holds
(
Funcs (
[:X,Y:],
Z),
Funcs (
X,
(Funcs (Y,Z)))
are_equipotent &
card (Funcs ([:X,Y:],Z)) = card (Funcs (X,(Funcs (Y,Z)))) )
theorem
for
Z,
X,
Y being
set holds
(
Funcs (
Z,
[:X,Y:]),
[:(Funcs (Z,X)),(Funcs (Z,Y)):] are_equipotent &
card (Funcs (Z,[:X,Y:])) = card [:(Funcs (Z,X)),(Funcs (Z,Y)):] )
theorem
theorem
begin
:: deftheorem defines op1 FUNCT_5:def 7 :
op1 = 0 .--> 0;
:: deftheorem defines op2 FUNCT_5:def 8 :
op2 = (0,0) :-> 0;