begin
:: deftheorem UNIALG_2:def 1 :
canceled;
:: deftheorem Def2 defines are_similar UNIALG_2:def 2 :
for U1, U2 being Universal_Algebra holds
( U1,U2 are_similar iff signature U1 = signature U2 );
theorem
canceled;
theorem
canceled;
theorem
theorem
theorem
:: deftheorem defines Operations UNIALG_2:def 3 :
for U0 being Universal_Algebra holds Operations U0 = rng the charact of U0;
:: deftheorem Def4 defines is_closed_on UNIALG_2:def 4 :
for U0 being Universal_Algebra
for A being Subset of U0
for o being operation of U0 holds
( A is_closed_on o iff for s being FinSequence of A st len s = arity o holds
o . s in A );
:: deftheorem Def5 defines opers_closed UNIALG_2:def 5 :
for U0 being Universal_Algebra
for A being Subset of U0 holds
( A is opers_closed iff for o being operation of U0 holds A is_closed_on o );
:: deftheorem Def6 defines /. UNIALG_2:def 6 :
for U0 being Universal_Algebra
for A being non empty Subset of U0
for o being operation of U0 st A is_closed_on o holds
o /. A = o | ((arity o) -tuples_on A);
:: deftheorem Def7 defines Opers UNIALG_2:def 7 :
for U0 being Universal_Algebra
for A being non empty Subset of U0
for b3 being PFuncFinSequence of A holds
( b3 = Opers (U0,A) iff ( dom b3 = dom the charact of U0 & ( for n being set
for o being operation of U0 st n in dom b3 & o = the charact of U0 . n holds
b3 . n = o /. A ) ) );
theorem
canceled;
theorem Th7:
theorem
:: deftheorem Def8 defines SubAlgebra UNIALG_2:def 8 :
for U0, b2 being Universal_Algebra holds
( b2 is SubAlgebra of U0 iff ( the carrier of b2 is Subset of U0 & ( for B being non empty Subset of U0 st B = the carrier of b2 holds
( the charact of b2 = Opers (U0,B) & B is opers_closed ) ) ) );
theorem Th9:
theorem Th10:
theorem
theorem
theorem Th13:
theorem Th14:
theorem Th15:
theorem
theorem Th17:
:: deftheorem Def9 defines UniAlgSetClosed UNIALG_2:def 9 :
for U0 being Universal_Algebra
for A being non empty Subset of U0 st A is opers_closed holds
UniAlgSetClosed A = UAStr(# A,(Opers (U0,A)) #);
:: deftheorem Def10 defines /\ UNIALG_2:def 10 :
for U0 being Universal_Algebra
for U1, U2 being SubAlgebra of U0 st the carrier of U1 meets the carrier of U2 holds
for b4 being strict SubAlgebra of U0 holds
( b4 = U1 /\ U2 iff ( the carrier of b4 = the carrier of U1 /\ the carrier of U2 & ( for B being non empty Subset of U0 st B = the carrier of b4 holds
( the charact of b4 = Opers (U0,B) & B is opers_closed ) ) ) );
:: deftheorem defines Constants UNIALG_2:def 11 :
for U0 being Universal_Algebra holds Constants U0 = { a where a is Element of U0 : ex o being operation of U0 st
( arity o = 0 & a in rng o ) } ;
:: deftheorem Def12 defines with_const_op UNIALG_2:def 12 :
for IT being Universal_Algebra holds
( IT is with_const_op iff ex o being operation of IT st arity o = 0 );
theorem Th18:
theorem
theorem Th20:
:: deftheorem Def13 defines GenUnivAlg UNIALG_2:def 13 :
for U0 being Universal_Algebra
for A being Subset of U0 st ( Constants U0 <> {} or A <> {} ) holds
for b3 being strict SubAlgebra of U0 holds
( b3 = GenUnivAlg A iff ( A c= the carrier of b3 & ( for U1 being SubAlgebra of U0 st A c= the carrier of U1 holds
b3 is SubAlgebra of U1 ) ) );
theorem
theorem Th22:
:: deftheorem Def14 defines "\/" UNIALG_2:def 14 :
for U0 being Universal_Algebra
for U1, U2 being SubAlgebra of U0
for b4 being strict SubAlgebra of U0 holds
( b4 = U1 "\/" U2 iff for A being non empty Subset of U0 st A = the carrier of U1 \/ the carrier of U2 holds
b4 = GenUnivAlg A );
theorem Th23:
theorem Th24:
theorem Th25:
theorem Th26:
:: deftheorem Def15 defines Sub UNIALG_2:def 15 :
for U0 being Universal_Algebra
for b2 being set holds
( b2 = Sub U0 iff for x being set holds
( x in b2 iff x is strict SubAlgebra of U0 ) );
definition
let U0 be
Universal_Algebra;
func UniAlg_join U0 -> BinOp of
(Sub U0) means :
Def16:
for
x,
y being
Element of
Sub U0 for
U1,
U2 being
strict SubAlgebra of
U0 st
x = U1 &
y = U2 holds
it . (
x,
y)
= U1 "\/" U2;
existence
ex b1 being BinOp of (Sub U0) st
for x, y being Element of Sub U0
for U1, U2 being strict SubAlgebra of U0 st x = U1 & y = U2 holds
b1 . (x,y) = U1 "\/" U2
uniqueness
for b1, b2 being BinOp of (Sub U0) st ( for x, y being Element of Sub U0
for U1, U2 being strict SubAlgebra of U0 st x = U1 & y = U2 holds
b1 . (x,y) = U1 "\/" U2 ) & ( for x, y being Element of Sub U0
for U1, U2 being strict SubAlgebra of U0 st x = U1 & y = U2 holds
b2 . (x,y) = U1 "\/" U2 ) holds
b1 = b2
end;
:: deftheorem Def16 defines UniAlg_join UNIALG_2:def 16 :
for U0 being Universal_Algebra
for b2 being BinOp of (Sub U0) holds
( b2 = UniAlg_join U0 iff for x, y being Element of Sub U0
for U1, U2 being strict SubAlgebra of U0 st x = U1 & y = U2 holds
b2 . (x,y) = U1 "\/" U2 );
definition
let U0 be
Universal_Algebra;
func UniAlg_meet U0 -> BinOp of
(Sub U0) means :
Def17:
for
x,
y being
Element of
Sub U0 for
U1,
U2 being
strict SubAlgebra of
U0 st
x = U1 &
y = U2 holds
it . (
x,
y)
= U1 /\ U2;
existence
ex b1 being BinOp of (Sub U0) st
for x, y being Element of Sub U0
for U1, U2 being strict SubAlgebra of U0 st x = U1 & y = U2 holds
b1 . (x,y) = U1 /\ U2
uniqueness
for b1, b2 being BinOp of (Sub U0) st ( for x, y being Element of Sub U0
for U1, U2 being strict SubAlgebra of U0 st x = U1 & y = U2 holds
b1 . (x,y) = U1 /\ U2 ) & ( for x, y being Element of Sub U0
for U1, U2 being strict SubAlgebra of U0 st x = U1 & y = U2 holds
b2 . (x,y) = U1 /\ U2 ) holds
b1 = b2
end;
:: deftheorem Def17 defines UniAlg_meet UNIALG_2:def 17 :
for U0 being Universal_Algebra
for b2 being BinOp of (Sub U0) holds
( b2 = UniAlg_meet U0 iff for x, y being Element of Sub U0
for U1, U2 being strict SubAlgebra of U0 st x = U1 & y = U2 holds
b2 . (x,y) = U1 /\ U2 );
theorem Th27:
theorem Th28:
theorem Th29:
theorem Th30:
:: deftheorem defines UnSubAlLattice UNIALG_2:def 18 :
for U0 being with_const_op Universal_Algebra holds UnSubAlLattice U0 = LattStr(# (Sub U0),(UniAlg_join U0),(UniAlg_meet U0) #);