Lm1:
for f being non constant standard special_circular_sequence holds (L~ f) ` = the carrier of ((TOP-REAL 2) | ((L~ f) `))
Lm2:
the carrier of (TOP-REAL 2) = REAL 2
by EUCLID:22;
Lm3:
for f being non constant standard special_circular_sequence holds
( Cl (Down ((LeftComp f),((L~ f) `))) is connected & Down ((LeftComp f),((L~ f) `)) = LeftComp f & Down ((LeftComp f),((L~ f) `)) is a_component )
Lm4:
for f being non constant standard special_circular_sequence holds
( Cl (Down ((RightComp f),((L~ f) `))) is connected & Down ((RightComp f),((L~ f) `)) = RightComp f & Down ((RightComp f),((L~ f) `)) is a_component )
Lm5:
for f being non constant standard special_circular_sequence
for i1, j1, i2, j2 being Nat st i1 <= len (GoB f) & j1 <= width (GoB f) & i2 <= len (GoB f) & j2 <= width (GoB f) & ( i2 = i1 + 1 or i1 = i2 + 1 ) & j1 = j2 & Int (cell ((GoB f),i1,j1)) c= (LeftComp f) \/ (RightComp f) holds
Int (cell ((GoB f),i2,j2)) c= (LeftComp f) \/ (RightComp f)
Lm6:
for f being non constant standard special_circular_sequence
for i1, j1, i2, j2 being Nat st i1 <= len (GoB f) & j1 <= width (GoB f) & i2 <= len (GoB f) & j2 <= width (GoB f) & ( j2 = j1 + 1 or j1 = j2 + 1 ) & i1 = i2 & Int (cell ((GoB f),i1,j1)) c= (LeftComp f) \/ (RightComp f) holds
Int (cell ((GoB f),i2,j2)) c= (LeftComp f) \/ (RightComp f)
theorem Th6:
for
f being non
constant standard special_circular_sequence for
i1,
j1,
i2,
j2 being
Nat st
i1 <= len (GoB f) &
j1 <= width (GoB f) &
i2 <= len (GoB f) &
j2 <= width (GoB f) &
i1,
j1,
i2,
j2 are_adjacent holds
(
Int (cell ((GoB f),i1,j1)) c= (LeftComp f) \/ (RightComp f) iff
Int (cell ((GoB f),i2,j2)) c= (LeftComp f) \/ (RightComp f) )