:: SUBSTLAT semantic presentation
definition
let c
1, c
2 be
set ;
func SubstitutionSet c
1,c
2 -> Subset of
(Fin (PFuncs a1,a2)) equals :: SUBSTLAT:def 1
{ b1 where B is Element of Fin (PFuncs a1,a2) : ( ( for b1 being set holds
( b2 in b1 implies b2 is finite ) ) & ( for b1, b2 being Element of PFuncs a1,a2 holds
( b2 in b1 & b3 in b1 & b2 c= b3 implies b2 = b3 ) ) ) } ;
coherence
{ b1 where B is Element of Fin (PFuncs c1,c2) : ( ( for b1 being set holds
( b2 in b1 implies b2 is finite ) ) & ( for b1, b2 being Element of PFuncs c1,c2 holds
( b2 in b1 & b3 in b1 & b2 c= b3 implies b2 = b3 ) ) ) } is Subset of (Fin (PFuncs c1,c2))
end;
:: deftheorem Def1 defines SubstitutionSet SUBSTLAT:def 1 :
Lemma1:
for b1, b2, b3, b4 being set holds
( b4 in SubstitutionSet b1,b2 & b3 in b4 implies b3 is finite )
theorem Th1: :: SUBSTLAT:1
theorem Th2: :: SUBSTLAT:2
:: deftheorem Def2 defines mi SUBSTLAT:def 2 :
definition
let c
1, c
2 be
set ;
let c
3, c
4 be
Element of
Fin (PFuncs c1,c2);
func c
3 ^ c
4 -> Element of
Fin (PFuncs a1,a2) equals :: SUBSTLAT:def 3
{ (b1 \/ b2) where B is Element of PFuncs a1,a2, B is Element of PFuncs a1,a2 : ( b1 in a3 & b2 in a4 & b1 tolerates b2 ) } ;
coherence
{ (b1 \/ b2) where B is Element of PFuncs c1,c2, B is Element of PFuncs c1,c2 : ( b1 in c3 & b2 in c4 & b1 tolerates b2 ) } is Element of Fin (PFuncs c1,c2)
end;
:: deftheorem Def3 defines ^ SUBSTLAT:def 3 :
theorem Th3: :: SUBSTLAT:3
theorem Th4: :: SUBSTLAT:4
theorem Th5: :: SUBSTLAT:5
theorem Th6: :: SUBSTLAT:6
for b
1, b
2 being
set for b
3 being
Element of
Fin (PFuncs b1,b2)for b
4 being
set holds
( b
4 in mi b
3 implies ( b
4 in b
3 & ( for b
5 being
set holds
( b
5 in b
3 & b
5 c= b
4 implies b
5 = b
4 ) ) ) )
Lemma8:
for b1, b2 being set
for b3, b4 being Element of Fin (PFuncs b1,b2) holds
( ( for b5 being set holds
( b5 in b3 implies b5 in b4 ) ) implies b3 c= b4 )
theorem Th7: :: SUBSTLAT:7
theorem Th8: :: SUBSTLAT:8
theorem Th9: :: SUBSTLAT:9
theorem Th10: :: SUBSTLAT:10
theorem Th11: :: SUBSTLAT:11
theorem Th12: :: SUBSTLAT:12
theorem Th13: :: SUBSTLAT:13
theorem Th14: :: SUBSTLAT:14
theorem Th15: :: SUBSTLAT:15
for b
1, b
2 being
set for b
3, b
4 being
Element of
Fin (PFuncs b1,b2)for b
5 being
set holds
not ( b
5 in b
3 ^ b
4 & ( for b
6, b
7 being
set holds
not ( b
6 in b
3 & b
7 in b
4 & b
5 = b
6 \/ b
7 ) ) )
theorem Th16: :: SUBSTLAT:16
Lemma17:
for b1, b2 being set
for b3, b4 being Element of Fin (PFuncs b1,b2)
for b5 being finite set holds
not ( b5 in b3 ^ b4 & ( for b6 being finite set holds
not ( b6 c= b5 & b6 in (mi b3) ^ b4 ) ) )
theorem Th17: :: SUBSTLAT:17
theorem Th18: :: SUBSTLAT:18
theorem Th19: :: SUBSTLAT:19
theorem Th20: :: SUBSTLAT:20
theorem Th21: :: SUBSTLAT:21
theorem Th22: :: SUBSTLAT:22
Lemma24:
for b1, b2 being set
for b3, b4 being Element of Fin (PFuncs b1,b2)
for b5 being set holds
not ( b5 in b3 ^ b4 & ( for b6 being set holds
not ( b6 in b4 & b6 c= b5 ) ) )
Lemma25:
for b1, b2 being set
for b3, b4 being Element of Fin (PFuncs b1,b2) holds mi ((b3 ^ b4) \/ b4) = mi b4
theorem Th23: :: SUBSTLAT:23
theorem Th24: :: SUBSTLAT:24
theorem Th25: :: SUBSTLAT:25
definition
let c
1, c
2 be
set ;
func SubstLatt c
1,c
2 -> strict LattStr means :
Def4:
:: SUBSTLAT:def 4
( the
carrier of a
3 = SubstitutionSet a
1,a
2 & ( for b
1, b
2 being
Element of
SubstitutionSet a
1,a
2 holds
( the
L_join of a
3 . b
1,b
2 = mi (b1 \/ b2) & the
L_meet of a
3 . b
1,b
2 = mi (b1 ^ b2) ) ) );
existence
ex b1 being strict LattStr st
( the carrier of b1 = SubstitutionSet c1,c2 & ( for b2, b3 being Element of SubstitutionSet c1,c2 holds
( the L_join of b1 . b2,b3 = mi (b2 \/ b3) & the L_meet of b1 . b2,b3 = mi (b2 ^ b3) ) ) )
uniqueness
for b1, b2 being strict LattStr holds
( the carrier of b1 = SubstitutionSet c1,c2 & ( for b3, b4 being Element of SubstitutionSet c1,c2 holds
( the L_join of b1 . b3,b4 = mi (b3 \/ b4) & the L_meet of b1 . b3,b4 = mi (b3 ^ b4) ) ) & the carrier of b2 = SubstitutionSet c1,c2 & ( for b3, b4 being Element of SubstitutionSet c1,c2 holds
( the L_join of b2 . b3,b4 = mi (b3 \/ b4) & the L_meet of b2 . b3,b4 = mi (b3 ^ b4) ) ) implies b1 = b2 )
end;
:: deftheorem Def4 defines SubstLatt SUBSTLAT:def 4 :
Lemma29:
for b1, b2 being set
for b3, b4 being Element of (SubstLatt b1,b2) holds b3 "\/" b4 = b4 "\/" b3
Lemma30:
for b1, b2 being set
for b3, b4, b5 being Element of (SubstLatt b1,b2) holds b3 "\/" (b4 "\/" b5) = (b3 "\/" b4) "\/" b5
Lemma31:
for b1, b2 being set
for b3, b4 being Element of SubstitutionSet b1,b2 holds the L_join of (SubstLatt b1,b2) . (the L_meet of (SubstLatt b1,b2) . b3,b4),b4 = b4
Lemma32:
for b1, b2 being set
for b3, b4 being Element of (SubstLatt b1,b2) holds (b3 "/\" b4) "\/" b4 = b4
Lemma33:
for b1, b2 being set
for b3, b4 being Element of (SubstLatt b1,b2) holds b3 "/\" b4 = b4 "/\" b3
Lemma34:
for b1, b2 being set
for b3, b4, b5 being Element of (SubstLatt b1,b2) holds b3 "/\" (b4 "/\" b5) = (b3 "/\" b4) "/\" b5
Lemma35:
for b1, b2 being set
for b3, b4, b5 being Element of SubstitutionSet b1,b2 holds the L_meet of (SubstLatt b1,b2) . b3,(the L_join of (SubstLatt b1,b2) . b4,b5) = the L_join of (SubstLatt b1,b2) . (the L_meet of (SubstLatt b1,b2) . b3,b4),(the L_meet of (SubstLatt b1,b2) . b3,b5)
Lemma36:
for b1, b2 being set
for b3, b4 being Element of (SubstLatt b1,b2) holds b3 "/\" (b3 "\/" b4) = b3
theorem Th26: :: SUBSTLAT:26
theorem Th27: :: SUBSTLAT:27