:: Projections in n-Dimensional Euclidean Space to Each Coordinates
:: by Roman Matuszewski and Yatsuka Nakamura
::
:: Received November 3, 1997
:: Copyright (c) 1997 Association of Mizar Users


begin

Lm1: for n, i being Nat
for p being Element of (TOP-REAL n) ex q being Real ex g being FinSequence of REAL st
( g = p & q = g /. i )
proof end;

registration
let n be Nat;
cluster -> REAL -valued Element of the carrier of (TOP-REAL n);
coherence
for b1 being Element of (TOP-REAL n) holds b1 is REAL -valued
proof end;
end;

theorem :: JORDAN2B:1
for i being Element of NAT
for n being Nat ex f being Function of (TOP-REAL n),R^1 st
for p being Element of (TOP-REAL n) holds f . p = p /. i
proof end;

theorem :: JORDAN2B:2
canceled;

theorem :: JORDAN2B:3
for n, i being Element of NAT st i in Seg n holds
(0. (TOP-REAL n)) /. i = 0
proof end;

theorem :: JORDAN2B:4
for n being Element of NAT
for r being real number
for p being Point of (TOP-REAL n)
for i being Element of NAT st i in Seg n holds
(r * p) /. i = r * (p /. i)
proof end;

theorem :: JORDAN2B:5
for n being Element of NAT
for p being Point of (TOP-REAL n)
for i being Element of NAT st i in Seg n holds
(- p) /. i = - (p /. i)
proof end;

theorem :: JORDAN2B:6
for n being Element of NAT
for p1, p2 being Point of (TOP-REAL n)
for i being Element of NAT st i in Seg n holds
(p1 + p2) /. i = (p1 /. i) + (p2 /. i)
proof end;

theorem Th7: :: JORDAN2B:7
for n being Element of NAT
for p1, p2 being Point of (TOP-REAL n)
for i being Element of NAT st i in Seg n holds
(p1 - p2) /. i = (p1 /. i) - (p2 /. i)
proof end;

theorem :: JORDAN2B:8
canceled;

theorem Th9: :: JORDAN2B:9
for i, n being Element of NAT st i <= n holds
(0* n) | i = 0* i
proof end;

theorem Th10: :: JORDAN2B:10
for n, i being Element of NAT holds (0* n) /^ i = 0* (n -' i)
proof end;

theorem Th11: :: JORDAN2B:11
for i being Element of NAT holds Sum (0* i) = 0
proof end;

theorem :: JORDAN2B:12
canceled;

theorem :: JORDAN2B:13
canceled;

theorem Th14: :: JORDAN2B:14
for i, n being Element of NAT
for r being Real st i in Seg n holds
Sum ((0* n) +* (i,r)) = r
proof end;

theorem Th15: :: JORDAN2B:15
for n being Element of NAT
for q being Element of REAL n
for p being Point of (TOP-REAL n)
for i being Element of NAT st i in Seg n & q = p holds
( p /. i <= |.q.| & (p /. i) ^2 <= |.q.| ^2 )
proof end;

begin

theorem Th16: :: JORDAN2B:16
for n being Element of NAT
for s1 being real number
for P being Subset of (TOP-REAL n)
for i being Element of NAT st P = { p where p is Point of (TOP-REAL n) : s1 > p /. i } & i in Seg n holds
P is open
proof end;

theorem Th17: :: JORDAN2B:17
for n being Element of NAT
for s1 being real number
for P being Subset of (TOP-REAL n)
for i being Element of NAT st P = { p where p is Point of (TOP-REAL n) : s1 < p /. i } & i in Seg n holds
P is open
proof end;

theorem Th18: :: JORDAN2B:18
for n being Element of NAT
for P being Subset of (TOP-REAL n)
for a, b being real number
for i being Element of NAT st P = { p where p is Element of (TOP-REAL n) : ( a < p /. i & p /. i < b ) } & i in Seg n holds
P is open
proof end;

theorem Th19: :: JORDAN2B:19
for n being Element of NAT
for a, b being real number
for f being Function of (TOP-REAL n),R^1
for i being Element of NAT st ( for p being Element of (TOP-REAL n) holds f . p = p /. i ) holds
f " { s where s is Real : ( a < s & s < b ) } = { p where p is Element of (TOP-REAL n) : ( a < p /. i & p /. i < b ) }
proof end;

theorem Th20: :: JORDAN2B:20
for X being non empty TopSpace
for M being non empty MetrSpace
for f being Function of X,(TopSpaceMetr M) st ( for r being real number
for u being Element of M
for P being Subset of (TopSpaceMetr M) st r > 0 & P = Ball (u,r) holds
f " P is open ) holds
f is continuous
proof end;

theorem Th21: :: JORDAN2B:21
for u being Point of RealSpace
for r, u1 being real number st u1 = u holds
Ball (u,r) = { s where s is Real : ( u1 - r < s & s < u1 + r ) }
proof end;

theorem Th22: :: JORDAN2B:22
for n being Element of NAT
for f being Function of (TOP-REAL n),R^1
for i being Element of NAT st i in Seg n & ( for p being Element of (TOP-REAL n) holds f . p = p /. i ) holds
f is continuous
proof end;

begin

theorem :: JORDAN2B:23
for s being Real holds abs <*s*> = <*(abs s)*>
proof end;

theorem Th24: :: JORDAN2B:24
for p being Element of (TOP-REAL 1) ex r being Real st p = <*r*>
proof end;

notation
let r be real number ;
synonym |[r]| for <*r*>;
end;

definition
let r be real number ;
:: original: |[
redefine func |[r]| -> Point of (TOP-REAL 1);
coherence
|[r]| is Point of (TOP-REAL 1)
proof end;
end;

theorem :: JORDAN2B:25
canceled;

theorem :: JORDAN2B:26
for r being real number
for s being Real holds s * |[r]| = |[(s * r)]|
proof end;

theorem :: JORDAN2B:27
for r1, r2 being real number holds |[(r1 + r2)]| = |[r1]| + |[r2]|
proof end;

theorem :: JORDAN2B:28
canceled;

theorem :: JORDAN2B:29
for r1, r2 being real number st |[r1]| = |[r2]| holds
r1 = r2 by FINSEQ_1:97;

theorem Th30: :: JORDAN2B:30
for P being Subset of R^1
for b being real number st P = { s where s is Real : s < b } holds
P is open
proof end;

theorem Th31: :: JORDAN2B:31
for P being Subset of R^1
for a being real number st P = { s where s is Real : a < s } holds
P is open
proof end;

theorem Th32: :: JORDAN2B:32
for P being Subset of R^1
for a, b being real number st P = { s where s is Real : ( a < s & s < b ) } holds
P is open
proof end;

theorem Th33: :: JORDAN2B:33
for u being Point of (Euclid 1)
for r, u1 being real number st <*u1*> = u holds
Ball (u,r) = { <*s*> where s is Real : ( u1 - r < s & s < u1 + r ) }
proof end;

theorem :: JORDAN2B:34
for f being Function of (TOP-REAL 1),R^1 st ( for p being Element of (TOP-REAL 1) holds f . p = p /. 1 ) holds
f is being_homeomorphism
proof end;

theorem Th35: :: JORDAN2B:35
for p being Element of (TOP-REAL 2) holds
( p /. 1 = p `1 & p /. 2 = p `2 )
proof end;

theorem :: JORDAN2B:36
for p being Element of (TOP-REAL 2) holds
( p /. 1 = proj1 . p & p /. 2 = proj2 . p )
proof end;