Lm1:
for Y being set
for g being Function st g is Y -valued & g is FinSequence-like holds
g is FinSequence of Y
by FINSEQ_1:def 4;
Lm2:
for A, B, Y being set st Y c= Funcs (A,B) holds
union Y c= [:A,B:]
Lm3:
for X being set
for f being Function holds
( f is X -one-to-one iff for x, y being set st x in (dom f) /\ X & y in (dom f) /\ X & f . x = f . y holds
x = y )
Lm4:
for A, B being set
for D being non empty set
for f being BinOp of D st A c= B & f is B -one-to-one holds
f is A -one-to-one
Lm5:
for A, B being set
for m, n being Nat st m -tuples_on A meets n -tuples_on B holds
m = n
Lm6:
for D being set holds 0 -tuples_on D = {{}}
Lm7:
for i being Nat
for Y being set holds i -tuples_on Y = Funcs ((Seg i),Y)
Lm8:
for A, B being set
for m being Nat holds (m -tuples_on A) /\ (B *) c= m -tuples_on B
Lm9:
for A, B, C being set holds (Funcs (A,B)) /\ (Funcs (A,C)) = Funcs (A,(B /\ C))
Lm10:
for D being non empty set
for x, y being FinSequence of D holds (D -concatenation) . (x,y) = x ^ y
Lm11:
for A, B being set
for D being non empty set
for f being BinOp of D st B is f -unambiguous & A c= B holds
A is f -unambiguous
Lm12:
for D being non empty set
for f being BinOp of D holds {} is f -unambiguous
Lm13:
for f, g being FinSequence holds dom <:f,g:> = Seg (min ((len f),(len g)))
Lm14:
for D being non empty set
for f being BinOp of D
for x being FinSequence of D
for m being Nat holds
( ( m + 1 <= len x implies (MultPlace (f,x)) . m in D ) & ( for n being Nat st n >= m + 1 holds
(MultPlace (f,(x | n))) . m = (MultPlace (f,x)) . m ) )
Lm15:
for D being non empty set
for f being BinOp of D
for x being non empty FinSequence of D
for y being Element of D holds
( (MultPlace f) . <*y*> = y & (MultPlace f) . (x ^ <*y*>) = f . (((MultPlace f) . x),y) )
Lm16:
for D being non empty set
for f being BinOp of D
for X being set holds
( f is [:X,D:] -one-to-one iff for x, y, d1, d2 being set st x in X /\ D & y in X /\ D & d1 in D & d2 in D & f . (x,d1) = f . (y,d2) holds
( x = y & d1 = d2 ) )
definition
let D be non
empty set ;
let f be
BinOp of
D;
let X be
set ;
redefine attr X is
f -unambiguous means :
Def10:
for
x,
y,
d1,
d2 being
set st
x in X /\ D &
y in X /\ D &
d1 in D &
d2 in D &
f . (
x,
d1)
= f . (
y,
d2) holds
(
x = y &
d1 = d2 );
compatibility
( X is f -unambiguous iff for x, y, d1, d2 being set st x in X /\ D & y in X /\ D & d1 in D & d2 in D & f . (x,d1) = f . (y,d2) holds
( x = y & d1 = d2 ) )
by Lm16;
end;
Lm17:
for D being non empty set
for f being BinOp of D st f is associative holds
for d being Element of D
for m being Nat
for x being Element of (m + 1) -tuples_on D holds (MultPlace f) . (<*d*> ^ x) = f . (d,((MultPlace f) . x))
Lm18:
for D being non empty set
for X being non empty Subset of D
for f being BinOp of D
for m being Nat st f is associative & X is f -unambiguous holds
(MultPlace f) .: ((m + 1) -tuples_on X) is f -unambiguous
Lm19:
for Y being set
for D being non empty set
for f being BinOp of D
for m being Nat st f is associative & Y is f -unambiguous holds
(MultPlace f) .: ((m + 1) -tuples_on Y) is f -unambiguous
Lm20:
for D being non empty set
for X being non empty Subset of D
for f being BinOp of D
for m being Nat st f is associative & X is f -unambiguous holds
MultPlace f is (m + 1) -tuples_on X -one-to-one
Lm21:
for Y being set
for D being non empty set
for f being BinOp of D
for m being Nat st f is associative & Y is f -unambiguous holds
MultPlace f is (m + 1) -tuples_on Y -one-to-one
Lm22:
for D being non empty set
for w being non empty FinSequence of D holds (D -firstChar) . w = w . 1
Lm23:
for D being non empty set holds (D -multiCat) | (((D *) *) \ {{}}) = MultPlace (D -concatenation)
Lm24:
for Y being set
for D being non empty set holds (D -multiCat) | (Y \ {{}}) = (MultPlace (D -concatenation)) | Y
Lm25:
for D being non empty set holds {} .--> {} tolerates MultPlace (D -concatenation)
Lm26:
for g being Function
for m being Nat
for f being Function st f c= g holds
iter (f,m) c= iter (g,m)
Lm27:
for R being Relation holds R [*] is_transitive_in field R
Lm28:
for R being Relation holds field (R [*]) c= field R
Lm29:
for R being Relation holds R [*] is_reflexive_in field R
Lm30:
for R being Relation holds field (R [*]) = field R
by LANG1:18, RELAT_1:16, Lm28;
Lm31:
for f being Function holds iter (f,0) c= f [*]
Lm32:
for m being Nat
for f being Function holds iter (f,(m + 1)) c= f [*]
Lm33:
for m being Nat
for f being Function holds iter (f,m) c= f [*]
Lm34:
for x being set
for g being Function
for m being Nat
for f being Function st rng f c= dom f & x in dom f & g . 0 = x & ( for i being Nat st i < m holds
g . (i + 1) = f . (g . i) ) holds
g . m = (iter (f,m)) . x
Lm35:
for x being set
for p being FinSequence
for m being Nat
for f being Function st rng f c= dom f & x in dom f & p . 1 = x & ( for i being Nat st i >= 1 & i < m + 1 holds
p . (i + 1) = f . (p . i) ) holds
p . (m + 1) = (iter (f,m)) . x
Lm36:
for z being set
for f being Function st z in f [*] & rng f c= dom f holds
ex n being Nat st z in iter (f,n)
Lm37:
for f being Function st rng f c= dom f holds
f [*] = union { (iter (f,mm)) where mm is Element of NAT : verum }
Lm38:
for A, B being set holds chi (A,B) = (B --> 0) +* ((A /\ B) --> 1)
Lm39:
for A, B being set holds chi (A,B) = ((B \ A) --> 0) +* ((A /\ B) --> 1)
Lm40:
for A, B being set holds chi (A,B) = ((B \ A) --> 0) \/ ((A /\ B) --> 1)
Lm41:
for Y being set
for R being b1 -defined total Relation ex f being Function of Y,(rng R) st
( f c= R & dom f = Y )
Lm42:
for R being Relation
for f being Function st dom f c= dom R & R c= f holds
R = f
Lm43:
for Y being set
for P, R being Relation
for Q being b1 -defined Relation st Q is total & R is Y -defined & ((P * Q) * (Q ~)) * R is Function-like & rng P c= dom R holds
((P * Q) * (Q ~)) * R = P * R
theorem Th29:
for
Y,
X being
set holds
( (
X \+\ Y = {} implies
X = Y ) & (
X = Y implies
X \+\ Y = {} ) & (
X \ Y = {} implies
X c= Y ) & (
X c= Y implies
X \ Y = {} ) & ( for
x being
object holds
(
{x} \ Y = {} iff
x in Y ) ) )
Lm44:
for x2 being set
for U being non empty set
for u, u1 being Element of U holds
( ( u = u1 implies (u1,x2) -SymbolSubstIn <*u*> = <*x2*> ) & ( u <> u1 implies (u1,x2) -SymbolSubstIn <*u*> = <*u*> ) )
Lm45:
for x being set
for U being non empty set
for u being Element of U
for q1, q2 being b2 -valued FinSequence holds (x,u) -SymbolSubstIn (q1 ^ q2) = ((x,u) -SymbolSubstIn q1) ^ ((x,u) -SymbolSubstIn q2)
Lm46:
for U being non empty set
for u, u1, u2 being Element of U holds
( ( u = u1 implies (u1 SubstWith u2) . <*u*> = <*u2*> ) & ( u <> u1 implies (u1 SubstWith u2) . <*u*> = <*u*> ) )
Lm47:
for x being set
for U being non empty set
for u being Element of U
for q1, q2 being b2 -valued FinSequence holds (x SubstWith u) . (q1 ^ q2) = ((x SubstWith u) . q1) ^ ((x SubstWith u) . q2)
Lm48:
for x being set
for U being non empty set
for u being Element of U
for p being FinSequence st p is U * -valued holds
(x SubstWith u) . ((U -multiCat) . p) = (U -multiCat) . ((x SubstWith u) * p)
Lm49:
for m being Nat
for p1, p2, q1, q2 being FinSequence st p1 is m -element & q1 is m -element & p1 ^ p2 = q1 ^ q2 holds
( p1 = q1 & p2 = q2 )
Lm50:
for X being functional set st X is finite & X is finite-membered holds
SymbolsOf X is finite
Lm51:
for B being functional set
for A being Subset of B holds { (rng x) where x is Element of A : x in A } c= { (rng x) where x is Element of B : x in B }
theorem Th48:
for
X,
Y being
set holds
X = (X \ Y) \/ (X /\ Y)
Lm52:
for Y being set
for U being non empty set
for p being FinSequence st p <> {} & p is Y -valued & Y c= U * & Y is with_non-empty_elements holds
(U -multiCat) . p <> {}
Lm53:
for A, B, X being set st X is Subset of (Funcs (A,B)) holds
X is Subset-Family of [:A,B:]
Lm54:
for Y being set
for U being non empty set
for x, y being Element of U holds
not ( ( x in Y implies y in Y ) & ( y in Y implies x in Y ) & not [((chi (Y,U)) . x),((chi (Y,U)) . y)] in id BOOLEAN )
Lm55:
for A, B being set
for R being Relation holds R | (A \ B) = (R | A) \ [:B,(rng R):]
Lm56:
for f, g being Function holds f +* g = (f \ [:(dom g),(rng f):]) \/ g
Lm57:
for A, B, C being set holds A /\ C = A \ ((A \/ B) \ C)
Lm58:
for X being set
for P being Relation holds P | ((dom P) \ X) = P \ [:X,(rng P):]
scheme
FraenkelInclusion{
F1()
-> set ,
F2(
set ,
set )
-> set ,
P1[
set ,
set ] } :
F1()
c= { F2(x,F1()) where x is Element of F1() : P1[x,F1()] }
provided
A1:
for
y being
set st
y in F1() holds
ex
x being
set st
(
x in F1() &
P1[
x,
F1()] &
y = F2(
x,
F1()) )
scheme
FraenkelInclusion2{
F1()
-> set ,
P1[
set ,
set ] } :
F1()
c= { x where x is Element of F1() : P1[x,F1()] }
provided
A1:
for
x being
set st
x in F1() holds
P1[
x,
F1()]
scheme
FraenkelEq{
F1()
-> non
empty set ,
P1[
set ] } :
F1()
= { x where x is Element of F1() : P1[x] }
provided
A1:
for
x being
set st
x in F1() holds
P1[
x]
registration
let X,
Y be
set ;
coherence
(proj1 [:X,Y:]) \ X is empty
by FUNCT_5:10, Th29;
coherence
((proj1 [:X,Y:]) /\ X) \+\ (proj1 [:X,Y:]) is empty
coherence
(proj2 [:X,Y:]) \ Y is empty
by FUNCT_5:10, Th29;
coherence
((proj2 [:X,Y:]) /\ Y) \+\ (proj2 [:X,Y:]) is empty
coherence
((proj2 X) \/ (proj2 Y)) \+\ (proj2 (X \/ Y)) is empty
by XTUPLE_0:27, Th29;
coherence
((proj1 X) \/ (proj1 Y)) \+\ (proj1 (X \/ Y)) is empty
by XTUPLE_0:23, Th29;
let y be
Subset of
Y;
coherence
(X \ Y) /\ y is empty
end;
Lm59:
for f, g being Function
for x being Element of [:f,g:] st f <> {} & g <> {} holds
( x `1 is pair & x `2 is pair )
Lm60:
for B, C being set
for P being Relation holds P * [:B,C:] c= [:(P " B),C:]
Lm61:
for B, C being set
for P being Relation holds P * [:B,C:] = [:(P " B),C:]
Lm62:
for X, Y, Z, y being set st y <> {} & y c= Y holds
[:X,y:] * [:Y,Z:] = [:X,Z:]
Lm63:
for X being set
for P being Relation holds P * [:(rng P),X:] = [:(dom P),X:]
Lm64:
for f being Function st rng f = dom f holds
f is onto Function of (dom f),(dom f)
Lm65:
for f being Function st f is involutive holds
f * f c= id (dom f)
Lm66:
for X being set holds X /\ (id ((proj1 X) /\ (proj2 X))) = X /\ (id (proj1 X))
Lm67:
for X, Y being set holds (id X) | Y = id (X /\ Y)
Lm68:
for b1, b2 being set
for f1, f2 being Function st (rng f1) /\ (rng f2) = {} & b1 <> b2 & rng f1 c= dom f2 & rng f2 c= dom f1 holds
[:(rng f1),{b2}:] \/ [:(rng f2),{b1}:] c= ((f1 \/ f2) >*> ([:(rng f1),{b2}:] \/ [:(rng f2),{b1}:])) >*> ((b1,b2) --> (b2,b1))
Lm69:
for a, a1 being set
for f being Function st a in {a1,(f . a1)} & f is involutive & a1 in dom f holds
{a,(f . a)} = {a1,(f . a1)}
Lm70:
for a, a1 being set
for f being Function st ( a in {a1,(f . a1)} or f . a in {a1,(f . a1)} ) & f is involutive & a in dom f & a1 in dom f holds
{a,(f . a)} = {a1,(f . a1)}
Lm71:
for X being set
for f being Function st X c= f holds
dom (f \ X) = (dom f) \ (proj1 X)
Lm72:
for X being set st X <> {} holds
{ x where x is Element of X : verum } = X
Lm73:
for a1, a2, b1, b2 being set
for f, g being Function holds
( [[a1,a2],[b1,b2]] in [:f,g:] iff ( [a1,b1] in f & [a2,b2] in g ) )
Lm74:
for X, Y being set holds
( X = Y iff for x being set st x in X \/ Y holds
( x in X iff x in Y ) )
Lm75:
for x being set
for f, g being Function st x in [:f,g:] holds
x = [[((x `1) `1),((x `1) `2)],[((x `2) `1),((x `2) `2)]]
Lm76:
for f1, f2, g1, g2 being Function holds [:f1,f2:] /\ [:g1,g2:] = [:(f1 /\ g1),(f2 /\ g2):]
Lm77:
for R being Relation holds fixpoints1 R = fixpoints1 (R ~)
Lm78:
for f being Function
for x being object holds
( x in fixpoints f iff x is_a_fixpoint_of f )
Lm79:
for f, g being Function holds fixpoints [:f,g:] = [:(fixpoints f),(fixpoints g):]
Lm80:
for X being set
for P being Relation st not X is empty & X is trivial & P is X -valued holds
P is constant Function
::#######################################################################