begin
theorem
canceled;
theorem Th2:
definition
let f be
Function;
func f ~ -> Function means :
Def1:
(
dom it = dom f & ( for
x being
set st
x in dom f holds
( ( for
y,
z being
set st
f . x = [y,z] holds
it . x = [z,y] ) & (
f . x = it . x or ex
y,
z being
set st
f . x = [y,z] ) ) ) );
existence
ex b1 being Function st
( dom b1 = dom f & ( for x being set st x in dom f holds
( ( for y, z being set st f . x = [y,z] holds
b1 . x = [z,y] ) & ( f . x = b1 . x or ex y, z being set st f . x = [y,z] ) ) ) )
uniqueness
for b1, b2 being Function st dom b1 = dom f & ( for x being set st x in dom f holds
( ( for y, z being set st f . x = [y,z] holds
b1 . x = [z,y] ) & ( f . x = b1 . x or ex y, z being set st f . x = [y,z] ) ) ) & dom b2 = dom f & ( for x being set st x in dom f holds
( ( for y, z being set st f . x = [y,z] holds
b2 . x = [z,y] ) & ( f . x = b2 . x or ex y, z being set st f . x = [y,z] ) ) ) holds
b1 = b2
involutiveness
for b1, b2 being Function st dom b1 = dom b2 & ( for x being set st x in dom b2 holds
( ( for y, z being set st b2 . x = [y,z] holds
b1 . x = [z,y] ) & ( b2 . x = b1 . x or ex y, z being set st b2 . x = [y,z] ) ) ) holds
( dom b2 = dom b1 & ( for x being set st x in dom b1 holds
( ( for y, z being set st b1 . x = [y,z] holds
b2 . x = [z,y] ) & ( b1 . x = b2 . x or ex y, z being set st b1 . x = [y,z] ) ) ) )
end;
:: deftheorem Def1 defines ~ FUNCOP_1:def 1 :
for f, b2 being Function holds
( b2 = f ~ iff ( dom b2 = dom f & ( for x being set st x in dom f holds
( ( for y, z being set st f . x = [y,z] holds
b2 . x = [z,y] ) & ( f . x = b2 . x or ex y, z being set st f . x = [y,z] ) ) ) ) );
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th6:
theorem Th7:
theorem
canceled;
theorem
theorem Th10:
theorem Th11:
:: deftheorem defines --> FUNCOP_1:def 2 :
for A, z being set holds A --> z = [:A,{z}:];
theorem
canceled;
theorem Th13:
for
A,
x,
z being
set st
x in A holds
(A --> z) . x = z
theorem Th14:
theorem Th15:
theorem
theorem Th17:
theorem Th18:
theorem Th19:
theorem Th20:
for
A,
x,
B being
set st
x in B holds
(A --> x) " B = A
theorem
theorem
for
A,
x,
B being
set st not
x in B holds
(A --> x) " B = {}
theorem
theorem
theorem
theorem
:: deftheorem defines .: FUNCOP_1:def 3 :
for F, f, g being Function holds F .: (f,g) = F * <:f,g:>;
Lm1:
for f, g, F being Function
for x being set st x in dom (F * <:f,g:>) holds
(F * <:f,g:>) . x = F . ((f . x),(g . x))
theorem
theorem
theorem Th29:
theorem Th30:
theorem Th31:
:: deftheorem defines [:] FUNCOP_1:def 4 :
for F, f being Function
for x being set holds F [:] (f,x) = F * <:f,((dom f) --> x):>;
theorem
canceled;
theorem
canceled;
theorem
theorem Th35:
theorem
theorem Th37:
theorem
canceled;
theorem
:: deftheorem defines [;] FUNCOP_1:def 5 :
for F being Function
for x being set
for g being Function holds F [;] (x,g) = F * <:((dom g) --> x),g:>;
theorem
canceled;
theorem
theorem Th42:
theorem
theorem Th44:
theorem
canceled;
theorem
theorem Th47:
theorem Th48:
theorem Th49:
theorem
canceled;
theorem
theorem
theorem
theorem
theorem
theorem
theorem Th57:
theorem
theorem Th59:
theorem Th60:
theorem Th61:
theorem
canceled;
theorem
theorem
theorem Th65:
theorem Th66:
theorem Th67:
theorem
canceled;
theorem
theorem
theorem
theorem
canceled;
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
:: deftheorem Def6 defines Function-yielding FUNCOP_1:def 6 :
for IT being Function holds
( IT is Function-yielding iff for x being set st x in dom IT holds
IT . x is Function );
theorem
:: deftheorem defines .--> FUNCOP_1:def 7 :
for a, b, c being set holds (a,b) .--> c = {[a,b]} --> c;
theorem
for
a,
b,
c being
set holds
((a,b) .--> c) . (
a,
b)
= c
:: deftheorem Def8 defines IFEQ FUNCOP_1:def 8 :
for x, y, a, b being set holds
( ( x = y implies IFEQ (x,y,a,b) = a ) & ( not x = y implies IFEQ (x,y,a,b) = b ) );
:: deftheorem defines .--> FUNCOP_1:def 9 :
for x, y being set holds x .--> y = {x} --> y;
theorem Th87:
theorem
Lm2:
for o, m, r being set holds (o,m) :-> r is Function of [:{o},{m}:],{r}
definition
let o,
m,
r be
set ;
.-->redefine func (
o,
m)
:-> r -> Function of
[:{o},{m}:],
{r} means
verum;
coherence
(o,m) .--> r is Function of [:{o},{m}:],{r}
by Lm2;
compatibility
for b1 being Function of [:{o},{m}:],{r} holds
( b1 = (o,m) .--> r iff verum )
end;
:: deftheorem defines :-> FUNCOP_1:def 10 :
for o, m, r being set
for b4 being Function of [:{o},{m}:],{r} holds
( b4 = (o,m) :-> r iff verum );
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem Th10:
theorem
theorem
theorem