begin
deffunc H1( 1-sorted ) -> set = the carrier of $1;
deffunc H2( multMagma ) -> Element of bool [:[: the carrier of $1, the carrier of $1:], the carrier of $1:] = the multF of $1;
:: deftheorem Def1 defines constituted-Functions MONOID_0:def 1 :
for IT being 1-sorted holds
( IT is constituted-Functions iff for a being Element of IT holds a is Function );
:: deftheorem Def2 defines constituted-FinSeqs MONOID_0:def 2 :
for IT being 1-sorted holds
( IT is constituted-FinSeqs iff for a being Element of IT holds a is FinSequence );
definition
let D be non
empty set ;
let IT be
BinOp of
D;
attr IT is
left-invertible means
for
a,
b being
Element of
D ex
l being
Element of
D st
IT . (
l,
a)
= b;
attr IT is
right-invertible means
for
a,
b being
Element of
D ex
r being
Element of
D st
IT . (
a,
r)
= b;
attr IT is
invertible means :
Def5:
for
a,
b being
Element of
D ex
r,
l being
Element of
D st
(
IT . (
a,
r)
= b &
IT . (
l,
a)
= b );
attr IT is
left-cancelable means
for
a,
b,
c being
Element of
D st
IT . (
a,
b)
= IT . (
a,
c) holds
b = c;
attr IT is
right-cancelable means
for
a,
b,
c being
Element of
D st
IT . (
b,
a)
= IT . (
c,
a) holds
b = c;
attr IT is
cancelable means
for
a,
b,
c being
Element of
D st (
IT . (
a,
b)
= IT . (
a,
c) or
IT . (
b,
a)
= IT . (
c,
a) ) holds
b = c;
attr IT is
uniquely-decomposable means
(
IT is
having_a_unity & ( for
a,
b being
Element of
D st
IT . (
a,
b)
= the_unity_wrt IT holds
(
a = b &
b = the_unity_wrt IT ) ) );
end;
:: deftheorem defines left-invertible MONOID_0:def 3 :
for D being non empty set
for IT being BinOp of D holds
( IT is left-invertible iff for a, b being Element of D ex l being Element of D st IT . (l,a) = b );
:: deftheorem defines right-invertible MONOID_0:def 4 :
for D being non empty set
for IT being BinOp of D holds
( IT is right-invertible iff for a, b being Element of D ex r being Element of D st IT . (a,r) = b );
:: deftheorem Def5 defines invertible MONOID_0:def 5 :
for D being non empty set
for IT being BinOp of D holds
( IT is invertible iff for a, b being Element of D ex r, l being Element of D st
( IT . (a,r) = b & IT . (l,a) = b ) );
:: deftheorem defines left-cancelable MONOID_0:def 6 :
for D being non empty set
for IT being BinOp of D holds
( IT is left-cancelable iff for a, b, c being Element of D st IT . (a,b) = IT . (a,c) holds
b = c );
:: deftheorem defines right-cancelable MONOID_0:def 7 :
for D being non empty set
for IT being BinOp of D holds
( IT is right-cancelable iff for a, b, c being Element of D st IT . (b,a) = IT . (c,a) holds
b = c );
:: deftheorem defines cancelable MONOID_0:def 8 :
for D being non empty set
for IT being BinOp of D holds
( IT is cancelable iff for a, b, c being Element of D st ( IT . (a,b) = IT . (a,c) or IT . (b,a) = IT . (c,a) ) holds
b = c );
:: deftheorem defines uniquely-decomposable MONOID_0:def 9 :
for D being non empty set
for IT being BinOp of D holds
( IT is uniquely-decomposable iff ( IT is having_a_unity & ( for a, b being Element of D st IT . (a,b) = the_unity_wrt IT holds
( a = b & b = the_unity_wrt IT ) ) ) );
theorem Th1:
theorem Th2:
theorem Th3:
begin
:: deftheorem Def10 defines unital MONOID_0:def 10 :
for IT being non empty multMagma holds
( IT is unital iff the multF of IT is having_a_unity );
:: deftheorem Def11 defines commutative MONOID_0:def 11 :
for G being non empty multMagma holds
( G is commutative iff the multF of G is commutative );
:: deftheorem Def12 defines associative MONOID_0:def 12 :
for G being non empty multMagma holds
( G is associative iff the multF of G is associative );
:: deftheorem defines idempotent MONOID_0:def 13 :
for IT being non empty multMagma holds
( IT is idempotent iff the multF of IT is idempotent );
:: deftheorem defines left-invertible MONOID_0:def 14 :
for IT being non empty multMagma holds
( IT is left-invertible iff the multF of IT is left-invertible );
:: deftheorem defines right-invertible MONOID_0:def 15 :
for IT being non empty multMagma holds
( IT is right-invertible iff the multF of IT is right-invertible );
:: deftheorem Def16 defines invertible MONOID_0:def 16 :
for IT being non empty multMagma holds
( IT is invertible iff the multF of IT is invertible );
:: deftheorem defines left-cancelable MONOID_0:def 17 :
for IT being non empty multMagma holds
( IT is left-cancelable iff the multF of IT is left-cancelable );
:: deftheorem defines right-cancelable MONOID_0:def 18 :
for IT being non empty multMagma holds
( IT is right-cancelable iff the multF of IT is right-cancelable );
:: deftheorem defines cancelable MONOID_0:def 19 :
for IT being non empty multMagma holds
( IT is cancelable iff the multF of IT is cancelable );
:: deftheorem Def20 defines uniquely-decomposable MONOID_0:def 20 :
for IT being non empty multMagma holds
( IT is uniquely-decomposable iff the multF of IT is uniquely-decomposable );
theorem Th4:
theorem
theorem Th6:
Lm1:
for G being non empty multMagma holds
( G is commutative iff for a, b being Element of G holds a * b = b * a )
Lm2:
for G being non empty multMagma holds
( G is associative iff for a, b, c being Element of G holds (a * b) * c = a * (b * c) )
theorem
canceled;
theorem
canceled;
theorem Th9:
theorem
theorem
theorem Th12:
theorem
theorem
theorem Th15:
theorem Th16:
theorem Th17:
Lm3:
for G being non empty multMagma holds
( G is invertible iff ( G is left-invertible & G is right-invertible ) )
Lm4:
for G being non empty multMagma holds
( G is cancelable iff ( G is left-cancelable & G is right-cancelable ) )
Lm5:
for G being non empty multMagma st G is associative & G is invertible holds
G is Group-like
begin
deffunc H3( multLoopStr ) -> Element of the carrier of $1 = 1. $1;
:: deftheorem Def21 defines well-unital MONOID_0:def 21 :
for IT being non empty multLoopStr holds
( IT is well-unital iff 1. IT is_a_unity_wrt the multF of IT );
theorem Th18:
theorem Th19:
:: deftheorem Def22 defines MonoidalExtension MONOID_0:def 22 :
for G being multMagma
for b2 being multLoopStr holds
( b2 is MonoidalExtension of G iff multMagma(# the carrier of b2, the multF of b2 #) = multMagma(# the carrier of G, the multF of G #) );
theorem Th20:
theorem Th21:
theorem Th22:
begin
:: deftheorem Def23 defines SubStr MONOID_0:def 23 :
for G, b2 being multMagma holds
( b2 is SubStr of G iff the multF of b2 c= the multF of G );
:: deftheorem Def24 defines MonoidalSubStr MONOID_0:def 24 :
for G being multMagma
for b2 being multLoopStr holds
( b2 is MonoidalSubStr of G iff ( the multF of b2 c= the multF of G & ( for M being multLoopStr st G = M holds
1. b2 = 1. M ) ) );
:: deftheorem Def25 defines MonoidalSubStr MONOID_0:def 25 :
for M, b2 being multLoopStr holds
( b2 is MonoidalSubStr of M iff ( the multF of b2 c= the multF of M & 1. b2 = 1. M ) );
theorem Th23:
theorem
theorem Th25:
theorem Th26:
theorem Th27:
theorem Th28:
theorem
theorem Th30:
theorem
theorem Th32:
theorem Th33:
theorem Th34:
theorem Th35:
theorem Th36:
theorem Th37:
theorem Th38:
theorem Th39:
Lm7:
now
let G be non
empty multLoopStr ;
for D being non empty Subset of G st ( for x, y being Element of D holds x * y in D ) & 1. G in D holds
ex H being non empty strict MonoidalSubStr of G st the carrier of H = Dlet D be non
empty Subset of
G;
( ( for x, y being Element of D holds x * y in D ) & 1. G in D implies ex H being non empty strict MonoidalSubStr of G st the carrier of H = D )assume that A1:
for
x,
y being
Element of
D holds
x * y in D
and A2:
1. G in D
;
ex H being non empty strict MonoidalSubStr of G st the carrier of H = Dthus
ex
H being non
empty strict MonoidalSubStr of
G st the
carrier of
H = D
verum
proof
consider H being non
empty strict SubStr of
G such that A4:
the
carrier of
H = D
by Lm6, A1;
reconsider e =
1. G as
Element of
H by A2, A4;
set N =
multLoopStr(# the
carrier of
H, the
multF of
H,
e #);
(
H2(
multLoopStr(# the
carrier of
H, the
multF of
H,
e #))
c= H2(
G) & ( for
M being
multLoopStr st
G = M holds
1. multLoopStr(# the
carrier of
H, the
multF of
H,
e #)
= 1. M ) )
by Def23;
then reconsider N =
multLoopStr(# the
carrier of
H, the
multF of
H,
e #) as non
empty strict MonoidalSubStr of
G by Def25;
take
N
;
the carrier of N = D
thus
the
carrier of
N = D
by A4;
verum
end;
end;
begin
:: deftheorem defines <REAL,+> MONOID_0:def 26 :
<REAL,+> = multMagma(# REAL,addreal #);
theorem
canceled;
theorem
theorem
canceled;
theorem Th43:
theorem Th44:
theorem
canceled;
theorem
theorem
:: deftheorem Def27 defines <NAT,+> MONOID_0:def 27 :
for b1 being non empty strict unital uniquely-decomposable SubStr of INT.Group holds
( b1 = <NAT,+> iff the carrier of b1 = NAT );
:: deftheorem defines <NAT,+,0> MONOID_0:def 28 :
for b1 being non empty strict well-unital MonoidalExtension of <NAT,+> holds
( b1 = <NAT,+,0> iff verum );
:: deftheorem defines addnat MONOID_0:def 29 :
addnat = the multF of <NAT,+>;
theorem
canceled;
theorem Th49:
theorem
theorem
theorem
theorem
theorem
:: deftheorem defines <REAL,*> MONOID_0:def 30 :
<REAL,*> = multMagma(# REAL,multreal #);
theorem
canceled;
theorem
theorem
canceled;
theorem Th58:
theorem
canceled;
theorem
:: deftheorem Def31 defines <NAT,*> MONOID_0:def 31 :
for b1 being non empty strict unital uniquely-decomposable SubStr of <REAL,*> holds
( b1 = <NAT,*> iff the carrier of b1 = NAT );
:: deftheorem defines <NAT,*,1> MONOID_0:def 32 :
for b1 being non empty strict well-unital MonoidalExtension of <NAT,*> holds
( b1 = <NAT,*,1> iff verum );
:: deftheorem defines multnat MONOID_0:def 33 :
multnat = the multF of <NAT,*>;
theorem Th61:
theorem
theorem Th63:
theorem
theorem
theorem
theorem
begin
:: deftheorem Def34 defines *+^ MONOID_0:def 34 :
for D being non empty set
for b2 being non empty strict unital associative constituted-FinSeqs cancelable uniquely-decomposable multMagma holds
( b2 = D *+^ iff ( the carrier of b2 = D * & ( for p, q being Element of b2 holds p [*] q = p ^ q ) ) );
:: deftheorem defines *+^+<0> MONOID_0:def 35 :
for D being non empty set
for b2 being non empty strict well-unital MonoidalExtension of D *+^ holds
( b2 = D *+^+<0> iff verum );
:: deftheorem defines -concatenation MONOID_0:def 36 :
for D being non empty set holds D -concatenation = the multF of (D *+^);
theorem
theorem Th69:
theorem
theorem
theorem Th72:
theorem
theorem
theorem
theorem
begin
:: deftheorem Def37 defines GPFuncs MONOID_0:def 37 :
for X being set
for b2 being strict constituted-Functions multMagma holds
( b2 = GPFuncs X iff ( the carrier of b2 = PFuncs (X,X) & ( for f, g being Element of b2 holds f [*] g = g (*) f ) ) );
:: deftheorem defines MPFuncs MONOID_0:def 38 :
for X being set
for b2 being non empty strict well-unital MonoidalExtension of GPFuncs X holds
( b2 = MPFuncs X iff verum );
:: deftheorem defines -composition MONOID_0:def 39 :
for X being set holds X -composition = the multF of (GPFuncs X);
theorem
theorem Th78:
theorem Th79:
theorem Th80:
theorem
:: deftheorem Def40 defines GFuncs MONOID_0:def 40 :
for X being set
for b2 being strict SubStr of GPFuncs X holds
( b2 = GFuncs X iff the carrier of b2 = Funcs (X,X) );
:: deftheorem defines MFuncs MONOID_0:def 41 :
for X being set
for b2 being strict well-unital MonoidalExtension of GFuncs X holds
( b2 = MFuncs X iff verum );
theorem
theorem Th83:
theorem Th84:
theorem
:: deftheorem Def42 defines GPerms MONOID_0:def 42 :
for X being set
for b2 being non empty strict SubStr of GFuncs X holds
( b2 = GPerms X iff for f being Element of (GFuncs X) holds
( f in the carrier of b2 iff f is Permutation of X ) );
theorem Th86:
theorem Th87:
theorem
theorem
theorem
theorem