:: MATRTOP3 semantic presentation begin theorem Th1: :: MATRTOP3:1 for n being Nat for K being Field for M being Matrix of n,K for P being Permutation of (Seg n) holds ( Det ((((M * P) @) * P) @) = Det M & ( for i, j being Nat st [i,j] in Indices M holds ((((M * P) @) * P) @) * (i,j) = M * ((P . i),(P . j)) ) ) proof let n be Nat; ::_thesis: for K being Field for M being Matrix of n,K for P being Permutation of (Seg n) holds ( Det ((((M * P) @) * P) @) = Det M & ( for i, j being Nat st [i,j] in Indices M holds ((((M * P) @) * P) @) * (i,j) = M * ((P . i),(P . j)) ) ) let K be Field; ::_thesis: for M being Matrix of n,K for P being Permutation of (Seg n) holds ( Det ((((M * P) @) * P) @) = Det M & ( for i, j being Nat st [i,j] in Indices M holds ((((M * P) @) * P) @) * (i,j) = M * ((P . i),(P . j)) ) ) let M be Matrix of n,K; ::_thesis: for P being Permutation of (Seg n) holds ( Det ((((M * P) @) * P) @) = Det M & ( for i, j being Nat st [i,j] in Indices M holds ((((M * P) @) * P) @) * (i,j) = M * ((P . i),(P . j)) ) ) let P be Permutation of (Seg n); ::_thesis: ( Det ((((M * P) @) * P) @) = Det M & ( for i, j being Nat st [i,j] in Indices M holds ((((M * P) @) * P) @) * (i,j) = M * ((P . i),(P . j)) ) ) reconsider p = P as Element of Permutations n by MATRIX_2:def_9; A1: - (- (Det M)) = Det M by RLVECT_1:17; A2: ( ( p is even & - ((Det M),p) = Det M ) or ( p is odd & - ((Det M),p) = - (Det M) ) ) by MATRIX_2:def_13; thus Det ((((M * P) @) * P) @) = Det (((M * P) @) * P) by MATRIXR2:43 .= - ((Det ((M * P) @)),p) by MATRIX11:46 .= - ((Det (M * P)),p) by MATRIXR2:43 .= - ((- ((Det M),p)),p) by MATRIX11:46 .= Det M by A1, A2, MATRIX_2:def_13 ; ::_thesis: for i, j being Nat st [i,j] in Indices M holds ((((M * P) @) * P) @) * (i,j) = M * ((P . i),(P . j)) let i, j be Nat; ::_thesis: ( [i,j] in Indices M implies ((((M * P) @) * P) @) * (i,j) = M * ((P . i),(P . j)) ) assume A3: [i,j] in Indices M ; ::_thesis: ((((M * P) @) * P) @) * (i,j) = M * ((P . i),(P . j)) Indices M = Indices ((((M * P) @) * P) @) by MATRIX_1:26; then A4: [j,i] in Indices (((M * P) @) * P) by A3, MATRIX_1:def_6; then A5: ((((M * P) @) * P) @) * (i,j) = (((M * P) @) * P) * (j,i) by MATRIX_1:def_6; ( Indices M = Indices (((M * P) @) * P) & Indices M = Indices ((M * P) @) ) by MATRIX_1:26; then A6: ex k being Nat st ( k = P . j & [k,i] in Indices ((M * P) @) & (((M * P) @) * P) * (j,i) = ((M * P) @) * (k,i) ) by A4, MATRIX11:37; then A7: [i,(P . j)] in Indices (M * P) by MATRIX_1:def_6; Indices (M * P) = Indices M by MATRIX_1:26; then (M * P) * (i,(P . j)) = M * ((P . i),(P . j)) by A7, MATRIX11:def_4; hence ((((M * P) @) * P) @) * (i,j) = M * ((P . i),(P . j)) by A5, A6, A7, MATRIX_1:def_6; ::_thesis: verum end; theorem Th2: :: MATRTOP3:2 for n being Nat for K being Field for M being V244(b2) Matrix of n,K holds M @ = M proof let n be Nat; ::_thesis: for K being Field for M being V244(b1) Matrix of n,K holds M @ = M let K be Field; ::_thesis: for M being V244(K) Matrix of n,K holds M @ = M let M be V244(K) Matrix of n,K; ::_thesis: M @ = M for i, j being Nat st [i,j] in Indices M holds M * (i,j) = (M @) * (i,j) proof let i, j be Nat; ::_thesis: ( [i,j] in Indices M implies M * (i,j) = (M @) * (i,j) ) assume A1: [i,j] in Indices M ; ::_thesis: M * (i,j) = (M @) * (i,j) then A2: [j,i] in Indices M by MATRIX_1:28; then A3: (M @) * (i,j) = M * (j,i) by MATRIX_1:def_6; percases ( i = j or i <> j ) ; suppose i = j ; ::_thesis: M * (i,j) = (M @) * (i,j) hence M * (i,j) = (M @) * (i,j) by A1, MATRIX_1:def_6; ::_thesis: verum end; supposeA4: i <> j ; ::_thesis: M * (i,j) = (M @) * (i,j) then M * (i,j) = 0. K by A1, MATRIX_1:def_14; hence M * (i,j) = (M @) * (i,j) by A2, A3, A4, MATRIX_1:def_14; ::_thesis: verum end; end; end; hence M @ = M by MATRIX_1:27; ::_thesis: verum end; theorem Th3: :: MATRTOP3:3 for r being real number for f being real-valued FinSequence for i being Nat st i in dom f holds Sum (sqr (f +* (i,r))) = ((Sum (sqr f)) - ((f . i) ^2)) + (r ^2) proof let r be real number ; ::_thesis: for f being real-valued FinSequence for i being Nat st i in dom f holds Sum (sqr (f +* (i,r))) = ((Sum (sqr f)) - ((f . i) ^2)) + (r ^2) let f be real-valued FinSequence; ::_thesis: for i being Nat st i in dom f holds Sum (sqr (f +* (i,r))) = ((Sum (sqr f)) - ((f . i) ^2)) + (r ^2) let i be Nat; ::_thesis: ( i in dom f implies Sum (sqr (f +* (i,r))) = ((Sum (sqr f)) - ((f . i) ^2)) + (r ^2) ) assume A1: i in dom f ; ::_thesis: Sum (sqr (f +* (i,r))) = ((Sum (sqr f)) - ((f . i) ^2)) + (r ^2) reconsider fi = f . i as Element of REAL ; set F = @ (@ f); set G = (@ (@ f)) | (i -' 1); set H = (@ (@ f)) /^ i; A2: sqr <*fi*> = <*(fi ^2)*> by RVSUM_1:55; @ (@ f) = (@ (@ f)) +* (i,fi) by FUNCT_7:35 .= (((@ (@ f)) | (i -' 1)) ^ <*fi*>) ^ ((@ (@ f)) /^ i) by A1, FUNCT_7:98 ; then sqr (@ (@ f)) = (sqr (((@ (@ f)) | (i -' 1)) ^ <*fi*>)) ^ (sqr ((@ (@ f)) /^ i)) by RVSUM_1:144 .= ((sqr ((@ (@ f)) | (i -' 1))) ^ (sqr <*fi*>)) ^ (sqr ((@ (@ f)) /^ i)) by RVSUM_1:144 ; then A3: Sum (sqr (@ (@ f))) = (Sum ((sqr ((@ (@ f)) | (i -' 1))) ^ (sqr <*fi*>))) + (Sum (sqr ((@ (@ f)) /^ i))) by RVSUM_1:75 .= ((Sum (sqr ((@ (@ f)) | (i -' 1)))) + (fi ^2)) + (Sum (sqr ((@ (@ f)) /^ i))) by A2, RVSUM_1:74 ; reconsider R = r as Element of REAL by XREAL_0:def_1; A4: sqr <*R*> = <*(R ^2)*> by RVSUM_1:55; (@ (@ f)) +* (i,R) = (((@ (@ f)) | (i -' 1)) ^ <*R*>) ^ ((@ (@ f)) /^ i) by A1, FUNCT_7:98; then sqr ((@ (@ f)) +* (i,R)) = (sqr (((@ (@ f)) | (i -' 1)) ^ <*R*>)) ^ (sqr ((@ (@ f)) /^ i)) by RVSUM_1:144 .= ((sqr ((@ (@ f)) | (i -' 1))) ^ (sqr <*R*>)) ^ (sqr ((@ (@ f)) /^ i)) by RVSUM_1:144 ; then Sum (sqr ((@ (@ f)) +* (i,R))) = (Sum ((sqr ((@ (@ f)) | (i -' 1))) ^ (sqr <*R*>))) + (Sum (sqr ((@ (@ f)) /^ i))) by RVSUM_1:75 .= ((Sum (sqr ((@ (@ f)) | (i -' 1)))) + (R ^2)) + (Sum (sqr ((@ (@ f)) /^ i))) by A4, RVSUM_1:74 ; hence Sum (sqr (f +* (i,r))) = ((Sum (sqr f)) - ((f . i) ^2)) + (r ^2) by A3; ::_thesis: verum end; definition let X be set ; let F be Function-yielding Function; attrF is X -support-yielding means :Def1: :: MATRTOP3:def 1 for f being Function for x being set st f in dom F & (F . f) . x <> f . x holds x in X; end; :: deftheorem Def1 defines -support-yielding MATRTOP3:def_1_:_ for X being set for F being Function-yielding Function holds ( F is X -support-yielding iff for f being Function for x being set st f in dom F & (F . f) . x <> f . x holds x in X ); registration let X be set ; cluster Relation-like Function-like Function-yielding V235() X -support-yielding for set ; existence ex b1 being Function-yielding Function st b1 is X -support-yielding proof reconsider F = {} as Function-yielding Function ; for f being Function for x being set st f in dom F & (F . f) . x <> f . x holds x in X ; then F is X -support-yielding by Def1; hence ex b1 being Function-yielding Function st b1 is X -support-yielding ; ::_thesis: verum end; end; registration let X be set ; let Y be Subset of X; cluster Relation-like Function-like Function-yielding Y -support-yielding -> Function-yielding X -support-yielding for set ; coherence for b1 being Function-yielding Function st b1 is Y -support-yielding holds b1 is X -support-yielding proof let F be Function-yielding Function; ::_thesis: ( F is Y -support-yielding implies F is X -support-yielding ) assume A1: F is Y -support-yielding ; ::_thesis: F is X -support-yielding let f be Function; :: according to MATRTOP3:def_1 ::_thesis: for x being set st f in dom F & (F . f) . x <> f . x holds x in X let x be set ; ::_thesis: ( f in dom F & (F . f) . x <> f . x implies x in X ) assume ( f in dom F & (F . f) . x <> f . x ) ; ::_thesis: x in X then x in Y by A1, Def1; hence x in X ; ::_thesis: verum end; end; registration let X, Y be set ; cluster Relation-like Function-like Function-yielding X -support-yielding Y -support-yielding -> Function-yielding X /\ Y -support-yielding for set ; coherence for b1 being Function-yielding Function st b1 is X -support-yielding & b1 is Y -support-yielding holds b1 is X /\ Y -support-yielding proof let F be Function-yielding Function; ::_thesis: ( F is X -support-yielding & F is Y -support-yielding implies F is X /\ Y -support-yielding ) assume A1: ( F is X -support-yielding & F is Y -support-yielding ) ; ::_thesis: F is X /\ Y -support-yielding let f be Function; :: according to MATRTOP3:def_1 ::_thesis: for x being set st f in dom F & (F . f) . x <> f . x holds x in X /\ Y let x be set ; ::_thesis: ( f in dom F & (F . f) . x <> f . x implies x in X /\ Y ) assume ( f in dom F & (F . f) . x <> f . x ) ; ::_thesis: x in X /\ Y then ( x in X & x in Y ) by A1, Def1; hence x in X /\ Y by XBOOLE_0:def_4; ::_thesis: verum end; let f be Function-yielding X -support-yielding Function; let g be Function-yielding Y -support-yielding Function; clusterg (#) f -> X \/ Y -support-yielding ; coherence f * g is X \/ Y -support-yielding proof set fg = f * g; let h be Function; :: according to MATRTOP3:def_1 ::_thesis: for x being set st h in dom (f * g) & ((f * g) . h) . x <> h . x holds x in X \/ Y let x be set ; ::_thesis: ( h in dom (f * g) & ((f * g) . h) . x <> h . x implies x in X \/ Y ) assume that A2: h in dom (f * g) and A3: ((f * g) . h) . x <> h . x ; ::_thesis: x in X \/ Y A4: h in dom g by A2, FUNCT_1:11; assume A5: not x in X \/ Y ; ::_thesis: contradiction then not x in Y by XBOOLE_0:def_3; then A6: (g . h) . x = h . x by A4, Def1; A7: ( (f * g) . h = f . (g . h) & g . h in dom f ) by A2, FUNCT_1:11, FUNCT_1:12; not x in X by A5, XBOOLE_0:def_3; hence contradiction by A3, A6, A7, Def1; ::_thesis: verum end; end; registration let n be Nat; cluster non empty Relation-like the carrier of (TOP-REAL n) -defined the carrier of (TOP-REAL n) -valued Function-like total quasi_total homogeneous for Element of bool [: the carrier of (TOP-REAL n), the carrier of (TOP-REAL n):]; existence ex b1 being Function of (TOP-REAL n),(TOP-REAL n) st b1 is homogeneous proof reconsider N = n as Element of NAT by ORDINAL1:def_12; the homogeneous Function of (TOP-REAL N),(TOP-REAL N) is homogeneous ; hence ex b1 being Function of (TOP-REAL n),(TOP-REAL n) st b1 is homogeneous ; ::_thesis: verum end; end; registration let n, m be Nat; cluster Function-like quasi_total -> FinSequence-yielding for Element of bool [: the carrier of (TOP-REAL n), the carrier of (TOP-REAL m):]; coherence for b1 being Function of (TOP-REAL n),(TOP-REAL m) holds b1 is FinSequence-yielding proof let F be Function of (TOP-REAL n),(TOP-REAL m); ::_thesis: F is FinSequence-yielding now__::_thesis:_for_x_being_set_st_x_in_dom_F_holds_ F_._x_is_FinSequence let x be set ; ::_thesis: ( x in dom F implies F . x is FinSequence ) assume x in dom F ; ::_thesis: F . x is FinSequence then ( rng F c= the carrier of (TOP-REAL m) & F . x in rng F ) by FUNCT_1:def_3, RELAT_1:def_19; hence F . x is FinSequence ; ::_thesis: verum end; hence F is FinSequence-yielding by PRE_POLY:def_3; ::_thesis: verum end; end; registration let n, m be Nat; let A be Matrix of n,m,F_Real; cluster Mx2Tran A -> additive ; coherence Mx2Tran A is additive proof let f1, f2 be Point of (TOP-REAL n); :: according to VECTSP_1:def_20 ::_thesis: (Mx2Tran A) . (f1 + f2) = ((Mx2Tran A) . f1) + ((Mx2Tran A) . f2) thus (Mx2Tran A) . (f1 + f2) = ((Mx2Tran A) . f1) + ((Mx2Tran A) . f2) by MATRTOP1:22; ::_thesis: verum end; end; registration let n be Nat; let A be Matrix of n,F_Real; cluster Mx2Tran A -> homogeneous ; coherence Mx2Tran A is homogeneous proof let r be real number ; :: according to TOPREAL9:def_4 ::_thesis: for b1 being Element of the carrier of (TOP-REAL n) holds (Mx2Tran A) . (r * b1) = r * ((Mx2Tran A) . b1) let p1 be Point of (TOP-REAL n); ::_thesis: (Mx2Tran A) . (r * p1) = r * ((Mx2Tran A) . p1) thus (Mx2Tran A) . (r * p1) = r * ((Mx2Tran A) . p1) by MATRTOP1:23; ::_thesis: verum end; end; registration let n be Nat; let f, g be homogeneous Function of (TOP-REAL n),(TOP-REAL n); clusterg (#) f -> homogeneous for Function of (TOP-REAL n),(TOP-REAL n); coherence for b1 being Function of (TOP-REAL n),(TOP-REAL n) st b1 = f * g holds b1 is homogeneous proof set TR = TOP-REAL n; now__::_thesis:_for_r_being_real_number_ for_p_being_Point_of_(TOP-REAL_n)_holds_(f_*_g)_._(r_*_p)_=_r_*_((f_*_g)_._p) let r be real number ; ::_thesis: for p being Point of (TOP-REAL n) holds (f * g) . (r * p) = r * ((f * g) . p) let p be Point of (TOP-REAL n); ::_thesis: (f * g) . (r * p) = r * ((f * g) . p) reconsider gp = g . p as Point of (TOP-REAL n) ; A1: dom (f * g) = the carrier of (TOP-REAL n) by FUNCT_2:52; hence (f * g) . (r * p) = f . (g . (r * p)) by FUNCT_1:12 .= f . (r * gp) by TOPREAL9:def_4 .= r * (f . gp) by TOPREAL9:def_4 .= r * ((f * g) . p) by A1, FUNCT_1:12 ; ::_thesis: verum end; hence for b1 being Function of (TOP-REAL n),(TOP-REAL n) st b1 = f * g holds b1 is homogeneous by TOPREAL9:def_4; ::_thesis: verum end; end; begin Lm1: for i, n being Nat st i in Seg n holds ex M being Matrix of n,F_Real st ( Det M = - (1. F_Real) & M * (i,i) = - (1. F_Real) & ( for k, m being Nat st [k,m] in Indices M holds ( ( k = m & k <> i implies M * (k,k) = 1. F_Real ) & ( k <> m implies M * (k,m) = 0. F_Real ) ) ) ) proof let i, n be Nat; ::_thesis: ( i in Seg n implies ex M being Matrix of n,F_Real st ( Det M = - (1. F_Real) & M * (i,i) = - (1. F_Real) & ( for k, m being Nat st [k,m] in Indices M holds ( ( k = m & k <> i implies M * (k,k) = 1. F_Real ) & ( k <> m implies M * (k,m) = 0. F_Real ) ) ) ) ) set FR = the carrier of F_Real; set mm = the multF of F_Real; reconsider N = n as Element of NAT by ORDINAL1:def_12; defpred S1[ set , set , set ] means ( ( $1 = $2 & $1 = i implies $3 = - (1. F_Real) ) & ( $1 = $2 & $1 <> i implies $3 = 1. F_Real ) & ( $1 <> $2 implies $3 = 0. F_Real ) ); A1: for k, m being Nat st [k,m] in [:(Seg N),(Seg N):] holds ex x being Element of F_Real st S1[k,m,x] proof let k, m be Nat; ::_thesis: ( [k,m] in [:(Seg N),(Seg N):] implies ex x being Element of F_Real st S1[k,m,x] ) assume [k,m] in [:(Seg N),(Seg N):] ; ::_thesis: ex x being Element of F_Real st S1[k,m,x] ( ( k = m & k = i ) or ( k = m & k <> i ) or k <> m ) ; hence ex x being Element of F_Real st S1[k,m,x] ; ::_thesis: verum end; consider M being Matrix of N,F_Real such that A2: for n, m being Nat st [n,m] in Indices M holds S1[n,m,M * (n,m)] from MATRIX_1:sch_2(A1); reconsider M = M as Matrix of n,F_Real ; A3: Indices M = [:(Seg n),(Seg n):] by MATRIX_1:24; now__::_thesis:_for_k,_m_being_Element_of_NAT_st_k_in_Seg_n_&_m_in_Seg_n_&_k_<>_m_holds_ M_*_(k,m)_=_0._F_Real let k, m be Element of NAT ; ::_thesis: ( k in Seg n & m in Seg n & k <> m implies M * (k,m) = 0. F_Real ) assume ( k in Seg n & m in Seg n ) ; ::_thesis: ( k <> m implies M * (k,m) = 0. F_Real ) then A4: [k,m] in Indices M by A3, ZFMISC_1:87; assume k <> m ; ::_thesis: M * (k,m) = 0. F_Real hence M * (k,m) = 0. F_Real by A2, A4; ::_thesis: verum end; then A5: M is V244( F_Real ) by MATRIX_7:def_2; set D = diagonal_of_Matrix M; defpred S2[ Nat] means ( $1 + i <= n implies the multF of F_Real "**" ((diagonal_of_Matrix M) | ($1 + i)) = - (1. F_Real) ); A6: len (diagonal_of_Matrix M) = n by MATRIX_3:def_10; A7: for k being Nat st S2[k] holds S2[k + 1] proof let k be Nat; ::_thesis: ( S2[k] implies S2[k + 1] ) assume A8: S2[k] ; ::_thesis: S2[k + 1] set k1 = k + 1; set ki = (k + 1) + i; assume A9: (k + 1) + i <= n ; ::_thesis: the multF of F_Real "**" ((diagonal_of_Matrix M) | ((k + 1) + i)) = - (1. F_Real) A10: (k + 1) + i = (k + i) + 1 ; then A11: 1 <= (k + 1) + i by NAT_1:11; then (k + 1) + i in dom (diagonal_of_Matrix M) by A6, A9, FINSEQ_3:25; then A12: (diagonal_of_Matrix M) | ((k + 1) + i) = ((diagonal_of_Matrix M) | (k + i)) ^ <*((diagonal_of_Matrix M) . ((k + 1) + i))*> by A10, FINSEQ_5:10; i <= k + i by NAT_1:11; then A13: i < (k + 1) + i by A10, NAT_1:13; A14: (k + 1) + i in Seg n by A9, A11; then [((k + 1) + i),((k + 1) + i)] in Indices M by A3, ZFMISC_1:87; then 1. F_Real = M * (((k + 1) + i),((k + 1) + i)) by A2, A13 .= (diagonal_of_Matrix M) . ((k + 1) + i) by A14, MATRIX_3:def_10 ; hence the multF of F_Real "**" ((diagonal_of_Matrix M) | ((k + 1) + i)) = (- (1. F_Real)) * (1. F_Real) by A8, A9, A10, A12, FINSOP_1:4, NAT_1:13 .= - (1. F_Real) ; ::_thesis: verum end; defpred S3[ Nat] means ( $1 < i implies the multF of F_Real "**" ((diagonal_of_Matrix M) | $1) = 1. F_Real ); A15: S3[ 0 ] proof assume 0 < i ; ::_thesis: the multF of F_Real "**" ((diagonal_of_Matrix M) | 0) = 1. F_Real ( (diagonal_of_Matrix M) | 0 = <*> the carrier of F_Real & the_unity_wrt the multF of F_Real = 1. F_Real ) by FVSUM_1:5; hence the multF of F_Real "**" ((diagonal_of_Matrix M) | 0) = 1. F_Real by FINSOP_1:10; ::_thesis: verum end; assume A16: i in Seg n ; ::_thesis: ex M being Matrix of n,F_Real st ( Det M = - (1. F_Real) & M * (i,i) = - (1. F_Real) & ( for k, m being Nat st [k,m] in Indices M holds ( ( k = m & k <> i implies M * (k,k) = 1. F_Real ) & ( k <> m implies M * (k,m) = 0. F_Real ) ) ) ) then A17: 1 <= i by FINSEQ_1:1; A18: i <= n by A16, FINSEQ_1:1; then A19: ( (n - i) + i = n & n - i is Nat ) by NAT_1:21; take M ; ::_thesis: ( Det M = - (1. F_Real) & M * (i,i) = - (1. F_Real) & ( for k, m being Nat st [k,m] in Indices M holds ( ( k = m & k <> i implies M * (k,k) = 1. F_Real ) & ( k <> m implies M * (k,m) = 0. F_Real ) ) ) ) A20: for k being Nat st S3[k] holds S3[k + 1] proof let k be Nat; ::_thesis: ( S3[k] implies S3[k + 1] ) assume A21: S3[k] ; ::_thesis: S3[k + 1] set k1 = k + 1; assume A22: k + 1 < i ; ::_thesis: the multF of F_Real "**" ((diagonal_of_Matrix M) | (k + 1)) = 1. F_Real then A23: ( 1 <= k + 1 & k + 1 <= n ) by A18, NAT_1:14, XXREAL_0:2; then k + 1 in dom (diagonal_of_Matrix M) by A6, FINSEQ_3:25; then A24: (diagonal_of_Matrix M) | (k + 1) = ((diagonal_of_Matrix M) | k) ^ <*((diagonal_of_Matrix M) . (k + 1))*> by FINSEQ_5:10; A25: k + 1 in Seg n by A23; then [(k + 1),(k + 1)] in Indices M by A3, ZFMISC_1:87; then 1. F_Real = M * ((k + 1),(k + 1)) by A2, A22 .= (diagonal_of_Matrix M) . (k + 1) by A25, MATRIX_3:def_10 ; hence the multF of F_Real "**" ((diagonal_of_Matrix M) | (k + 1)) = (1. F_Real) * (1. F_Real) by A21, A22, A24, FINSOP_1:4, NAT_1:13 .= 1. F_Real ; ::_thesis: verum end; A26: for k being Nat holds S3[k] from NAT_1:sch_2(A15, A20); A27: S2[ 0 ] proof reconsider I = i - 1 as Nat by A17, NAT_1:21; assume 0 + i <= n ; ::_thesis: the multF of F_Real "**" ((diagonal_of_Matrix M) | (0 + i)) = - (1. F_Real) A28: I + 1 = i ; then I < i by NAT_1:13; then A29: the multF of F_Real "**" ((diagonal_of_Matrix M) | I) = 1. F_Real by A26; 1 <= i by A16, FINSEQ_1:1; then i in dom (diagonal_of_Matrix M) by A6, A18, FINSEQ_3:25; then A30: (diagonal_of_Matrix M) | i = ((diagonal_of_Matrix M) | I) ^ <*((diagonal_of_Matrix M) . i)*> by A28, FINSEQ_5:10; [i,i] in Indices M by A3, A16, ZFMISC_1:87; then - (1. F_Real) = M * (i,i) by A2 .= (diagonal_of_Matrix M) . i by A16, MATRIX_3:def_10 ; hence the multF of F_Real "**" ((diagonal_of_Matrix M) | (0 + i)) = (1. F_Real) * (- (1. F_Real)) by A29, A30, FINSOP_1:4 .= - (1. F_Real) ; ::_thesis: verum end; for k being Nat holds S2[k] from NAT_1:sch_2(A27, A7); hence - (1. F_Real) = the multF of F_Real "**" ((diagonal_of_Matrix M) | n) by A19 .= the multF of F_Real "**" (diagonal_of_Matrix M) by A6, FINSEQ_1:58 .= Det M by A5, A17, A18, MATRIX_7:17, XXREAL_0:2 ; ::_thesis: ( M * (i,i) = - (1. F_Real) & ( for k, m being Nat st [k,m] in Indices M holds ( ( k = m & k <> i implies M * (k,k) = 1. F_Real ) & ( k <> m implies M * (k,m) = 0. F_Real ) ) ) ) [i,i] in Indices M by A3, A16, ZFMISC_1:87; hence M * (i,i) = - (1. F_Real) by A2; ::_thesis: for k, m being Nat st [k,m] in Indices M holds ( ( k = m & k <> i implies M * (k,k) = 1. F_Real ) & ( k <> m implies M * (k,m) = 0. F_Real ) ) let k, m be Nat; ::_thesis: ( [k,m] in Indices M implies ( ( k = m & k <> i implies M * (k,k) = 1. F_Real ) & ( k <> m implies M * (k,m) = 0. F_Real ) ) ) assume [k,m] in Indices M ; ::_thesis: ( ( k = m & k <> i implies M * (k,k) = 1. F_Real ) & ( k <> m implies M * (k,m) = 0. F_Real ) ) hence ( ( k = m & k <> i implies M * (k,k) = 1. F_Real ) & ( k <> m implies M * (k,m) = 0. F_Real ) ) by A2; ::_thesis: verum end; definition let n, i be Nat; assume B1: i in Seg n ; func AxialSymmetry (i,n) -> invertible Matrix of n,F_Real means :Def2: :: MATRTOP3:def 2 ( it * (i,i) = - (1. F_Real) & ( for k, m being Nat st [k,m] in Indices it holds ( ( k = m & k <> i implies it * (k,k) = 1. F_Real ) & ( k <> m implies it * (k,m) = 0. F_Real ) ) ) ); existence ex b1 being invertible Matrix of n,F_Real st ( b1 * (i,i) = - (1. F_Real) & ( for k, m being Nat st [k,m] in Indices b1 holds ( ( k = m & k <> i implies b1 * (k,k) = 1. F_Real ) & ( k <> m implies b1 * (k,m) = 0. F_Real ) ) ) ) proof consider M being Matrix of n,F_Real such that A1: Det M = - (1. F_Real) and A2: ( M * (i,i) = - (1. F_Real) & ( for k, m being Nat st [k,m] in Indices M holds ( ( k = m & k <> i implies M * (k,k) = 1. F_Real ) & ( k <> m implies M * (k,m) = 0. F_Real ) ) ) ) by B1, Lm1; Det M <> 0. F_Real by A1; then M is invertible by LAPLACE:34; hence ex b1 being invertible Matrix of n,F_Real st ( b1 * (i,i) = - (1. F_Real) & ( for k, m being Nat st [k,m] in Indices b1 holds ( ( k = m & k <> i implies b1 * (k,k) = 1. F_Real ) & ( k <> m implies b1 * (k,m) = 0. F_Real ) ) ) ) by A2; ::_thesis: verum end; uniqueness for b1, b2 being invertible Matrix of n,F_Real st b1 * (i,i) = - (1. F_Real) & ( for k, m being Nat st [k,m] in Indices b1 holds ( ( k = m & k <> i implies b1 * (k,k) = 1. F_Real ) & ( k <> m implies b1 * (k,m) = 0. F_Real ) ) ) & b2 * (i,i) = - (1. F_Real) & ( for k, m being Nat st [k,m] in Indices b2 holds ( ( k = m & k <> i implies b2 * (k,k) = 1. F_Real ) & ( k <> m implies b2 * (k,m) = 0. F_Real ) ) ) holds b1 = b2 proof let A1, A2 be invertible Matrix of n,F_Real; ::_thesis: ( A1 * (i,i) = - (1. F_Real) & ( for k, m being Nat st [k,m] in Indices A1 holds ( ( k = m & k <> i implies A1 * (k,k) = 1. F_Real ) & ( k <> m implies A1 * (k,m) = 0. F_Real ) ) ) & A2 * (i,i) = - (1. F_Real) & ( for k, m being Nat st [k,m] in Indices A2 holds ( ( k = m & k <> i implies A2 * (k,k) = 1. F_Real ) & ( k <> m implies A2 * (k,m) = 0. F_Real ) ) ) implies A1 = A2 ) assume that A3: A1 * (i,i) = - (1. F_Real) and A4: for k, m being Nat st [k,m] in Indices A1 holds ( ( k = m & k <> i implies A1 * (k,k) = 1. F_Real ) & ( k <> m implies A1 * (k,m) = 0. F_Real ) ) and A5: A2 * (i,i) = - (1. F_Real) and A6: for k, m being Nat st [k,m] in Indices A2 holds ( ( k = m & k <> i implies A2 * (k,k) = 1. F_Real ) & ( k <> m implies A2 * (k,m) = 0. F_Real ) ) ; ::_thesis: A1 = A2 for k, m being Nat st [k,m] in Indices A1 holds A1 * (k,m) = A2 * (k,m) proof let k, m be Nat; ::_thesis: ( [k,m] in Indices A1 implies A1 * (k,m) = A2 * (k,m) ) assume A7: [k,m] in Indices A1 ; ::_thesis: A1 * (k,m) = A2 * (k,m) then A8: [k,m] in Indices A2 by MATRIX_1:26; percases ( k <> m or ( k = m & k <> i ) or ( k = m & k = i ) ) ; supposeA9: k <> m ; ::_thesis: A1 * (k,m) = A2 * (k,m) then A1 * (k,m) = 0. F_Real by A4, A7; hence A1 * (k,m) = A2 * (k,m) by A6, A8, A9; ::_thesis: verum end; supposeA10: ( k = m & k <> i ) ; ::_thesis: A1 * (k,m) = A2 * (k,m) then A1 * (k,m) = 1. F_Real by A4, A7; hence A1 * (k,m) = A2 * (k,m) by A6, A8, A10; ::_thesis: verum end; suppose ( k = m & k = i ) ; ::_thesis: A1 * (k,m) = A2 * (k,m) hence A1 * (k,m) = A2 * (k,m) by A3, A5; ::_thesis: verum end; end; end; hence A1 = A2 by MATRIX_1:27; ::_thesis: verum end; end; :: deftheorem Def2 defines AxialSymmetry MATRTOP3:def_2_:_ for n, i being Nat st i in Seg n holds for b3 being invertible Matrix of n,F_Real holds ( b3 = AxialSymmetry (i,n) iff ( b3 * (i,i) = - (1. F_Real) & ( for k, m being Nat st [k,m] in Indices b3 holds ( ( k = m & k <> i implies b3 * (k,k) = 1. F_Real ) & ( k <> m implies b3 * (k,m) = 0. F_Real ) ) ) ) ); theorem Th4: :: MATRTOP3:4 for i, n being Nat st i in Seg n holds Det (AxialSymmetry (i,n)) = - (1. F_Real) proof let i, n be Nat; ::_thesis: ( i in Seg n implies Det (AxialSymmetry (i,n)) = - (1. F_Real) ) assume A1: i in Seg n ; ::_thesis: Det (AxialSymmetry (i,n)) = - (1. F_Real) then consider M being Matrix of n,F_Real such that A2: Det M = - (1. F_Real) and A3: ( M * (i,i) = - (1. F_Real) & ( for k, m being Nat st [k,m] in Indices M holds ( ( k = m & k <> i implies M * (k,k) = 1. F_Real ) & ( k <> m implies M * (k,m) = 0. F_Real ) ) ) ) by Lm1; Det M <> 0. F_Real by A2; then M is invertible by LAPLACE:34; hence Det (AxialSymmetry (i,n)) = - (1. F_Real) by A1, A2, A3, Def2; ::_thesis: verum end; theorem Th5: :: MATRTOP3:5 for i, n, j being Nat for p being Point of (TOP-REAL n) st i in Seg n & j in Seg n & i <> j holds (@ p) "*" (Col ((AxialSymmetry (i,n)),j)) = p . j proof let i, n, j be Nat; ::_thesis: for p being Point of (TOP-REAL n) st i in Seg n & j in Seg n & i <> j holds (@ p) "*" (Col ((AxialSymmetry (i,n)),j)) = p . j let p be Point of (TOP-REAL n); ::_thesis: ( i in Seg n & j in Seg n & i <> j implies (@ p) "*" (Col ((AxialSymmetry (i,n)),j)) = p . j ) set S = Seg n; assume that A1: i in Seg n and A2: j in Seg n and A3: i <> j ; ::_thesis: (@ p) "*" (Col ((AxialSymmetry (i,n)),j)) = p . j set A = AxialSymmetry (i,n); set C = Col ((AxialSymmetry (i,n)),j); A4: Indices (AxialSymmetry (i,n)) = [:(Seg n),(Seg n):] by MATRIX_1:24; then A5: [j,j] in Indices (AxialSymmetry (i,n)) by A2, ZFMISC_1:87; A6: len (AxialSymmetry (i,n)) = n by MATRIX_1:25; then A7: dom (AxialSymmetry (i,n)) = Seg n by FINSEQ_1:def_3; len (Col ((AxialSymmetry (i,n)),j)) = n by A6, MATRIX_1:def_8; then A8: dom (Col ((AxialSymmetry (i,n)),j)) = Seg n by FINSEQ_1:def_3; A9: now__::_thesis:_for_m_being_Nat_st_m_in_dom_(Col_((AxialSymmetry_(i,n)),j))_&_m_<>_j_holds_ (Col_((AxialSymmetry_(i,n)),j))_._m_=_0._F_Real let m be Nat; ::_thesis: ( m in dom (Col ((AxialSymmetry (i,n)),j)) & m <> j implies (Col ((AxialSymmetry (i,n)),j)) . m = 0. F_Real ) assume that A10: m in dom (Col ((AxialSymmetry (i,n)),j)) and A11: m <> j ; ::_thesis: (Col ((AxialSymmetry (i,n)),j)) . m = 0. F_Real A12: [m,j] in Indices (AxialSymmetry (i,n)) by A2, A4, A8, A10, ZFMISC_1:87; thus (Col ((AxialSymmetry (i,n)),j)) . m = (AxialSymmetry (i,n)) * (m,j) by A7, A8, A10, MATRIX_1:def_8 .= 0. F_Real by A1, A11, A12, Def2 ; ::_thesis: verum end; len p = n by CARD_1:def_7; then A13: dom p = Seg n by FINSEQ_1:def_3; (Col ((AxialSymmetry (i,n)),j)) . j = (AxialSymmetry (i,n)) * (j,j) by A2, A7, MATRIX_1:def_8 .= 1. F_Real by A1, A3, A5, Def2 ; hence p . j = Sum (mlt ((Col ((AxialSymmetry (i,n)),j)),(@ p))) by A2, A8, A9, A13, MATRIX_3:17 .= (@ p) "*" (Col ((AxialSymmetry (i,n)),j)) by FVSUM_1:64 ; ::_thesis: verum end; theorem Th6: :: MATRTOP3:6 for i, n being Nat for p being Point of (TOP-REAL n) st i in Seg n holds (@ p) "*" (Col ((AxialSymmetry (i,n)),i)) = - (p . i) proof let i, n be Nat; ::_thesis: for p being Point of (TOP-REAL n) st i in Seg n holds (@ p) "*" (Col ((AxialSymmetry (i,n)),i)) = - (p . i) let p be Point of (TOP-REAL n); ::_thesis: ( i in Seg n implies (@ p) "*" (Col ((AxialSymmetry (i,n)),i)) = - (p . i) ) set S = Seg n; assume A1: i in Seg n ; ::_thesis: (@ p) "*" (Col ((AxialSymmetry (i,n)),i)) = - (p . i) reconsider pI = (@ p) . i as Element of F_Real ; set A = AxialSymmetry (i,n); set C = Col ((AxialSymmetry (i,n)),i); A2: len (AxialSymmetry (i,n)) = n by MATRIX_1:25; then A3: dom (AxialSymmetry (i,n)) = Seg n by FINSEQ_1:def_3; then A4: (Col ((AxialSymmetry (i,n)),i)) . i = (AxialSymmetry (i,n)) * (i,i) by A1, MATRIX_1:def_8; ( len p = n & len (Col ((AxialSymmetry (i,n)),i)) = n ) by A2, CARD_1:def_7; then len (mlt ((@ p),(Col ((AxialSymmetry (i,n)),i)))) = n by MATRIX_3:6; then A5: dom (mlt ((@ p),(Col ((AxialSymmetry (i,n)),i)))) = Seg n by FINSEQ_1:def_3; A6: Indices (AxialSymmetry (i,n)) = [:(Seg n),(Seg n):] by MATRIX_1:24; A7: for k being Nat st k in dom (mlt ((@ p),(Col ((AxialSymmetry (i,n)),i)))) & k <> i holds (mlt ((@ p),(Col ((AxialSymmetry (i,n)),i)))) . k = 0. F_Real proof let k be Nat; ::_thesis: ( k in dom (mlt ((@ p),(Col ((AxialSymmetry (i,n)),i)))) & k <> i implies (mlt ((@ p),(Col ((AxialSymmetry (i,n)),i)))) . k = 0. F_Real ) assume that A8: k in dom (mlt ((@ p),(Col ((AxialSymmetry (i,n)),i)))) and A9: k <> i ; ::_thesis: (mlt ((@ p),(Col ((AxialSymmetry (i,n)),i)))) . k = 0. F_Real reconsider pk = (@ p) . k as Element of F_Real ; A10: [k,i] in Indices (AxialSymmetry (i,n)) by A1, A5, A6, A8, ZFMISC_1:87; (Col ((AxialSymmetry (i,n)),i)) . k = (AxialSymmetry (i,n)) * (k,i) by A3, A5, A8, MATRIX_1:def_8; hence (mlt ((@ p),(Col ((AxialSymmetry (i,n)),i)))) . k = pk * ((AxialSymmetry (i,n)) * (k,i)) by A8, FVSUM_1:60 .= pk * (0. F_Real) by A1, A9, A10, Def2 .= 0. F_Real ; ::_thesis: verum end; thus (@ p) "*" (Col ((AxialSymmetry (i,n)),i)) = (mlt ((@ p),(Col ((AxialSymmetry (i,n)),i)))) . i by A1, A5, A7, MATRIX_3:12 .= pI * ((AxialSymmetry (i,n)) * (i,i)) by A1, A4, A5, FVSUM_1:60 .= pI * (- (1. F_Real)) by A1, Def2 .= - (p . i) ; ::_thesis: verum end; theorem Th7: :: MATRTOP3:7 for i, n being Nat st i in Seg n holds ( AxialSymmetry (i,n) is V244( F_Real ) & (AxialSymmetry (i,n)) ~ = AxialSymmetry (i,n) ) proof let i, n be Nat; ::_thesis: ( i in Seg n implies ( AxialSymmetry (i,n) is V244( F_Real ) & (AxialSymmetry (i,n)) ~ = AxialSymmetry (i,n) ) ) set S = Seg n; set A = AxialSymmetry (i,n); set ONE = 1. (F_Real,n); set AA = (AxialSymmetry (i,n)) * (AxialSymmetry (i,n)); assume A1: i in Seg n ; ::_thesis: ( AxialSymmetry (i,n) is V244( F_Real ) & (AxialSymmetry (i,n)) ~ = AxialSymmetry (i,n) ) for k, m being Nat st [k,m] in Indices (AxialSymmetry (i,n)) & (AxialSymmetry (i,n)) * (k,m) <> 0. F_Real holds k = m by A1, Def2; hence AxialSymmetry (i,n) is V244( F_Real ) by MATRIX_1:def_14; ::_thesis: (AxialSymmetry (i,n)) ~ = AxialSymmetry (i,n) for k, m being Nat st [k,m] in Indices ((AxialSymmetry (i,n)) * (AxialSymmetry (i,n))) holds ((AxialSymmetry (i,n)) * (AxialSymmetry (i,n))) * (k,m) = (1. (F_Real,n)) * (k,m) proof let k, m be Nat; ::_thesis: ( [k,m] in Indices ((AxialSymmetry (i,n)) * (AxialSymmetry (i,n))) implies ((AxialSymmetry (i,n)) * (AxialSymmetry (i,n))) * (k,m) = (1. (F_Real,n)) * (k,m) ) assume A2: [k,m] in Indices ((AxialSymmetry (i,n)) * (AxialSymmetry (i,n))) ; ::_thesis: ((AxialSymmetry (i,n)) * (AxialSymmetry (i,n))) * (k,m) = (1. (F_Real,n)) * (k,m) A3: width (AxialSymmetry (i,n)) = n by MATRIX_1:24; then len (@ (Line ((AxialSymmetry (i,n)),k))) = n by CARD_1:def_7; then reconsider L = @ (Line ((AxialSymmetry (i,n)),k)) as Element of (TOP-REAL n) by TOPREAL3:46; len (AxialSymmetry (i,n)) = n by MATRIX_1:25; then A4: ((AxialSymmetry (i,n)) * (AxialSymmetry (i,n))) * (k,m) = (@ L) "*" (Col ((AxialSymmetry (i,n)),m)) by A2, A3, MATRIX_3:def_4; A5: Indices ((AxialSymmetry (i,n)) * (AxialSymmetry (i,n))) = [:(Seg n),(Seg n):] by MATRIX_1:24; then A6: m in Seg n by A2, ZFMISC_1:87; then A7: (Line ((AxialSymmetry (i,n)),k)) . m = (AxialSymmetry (i,n)) * (k,m) by A3, MATRIX_1:def_7; A8: Indices (AxialSymmetry (i,n)) = [:(Seg n),(Seg n):] by MATRIX_1:24; A9: Indices (1. (F_Real,n)) = [:(Seg n),(Seg n):] by MATRIX_1:24; percases ( m <> i or m = i ) ; supposeA10: m <> i ; ::_thesis: ((AxialSymmetry (i,n)) * (AxialSymmetry (i,n))) * (k,m) = (1. (F_Real,n)) * (k,m) then A11: ((AxialSymmetry (i,n)) * (AxialSymmetry (i,n))) * (k,m) = (AxialSymmetry (i,n)) * (k,m) by A1, A4, A6, A7, Th5; percases ( k <> m or k = m ) ; supposeA12: k <> m ; ::_thesis: ((AxialSymmetry (i,n)) * (AxialSymmetry (i,n))) * (k,m) = (1. (F_Real,n)) * (k,m) then (1. (F_Real,n)) * (k,m) = 0. F_Real by A2, A5, A9, MATRIX_1:def_11; hence ((AxialSymmetry (i,n)) * (AxialSymmetry (i,n))) * (k,m) = (1. (F_Real,n)) * (k,m) by A1, A2, A5, A8, A11, A12, Def2; ::_thesis: verum end; supposeA13: k = m ; ::_thesis: ((AxialSymmetry (i,n)) * (AxialSymmetry (i,n))) * (k,m) = (1. (F_Real,n)) * (k,m) then (1. (F_Real,n)) * (k,m) = 1. F_Real by A2, A5, A9, MATRIX_1:def_11; hence ((AxialSymmetry (i,n)) * (AxialSymmetry (i,n))) * (k,m) = (1. (F_Real,n)) * (k,m) by A1, A2, A5, A8, A10, A11, A13, Def2; ::_thesis: verum end; end; end; supposeA14: m = i ; ::_thesis: ((AxialSymmetry (i,n)) * (AxialSymmetry (i,n))) * (k,m) = (1. (F_Real,n)) * (k,m) then A15: ((AxialSymmetry (i,n)) * (AxialSymmetry (i,n))) * (k,m) = - ((AxialSymmetry (i,n)) * (k,m)) by A4, A6, A7, Th6; percases ( k <> m or k = m ) ; supposeA16: k <> m ; ::_thesis: ((AxialSymmetry (i,n)) * (AxialSymmetry (i,n))) * (k,m) = (1. (F_Real,n)) * (k,m) then (AxialSymmetry (i,n)) * (k,m) = 0. F_Real by A1, A2, A5, A8, Def2; hence ((AxialSymmetry (i,n)) * (AxialSymmetry (i,n))) * (k,m) = (1. (F_Real,n)) * (k,m) by A2, A5, A9, A15, A16, MATRIX_1:def_11; ::_thesis: verum end; supposeA17: k = m ; ::_thesis: ((AxialSymmetry (i,n)) * (AxialSymmetry (i,n))) * (k,m) = (1. (F_Real,n)) * (k,m) then ((AxialSymmetry (i,n)) * (AxialSymmetry (i,n))) * (k,m) = - (- (1. F_Real)) by A1, A14, A15, Def2; hence ((AxialSymmetry (i,n)) * (AxialSymmetry (i,n))) * (k,m) = (1. (F_Real,n)) * (k,m) by A2, A5, A9, A17, MATRIX_1:def_11; ::_thesis: verum end; end; end; end; end; then (AxialSymmetry (i,n)) * (AxialSymmetry (i,n)) = 1. (F_Real,n) by MATRIX_1:27; then AxialSymmetry (i,n) is_reverse_of AxialSymmetry (i,n) by MATRIX_6:def_2; hence (AxialSymmetry (i,n)) ~ = AxialSymmetry (i,n) by MATRIX_6:def_4; ::_thesis: verum end; theorem Th8: :: MATRTOP3:8 for i, n, j being Nat for p being Point of (TOP-REAL n) st i in Seg n & i <> j holds ((Mx2Tran (AxialSymmetry (i,n))) . p) . j = p . j proof let i, n, j be Nat; ::_thesis: for p being Point of (TOP-REAL n) st i in Seg n & i <> j holds ((Mx2Tran (AxialSymmetry (i,n))) . p) . j = p . j let p be Point of (TOP-REAL n); ::_thesis: ( i in Seg n & i <> j implies ((Mx2Tran (AxialSymmetry (i,n))) . p) . j = p . j ) set A = AxialSymmetry (i,n); set M = Mx2Tran (AxialSymmetry (i,n)); set Mp = (Mx2Tran (AxialSymmetry (i,n))) . p; set S = Seg n; assume A1: ( i in Seg n & i <> j ) ; ::_thesis: ((Mx2Tran (AxialSymmetry (i,n))) . p) . j = p . j len ((Mx2Tran (AxialSymmetry (i,n))) . p) = n by CARD_1:def_7; then A2: dom ((Mx2Tran (AxialSymmetry (i,n))) . p) = Seg n by FINSEQ_1:def_3; len p = n by CARD_1:def_7; then A3: dom p = Seg n by FINSEQ_1:def_3; percases ( j in Seg n or not j in Seg n ) ; supposeA4: j in Seg n ; ::_thesis: ((Mx2Tran (AxialSymmetry (i,n))) . p) . j = p . j then ( 1 <= j & j <= n ) by FINSEQ_1:1; hence ((Mx2Tran (AxialSymmetry (i,n))) . p) . j = (@ p) "*" (Col ((AxialSymmetry (i,n)),j)) by MATRTOP1:18 .= p . j by A1, A4, Th5 ; ::_thesis: verum end; supposeA5: not j in Seg n ; ::_thesis: ((Mx2Tran (AxialSymmetry (i,n))) . p) . j = p . j then ((Mx2Tran (AxialSymmetry (i,n))) . p) . j = {} by A2, FUNCT_1:def_2; hence ((Mx2Tran (AxialSymmetry (i,n))) . p) . j = p . j by A3, A5, FUNCT_1:def_2; ::_thesis: verum end; end; end; theorem Th9: :: MATRTOP3:9 for i, n being Nat for p being Point of (TOP-REAL n) st i in Seg n holds ((Mx2Tran (AxialSymmetry (i,n))) . p) . i = - (p . i) proof let i, n be Nat; ::_thesis: for p being Point of (TOP-REAL n) st i in Seg n holds ((Mx2Tran (AxialSymmetry (i,n))) . p) . i = - (p . i) let p be Point of (TOP-REAL n); ::_thesis: ( i in Seg n implies ((Mx2Tran (AxialSymmetry (i,n))) . p) . i = - (p . i) ) set A = AxialSymmetry (i,n); set M = Mx2Tran (AxialSymmetry (i,n)); set Mp = (Mx2Tran (AxialSymmetry (i,n))) . p; set S = Seg n; assume A1: i in Seg n ; ::_thesis: ((Mx2Tran (AxialSymmetry (i,n))) . p) . i = - (p . i) then ( 1 <= i & i <= n ) by FINSEQ_1:1; hence ((Mx2Tran (AxialSymmetry (i,n))) . p) . i = (@ p) "*" (Col ((AxialSymmetry (i,n)),i)) by MATRTOP1:18 .= - (p . i) by A1, Th6 ; ::_thesis: verum end; theorem Th10: :: MATRTOP3:10 for i, n being Nat for p being Point of (TOP-REAL n) st i in Seg n holds (Mx2Tran (AxialSymmetry (i,n))) . p = p +* (i,(- (p . i))) proof let i, n be Nat; ::_thesis: for p being Point of (TOP-REAL n) st i in Seg n holds (Mx2Tran (AxialSymmetry (i,n))) . p = p +* (i,(- (p . i))) let p be Point of (TOP-REAL n); ::_thesis: ( i in Seg n implies (Mx2Tran (AxialSymmetry (i,n))) . p = p +* (i,(- (p . i))) ) set S = Seg n; set Mp = (Mx2Tran (AxialSymmetry (i,n))) . p; set p0 = p +* (i,(- (p . i))); A1: len p = n by CARD_1:def_7; assume A2: i in Seg n ; ::_thesis: (Mx2Tran (AxialSymmetry (i,n))) . p = p +* (i,(- (p . i))) A3: for j being Nat st 1 <= j & j <= n holds ((Mx2Tran (AxialSymmetry (i,n))) . p) . j = (p +* (i,(- (p . i)))) . j proof let j be Nat; ::_thesis: ( 1 <= j & j <= n implies ((Mx2Tran (AxialSymmetry (i,n))) . p) . j = (p +* (i,(- (p . i)))) . j ) assume A4: ( 1 <= j & j <= n ) ; ::_thesis: ((Mx2Tran (AxialSymmetry (i,n))) . p) . j = (p +* (i,(- (p . i)))) . j A5: j in Seg n by A4, FINSEQ_1:1; A6: j in dom p by A1, A4, FINSEQ_3:25; percases ( j <> i or j = i ) ; supposeA7: j <> i ; ::_thesis: ((Mx2Tran (AxialSymmetry (i,n))) . p) . j = (p +* (i,(- (p . i)))) . j then (p +* (i,(- (p . i)))) . j = p . j by FUNCT_7:32; hence ((Mx2Tran (AxialSymmetry (i,n))) . p) . j = (p +* (i,(- (p . i)))) . j by A2, A7, Th8; ::_thesis: verum end; supposeA8: j = i ; ::_thesis: ((Mx2Tran (AxialSymmetry (i,n))) . p) . j = (p +* (i,(- (p . i)))) . j then (p +* (i,(- (p . i)))) . j = - (p . i) by A6, FUNCT_7:31; hence ((Mx2Tran (AxialSymmetry (i,n))) . p) . j = (p +* (i,(- (p . i)))) . j by A5, A8, Th9; ::_thesis: verum end; end; end; ( len (p +* (i,(- (p . i)))) = len p & len ((Mx2Tran (AxialSymmetry (i,n))) . p) = n ) by CARD_1:def_7, FUNCT_7:97; hence (Mx2Tran (AxialSymmetry (i,n))) . p = p +* (i,(- (p . i))) by A1, A3, FINSEQ_1:14; ::_thesis: verum end; theorem Th11: :: MATRTOP3:11 for i, n being Nat st i in Seg n holds Mx2Tran (AxialSymmetry (i,n)) is {i} -support-yielding proof let i, n be Nat; ::_thesis: ( i in Seg n implies Mx2Tran (AxialSymmetry (i,n)) is {i} -support-yielding ) set M = Mx2Tran (AxialSymmetry (i,n)); assume A1: i in Seg n ; ::_thesis: Mx2Tran (AxialSymmetry (i,n)) is {i} -support-yielding let f be Function; :: according to MATRTOP3:def_1 ::_thesis: for x being set st f in dom (Mx2Tran (AxialSymmetry (i,n))) & ((Mx2Tran (AxialSymmetry (i,n))) . f) . x <> f . x holds x in {i} let x be set ; ::_thesis: ( f in dom (Mx2Tran (AxialSymmetry (i,n))) & ((Mx2Tran (AxialSymmetry (i,n))) . f) . x <> f . x implies x in {i} ) assume f in dom (Mx2Tran (AxialSymmetry (i,n))) ; ::_thesis: ( not ((Mx2Tran (AxialSymmetry (i,n))) . f) . x <> f . x or x in {i} ) then reconsider F = f as Point of (TOP-REAL n) by FUNCT_2:52; assume A2: ((Mx2Tran (AxialSymmetry (i,n))) . f) . x <> f . x ; ::_thesis: x in {i} len ((Mx2Tran (AxialSymmetry (i,n))) . F) = n by CARD_1:def_7; then A3: dom ((Mx2Tran (AxialSymmetry (i,n))) . F) = Seg n by FINSEQ_1:def_3; A4: len F = n by CARD_1:def_7; then A5: dom F = Seg n by FINSEQ_1:def_3; percases ( not x in Seg n or x in Seg n ) ; supposeA6: not x in Seg n ; ::_thesis: x in {i} then ((Mx2Tran (AxialSymmetry (i,n))) . F) . x = {} by A3, FUNCT_1:def_2 .= F . x by A5, A6, FUNCT_1:def_2 ; hence x in {i} by A2; ::_thesis: verum end; suppose x in Seg n ; ::_thesis: x in {i} then x = i by A1, A2, A4, Th8; hence x in {i} by TARSKI:def_1; ::_thesis: verum end; end; end; theorem Th12: :: MATRTOP3:12 for r being real number for n being Nat for a, b being Element of F_Real st a = cos r & b = sin r holds Det (block_diagonal (<*((a,b) ][ ((- b),a)),(1. (F_Real,n))*>,(0. F_Real))) = 1. F_Real proof let r be real number ; ::_thesis: for n being Nat for a, b being Element of F_Real st a = cos r & b = sin r holds Det (block_diagonal (<*((a,b) ][ ((- b),a)),(1. (F_Real,n))*>,(0. F_Real))) = 1. F_Real let n be Nat; ::_thesis: for a, b being Element of F_Real st a = cos r & b = sin r holds Det (block_diagonal (<*((a,b) ][ ((- b),a)),(1. (F_Real,n))*>,(0. F_Real))) = 1. F_Real let a, b be Element of F_Real; ::_thesis: ( a = cos r & b = sin r implies Det (block_diagonal (<*((a,b) ][ ((- b),a)),(1. (F_Real,n))*>,(0. F_Real))) = 1. F_Real ) set A = (a,b) ][ ((- b),a); set ONE = 1. (F_Real,n); set B = block_diagonal (<*((a,b) ][ ((- b),a)),(1. (F_Real,n))*>,(0. F_Real)); A1: ( n = 0 or n >= 1 ) by NAT_1:14; n in NAT by ORDINAL1:def_12; then A2: ( Det (1. (F_Real,n)) = 1_ F_Real or Det (1. (F_Real,n)) = 1. F_Real ) by A1, MATRIXR2:41, MATRIX_7:16; assume ( a = cos r & b = sin r ) ; ::_thesis: Det (block_diagonal (<*((a,b) ][ ((- b),a)),(1. (F_Real,n))*>,(0. F_Real))) = 1. F_Real then A3: ((cos r) * (cos r)) + ((sin r) * (sin r)) = (a * a) - (b * (- b)) .= Det ((a,b) ][ ((- b),a)) by MATRIX_9:13 ; A4: ( cos r = cos . r & sin r = sin . r ) by SIN_COS:def_17, SIN_COS:def_19; thus Det (block_diagonal (<*((a,b) ][ ((- b),a)),(1. (F_Real,n))*>,(0. F_Real))) = (Det ((a,b) ][ ((- b),a))) * (Det (1. (F_Real,n))) by MATRIXJ1:52 .= 1. F_Real by A2, A3, A4, SIN_COS:28 ; ::_thesis: verum end; Lm2: for i, j, n being Nat st 1 <= i & i < j & j <= n holds ex P being Permutation of (Seg n) st ( P . 1 = i & P . 2 = j & ( for k being Nat st k in Seg n & k > 2 holds ( ( k <= i + 1 implies P . k = k - 2 ) & ( i + 1 < k & k <= j implies P . k = k - 1 ) & ( k > j implies P . k = k ) ) ) ) proof let i, j, n be Nat; ::_thesis: ( 1 <= i & i < j & j <= n implies ex P being Permutation of (Seg n) st ( P . 1 = i & P . 2 = j & ( for k being Nat st k in Seg n & k > 2 holds ( ( k <= i + 1 implies P . k = k - 2 ) & ( i + 1 < k & k <= j implies P . k = k - 1 ) & ( k > j implies P . k = k ) ) ) ) ) assume that A1: 1 <= i and A2: i < j and A3: j <= n ; ::_thesis: ex P being Permutation of (Seg n) st ( P . 1 = i & P . 2 = j & ( for k being Nat st k in Seg n & k > 2 holds ( ( k <= i + 1 implies P . k = k - 2 ) & ( i + 1 < k & k <= j implies P . k = k - 1 ) & ( k > j implies P . k = k ) ) ) ) i <= n by A2, A3, XXREAL_0:2; then A4: i in Seg n by A1, FINSEQ_1:1; 1 <= j by A1, A2, XXREAL_0:2; then A5: j in Seg n by A3, FINSEQ_1:1; reconsider S = Seg n as non empty Subset of NAT by A1, A2, A3; defpred S1[ Nat, Nat] means ( ( $1 = 1 implies $2 = i ) & ( $1 = 2 implies $2 = j ) & ( $1 > 2 implies ( ( $1 <= i + 1 implies $2 = $1 - 2 ) & ( i + 1 < $1 & $1 <= j implies $2 = $1 - 1 ) & ( $1 > j implies $2 = $1 ) ) ) ); A6: i + 1 < j + 1 by A2, XREAL_1:8; A7: for k being Nat st k in Seg n holds ex d being Element of S st S1[k,d] proof let k be Nat; ::_thesis: ( k in Seg n implies ex d being Element of S st S1[k,d] ) assume A8: k in Seg n ; ::_thesis: ex d being Element of S st S1[k,d] then A9: k <= n by FINSEQ_1:1; A10: k <> 0 by A8, FINSEQ_1:1; percases ( k = 1 or k = 2 or ( k > 2 & k <> 1 & k <> 2 ) ) by A10, NAT_1:26; suppose ( k = 1 or k = 2 ) ; ::_thesis: ex d being Element of S st S1[k,d] hence ex d being Element of S st S1[k,d] by A4, A5; ::_thesis: verum end; supposeA11: ( k > 2 & k <> 1 & k <> 2 ) ; ::_thesis: ex d being Element of S st S1[k,d] then reconsider k2 = k - 2 as Nat by NAT_1:21; k2 > 2 - 2 by A11, XREAL_1:8; then A12: k2 >= 1 by NAT_1:14; A13: k2 <= k - 0 by XREAL_1:10; percases ( k <= i + 1 or ( k > i + 1 & k <= j ) or ( k > i + 1 & k > j & k in S ) ) by A8; supposeA14: k <= i + 1 ; ::_thesis: ex d being Element of S st S1[k,d] then k < j + 1 by A6, XXREAL_0:2; then A15: k <= j by NAT_1:13; k2 <= n by A9, A13, XXREAL_0:2; then k2 in S by A12, FINSEQ_1:1; hence ex d being Element of S st S1[k,d] by A11, A14, A15; ::_thesis: verum end; supposeA16: ( k > i + 1 & k <= j ) ; ::_thesis: ex d being Element of S st S1[k,d] set k1 = k2 + 1; k2 + 1 <= (k2 + 1) + 1 by NAT_1:11; then A17: k2 + 1 <= n by A9, XXREAL_0:2; k2 + 1 >= 1 by NAT_1:11; then k2 + 1 in S by A17; hence ex d being Element of S st S1[k,d] by A11, A16; ::_thesis: verum end; suppose ( k > i + 1 & k > j & k in S ) ; ::_thesis: ex d being Element of S st S1[k,d] hence ex d being Element of S st S1[k,d] by A11; ::_thesis: verum end; end; end; end; end; consider f being FinSequence of S such that A18: ( len f = n & ( for k being Nat st k in Seg n holds S1[k,f /. k] ) ) from FINSEQ_4:sch_1(A7); A19: 1 < j by A1, A2, XXREAL_0:2; then 1 <= n by A3, XXREAL_0:2; then A20: 1 in S ; then A21: S1[1,f /. 1] by A18; 1 + 1 <= j by A19, NAT_1:13; then 2 <= n by A3, XXREAL_0:2; then A22: 2 in S ; then A23: S1[2,f /. 2] by A18; A24: dom f = S by A18, FINSEQ_1:def_3; then A25: f /. 1 = f . 1 by A20, PARTFUN1:def_6; A26: rng f c= S by FINSEQ_1:def_4; then reconsider F = f as Function of S,S by A24, FUNCT_2:2; A27: f /. 2 = f . 2 by A22, A24, PARTFUN1:def_6; S c= rng f proof let x be set ; :: according to TARSKI:def_3 ::_thesis: ( not x in S or x in rng f ) assume A28: x in S ; ::_thesis: x in rng f then reconsider k = x as Nat ; set k1 = k + 1; set k2 = (k + 1) + 1; A29: 1 <= k by A28, FINSEQ_1:1; percases ( k < i or k = i or ( k > i & k < j ) or k = j or k > j ) by XXREAL_0:1; supposeA30: k < i ; ::_thesis: x in rng f A31: k + 2 > 0 + 2 by A29, XREAL_1:8; A32: k + 1 <= i by A30, NAT_1:13; then k + 1 < j by A2, XXREAL_0:2; then (k + 1) + 1 <= j by NAT_1:13; then ( 1 <= (k + 1) + 1 & (k + 1) + 1 <= n ) by A3, NAT_1:11, XXREAL_0:2; then A33: (k + 1) + 1 in S ; then S1[(k + 1) + 1,f /. ((k + 1) + 1)] by A18; then f . ((k + 1) + 1) = k by A24, A31, A32, A33, PARTFUN1:def_6, XREAL_1:6; hence x in rng f by A24, A33, FUNCT_1:def_3; ::_thesis: verum end; suppose k = i ; ::_thesis: x in rng f hence x in rng f by A20, A21, A24, A25, FUNCT_1:def_3; ::_thesis: verum end; supposeA34: ( k > i & k < j ) ; ::_thesis: x in rng f then k > 1 by A1, A29, XXREAL_0:1; then A35: k >= 1 + 1 by NAT_1:13; k + 1 <= j by A34, NAT_1:13; then ( 1 <= k + 1 & k + 1 <= n ) by A3, NAT_1:11, XXREAL_0:2; then A36: k + 1 in S ; then S1[k + 1,f /. (k + 1)] by A18; then f . (k + 1) = k by A24, A34, A35, A36, NAT_1:13, PARTFUN1:def_6, XREAL_1:8; hence x in rng f by A24, A36, FUNCT_1:def_3; ::_thesis: verum end; suppose k = j ; ::_thesis: x in rng f hence x in rng f by A22, A23, A24, A27, FUNCT_1:def_3; ::_thesis: verum end; supposeA37: k > j ; ::_thesis: x in rng f j > 1 by A1, A2, XXREAL_0:2; then A38: j >= 1 + 1 by NAT_1:13; S1[k,f /. k] by A18, A28; then f . k = k by A24, A37, A38, PARTFUN1:def_6, XXREAL_0:2; hence x in rng f by A24, A28, FUNCT_1:def_3; ::_thesis: verum end; end; end; then rng F = S by A26, XBOOLE_0:def_10; then A39: F is onto by FUNCT_2:def_3; card S = card S ; then F is one-to-one by A39, STIRL2_1:60; then reconsider F = F as Permutation of (Seg n) by A39; take F ; ::_thesis: ( F . 1 = i & F . 2 = j & ( for k being Nat st k in Seg n & k > 2 holds ( ( k <= i + 1 implies F . k = k - 2 ) & ( i + 1 < k & k <= j implies F . k = k - 1 ) & ( k > j implies F . k = k ) ) ) ) thus ( F . 1 = i & F . 2 = j ) by A20, A21, A22, A23, A24, PARTFUN1:def_6; ::_thesis: for k being Nat st k in Seg n & k > 2 holds ( ( k <= i + 1 implies F . k = k - 2 ) & ( i + 1 < k & k <= j implies F . k = k - 1 ) & ( k > j implies F . k = k ) ) let k be Nat; ::_thesis: ( k in Seg n & k > 2 implies ( ( k <= i + 1 implies F . k = k - 2 ) & ( i + 1 < k & k <= j implies F . k = k - 1 ) & ( k > j implies F . k = k ) ) ) assume that A40: k in Seg n and A41: k > 2 ; ::_thesis: ( ( k <= i + 1 implies F . k = k - 2 ) & ( i + 1 < k & k <= j implies F . k = k - 1 ) & ( k > j implies F . k = k ) ) f /. k = f . k by A24, A40, PARTFUN1:def_6; hence ( ( k <= i + 1 implies F . k = k - 2 ) & ( i + 1 < k & k <= j implies F . k = k - 1 ) & ( k > j implies F . k = k ) ) by A18, A40, A41; ::_thesis: verum end; Lm3: for r being real number for i, j, n being Nat st 1 <= i & i < j & j <= n holds ex A being Matrix of n,F_Real st ( Det A = 1. F_Real & A * (i,i) = cos r & A * (j,j) = cos r & A * (i,j) = sin r & A * (j,i) = - (sin r) & ( for k, m being Nat st [k,m] in Indices A holds ( ( k = m & k <> i & k <> j implies A * (k,k) = 1. F_Real ) & ( k <> m & {k,m} <> {i,j} implies A * (k,m) = 0. F_Real ) ) ) ) proof let r be real number ; ::_thesis: for i, j, n being Nat st 1 <= i & i < j & j <= n holds ex A being Matrix of n,F_Real st ( Det A = 1. F_Real & A * (i,i) = cos r & A * (j,j) = cos r & A * (i,j) = sin r & A * (j,i) = - (sin r) & ( for k, m being Nat st [k,m] in Indices A holds ( ( k = m & k <> i & k <> j implies A * (k,k) = 1. F_Real ) & ( k <> m & {k,m} <> {i,j} implies A * (k,m) = 0. F_Real ) ) ) ) let i, j, n be Nat; ::_thesis: ( 1 <= i & i < j & j <= n implies ex A being Matrix of n,F_Real st ( Det A = 1. F_Real & A * (i,i) = cos r & A * (j,j) = cos r & A * (i,j) = sin r & A * (j,i) = - (sin r) & ( for k, m being Nat st [k,m] in Indices A holds ( ( k = m & k <> i & k <> j implies A * (k,k) = 1. F_Real ) & ( k <> m & {k,m} <> {i,j} implies A * (k,m) = 0. F_Real ) ) ) ) ) A1: now__::_thesis:_for_k_being_Nat_st_k_>=_1_&_k_<>_1_&_k_<>_2_holds_ k_>_2 let k be Nat; ::_thesis: ( k >= 1 & k <> 1 & k <> 2 implies k > 2 ) assume that A2: ( k >= 1 & k <> 1 ) and A3: k <> 2 ; ::_thesis: k > 2 k > 1 by A2, XXREAL_0:1; then k >= 1 + 1 by NAT_1:13; hence k > 2 by A3, XXREAL_0:1; ::_thesis: verum end; reconsider s = sin r, c = cos r as Element of F_Real by XREAL_0:def_1; set S = Seg n; assume that A4: 1 <= i and A5: i < j and A6: j <= n ; ::_thesis: ex A being Matrix of n,F_Real st ( Det A = 1. F_Real & A * (i,i) = cos r & A * (j,j) = cos r & A * (i,j) = sin r & A * (j,i) = - (sin r) & ( for k, m being Nat st [k,m] in Indices A holds ( ( k = m & k <> i & k <> j implies A * (k,k) = 1. F_Real ) & ( k <> m & {k,m} <> {i,j} implies A * (k,m) = 0. F_Real ) ) ) ) A7: 1 < j by A4, A5, XXREAL_0:2; then A8: 1 + 1 <= j by NAT_1:13; then reconsider n2 = n - 2 as Nat by A6, NAT_1:21, XXREAL_0:2; consider P being Permutation of (Seg n) such that A9: P . 1 = i and A10: P . 2 = j and for k being Nat st k in Seg n & k > 2 holds ( ( k <= i + 1 implies P . k = k - 2 ) & ( i + 1 < k & k <= j implies P . k = k - 1 ) & ( k > j implies P . k = k ) ) by A4, A5, A6, Lm2; reconsider p = P as one-to-one Function ; dom P = Seg n by FUNCT_2:52; then A11: rng (p ") = Seg n by FUNCT_1:33; rng P = Seg n by FUNCT_2:def_3; then dom (p ") = Seg n by FUNCT_1:33; then reconsider P1 = p " as one-to-one Function of (Seg n),(Seg n) by A11, FUNCT_2:1; P1 is onto by A11, FUNCT_2:def_3; then reconsider P1 = P1 as Permutation of (Seg n) ; A12: dom P = Seg n by FUNCT_2:52; 1 <= n by A6, A7, XXREAL_0:2; then A13: 1 in Seg n ; then A14: P1 . (P . 1) = 1 by A12, FUNCT_1:34; set A = (c,s) ][ ((- s),c); set ONE = 1. (F_Real,n2); set ao = <*((c,s) ][ ((- s),c)),(1. (F_Real,n2))*>; set B = block_diagonal (<*((c,s) ][ ((- s),c)),(1. (F_Real,n2))*>,(0. F_Real)); A15: len ((c,s) ][ ((- s),c)) = 2 by MATRIX_1:25; then Len <*((c,s) ][ ((- s),c))*> = <*2*> by MATRIXJ1:15; then A16: Sum (Len <*((c,s) ][ ((- s),c))*>) = 2 by RVSUM_1:73; len (1. (F_Real,n2)) = n2 by MATRIX_1:25; then A17: Sum (Len <*((c,s) ][ ((- s),c)),(1. (F_Real,n2))*>) = 2 + n2 by A15, MATRIXJ1:16; then reconsider B = block_diagonal (<*((c,s) ][ ((- s),c)),(1. (F_Real,n2))*>,(0. F_Real)) as Matrix of n,F_Real ; A18: Indices B = [:(Seg n),(Seg n):] by MATRIX_1:24; 2 <= n by A6, A8, XXREAL_0:2; then A19: 2 in Seg n ; then A20: P1 . (P . 2) = 2 by A12, FUNCT_1:34; set pBp = (((B * P1) @) * P1) @ ; A21: dom P1 = Seg n by FUNCT_2:52; i < n by A5, A6, XXREAL_0:2; then A22: i in Seg n by A4, FINSEQ_1:1; then [i,i] in Indices B by A18, ZFMISC_1:87; then A23: ((((B * P1) @) * P1) @) * (i,i) = B * (1,1) by A9, A14, Th1; take (((B * P1) @) * P1) @ ; ::_thesis: ( Det ((((B * P1) @) * P1) @) = 1. F_Real & ((((B * P1) @) * P1) @) * (i,i) = cos r & ((((B * P1) @) * P1) @) * (j,j) = cos r & ((((B * P1) @) * P1) @) * (i,j) = sin r & ((((B * P1) @) * P1) @) * (j,i) = - (sin r) & ( for k, m being Nat st [k,m] in Indices ((((B * P1) @) * P1) @) holds ( ( k = m & k <> i & k <> j implies ((((B * P1) @) * P1) @) * (k,k) = 1. F_Real ) & ( k <> m & {k,m} <> {i,j} implies ((((B * P1) @) * P1) @) * (k,m) = 0. F_Real ) ) ) ) A24: Indices ((((B * P1) @) * P1) @) = [:(Seg n),(Seg n):] by MATRIX_1:24; A25: Det B = 1. F_Real by A17, Th12; A26: block_diagonal (<*((c,s) ][ ((- s),c))*>,(0. F_Real)) = (c,s) ][ ((- s),c) by MATRIXJ1:34; width ((c,s) ][ ((- s),c)) = 2 by MATRIX_1:24; then Width <*((c,s) ][ ((- s),c))*> = <*2*> by MATRIXJ1:19; then A27: Sum (Width <*((c,s) ][ ((- s),c))*>) = 2 by RVSUM_1:73; 1 < j by A4, A5, XXREAL_0:2; then A28: j in Seg n by A6, FINSEQ_1:1; then [j,j] in Indices B by A18, ZFMISC_1:87; then A29: ((((B * P1) @) * P1) @) * (j,j) = B * (2,2) by A10, A20, Th1; A30: block_diagonal (<*(1. (F_Real,n2))*>,(0. F_Real)) = 1. (F_Real,n2) by MATRIXJ1:34; [1,1] in Indices ((c,s) ][ ((- s),c)) by MATRIX_2:4; then A31: B * (1,1) = ((c,s) ][ ((- s),c)) * (1,1) by A26, A30, MATRIXJ1:26; [2,1] in Indices ((c,s) ][ ((- s),c)) by MATRIX_2:4; then A32: B * (2,1) = ((c,s) ][ ((- s),c)) * (2,1) by A26, A30, MATRIXJ1:26; [1,2] in Indices ((c,s) ][ ((- s),c)) by MATRIX_2:4; then A33: B * (1,2) = ((c,s) ][ ((- s),c)) * (1,2) by A26, A30, MATRIXJ1:26; [2,2] in Indices ((c,s) ][ ((- s),c)) by MATRIX_2:4; then A34: B * (2,2) = ((c,s) ][ ((- s),c)) * (2,2) by A26, A30, MATRIXJ1:26; A35: <*((c,s) ][ ((- s),c))*> ^ <*(1. (F_Real,n2))*> = <*((c,s) ][ ((- s),c)),(1. (F_Real,n2))*> ; [i,j] in Indices B by A18, A22, A28, ZFMISC_1:87; then A36: ((((B * P1) @) * P1) @) * (i,j) = B * (1,2) by A9, A10, A14, A20, Th1; [j,i] in Indices B by A18, A22, A28, ZFMISC_1:87; then ((((B * P1) @) * P1) @) * (j,i) = B * (2,1) by A9, A10, A14, A20, Th1; hence ( Det ((((B * P1) @) * P1) @) = 1. F_Real & ((((B * P1) @) * P1) @) * (i,i) = cos r & ((((B * P1) @) * P1) @) * (j,j) = cos r & ((((B * P1) @) * P1) @) * (i,j) = sin r & ((((B * P1) @) * P1) @) * (j,i) = - (sin r) ) by A23, A25, A29, A34, A33, A32, A31, A36, Th1, MATRIX_2:6; ::_thesis: for k, m being Nat st [k,m] in Indices ((((B * P1) @) * P1) @) holds ( ( k = m & k <> i & k <> j implies ((((B * P1) @) * P1) @) * (k,k) = 1. F_Real ) & ( k <> m & {k,m} <> {i,j} implies ((((B * P1) @) * P1) @) * (k,m) = 0. F_Real ) ) let k, m be Nat; ::_thesis: ( [k,m] in Indices ((((B * P1) @) * P1) @) implies ( ( k = m & k <> i & k <> j implies ((((B * P1) @) * P1) @) * (k,k) = 1. F_Real ) & ( k <> m & {k,m} <> {i,j} implies ((((B * P1) @) * P1) @) * (k,m) = 0. F_Real ) ) ) assume A37: [k,m] in Indices ((((B * P1) @) * P1) @) ; ::_thesis: ( ( k = m & k <> i & k <> j implies ((((B * P1) @) * P1) @) * (k,k) = 1. F_Real ) & ( k <> m & {k,m} <> {i,j} implies ((((B * P1) @) * P1) @) * (k,m) = 0. F_Real ) ) A38: k in Seg n by A24, A37, ZFMISC_1:87; set Pk = P1 . k; set Pm = P1 . m; A39: rng P1 = Seg n by FUNCT_2:def_3; then A40: P1 . k in Seg n by A21, A38, FUNCT_1:def_3; then A41: P1 . k >= 1 by FINSEQ_1:1; A42: m in Seg n by A24, A37, ZFMISC_1:87; then A43: P1 . m in Seg n by A21, A39, FUNCT_1:def_3; then A44: [(P1 . k),(P1 . m)] in [:(Seg n),(Seg n):] by A40, ZFMISC_1:87; A45: ((((B * P1) @) * P1) @) * (k,m) = B * ((P1 . k),(P1 . m)) by A18, A24, A37, Th1; thus ( k = m & k <> i & k <> j implies ((((B * P1) @) * P1) @) * (k,k) = 1. F_Real ) ::_thesis: ( k <> m & {k,m} <> {i,j} implies ((((B * P1) @) * P1) @) * (k,m) = 0. F_Real ) proof assume that A46: k = m and A47: ( k <> i & k <> j ) ; ::_thesis: ((((B * P1) @) * P1) @) * (k,k) = 1. F_Real ( P1 . k <> 1 & P1 . k <> 2 ) by A9, A10, A14, A20, A21, A22, A28, A38, A47, FUNCT_1:def_4; then A48: P1 . k > 2 by A1, A41; then reconsider Pk2 = (P1 . k) - 2 as Nat by NAT_1:21; ( P1 . k = Pk2 + 2 & Pk2 > 2 - 2 ) by A48, XREAL_1:8; then A49: [Pk2,Pk2] in Indices (1. (F_Real,n2)) by A16, A18, A27, A30, A44, A46, MATRIXJ1:27; then (1. (F_Real,n2)) * (Pk2,Pk2) = B * ((Pk2 + 2),(Pk2 + 2)) by A16, A27, A30, MATRIXJ1:28; hence ((((B * P1) @) * P1) @) * (k,k) = 1. F_Real by A45, A46, A49, MATRIX_1:def_11; ::_thesis: verum end; A50: P1 . m >= 1 by A43, FINSEQ_1:1; thus ( k <> m & {k,m} <> {i,j} implies ((((B * P1) @) * P1) @) * (k,m) = 0. F_Real ) ::_thesis: verum proof assume that A51: k <> m and A52: {k,m} <> {i,j} ; ::_thesis: ((((B * P1) @) * P1) @) * (k,m) = 0. F_Real A53: P1 . k <> P1 . m by A21, A38, A42, A51, FUNCT_1:def_4; percases ( ( k <> i & k <> j & m <> i & m <> j ) or ( k = i & m <> j ) or ( k = j & m <> i ) or ( m = i & k <> j ) or ( m = j & k <> i ) ) by A52; supposeA54: ( k <> i & k <> j & m <> i & m <> j ) ; ::_thesis: ((((B * P1) @) * P1) @) * (k,m) = 0. F_Real then ( P1 . k <> 1 & P1 . k <> 2 ) by A9, A10, A14, A20, A21, A22, A28, A38, FUNCT_1:def_4; then A55: P1 . k > 2 by A1, A41; ( P1 . m <> 1 & P1 . m <> 2 ) by A9, A10, A14, A20, A21, A22, A28, A42, A54, FUNCT_1:def_4; then A56: P1 . m > 2 by A1, A50; then reconsider Pk2 = (P1 . k) - 2, Pm2 = (P1 . m) - 2 as Nat by A55, NAT_1:21; A57: Pk2 > 2 - 2 by A55, XREAL_1:8; A58: ( P1 . k = Pk2 + 2 & P1 . m = Pm2 + 2 ) ; Pm2 > 2 - 2 by A56, XREAL_1:8; then A59: [Pk2,Pm2] in Indices (1. (F_Real,n2)) by A16, A18, A27, A30, A44, A58, A57, MATRIXJ1:27; then (1. (F_Real,n2)) * (Pk2,Pm2) = B * ((Pk2 + 2),(Pm2 + 2)) by A16, A27, A30, MATRIXJ1:28; hence ((((B * P1) @) * P1) @) * (k,m) = 0. F_Real by A45, A53, A59, MATRIX_1:def_11; ::_thesis: verum end; supposeA60: ( k = i & m <> j ) ; ::_thesis: ((((B * P1) @) * P1) @) * (k,m) = 0. F_Real then P1 . m <> 2 by A10, A20, A21, A28, A42, FUNCT_1:def_4; then A61: P1 . m > 2 by A1, A9, A14, A50, A53, A60; P1 . k = 1 by A9, A12, A13, A60, FUNCT_1:34; hence ((((B * P1) @) * P1) @) * (k,m) = 0. F_Real by A16, A18, A27, A35, A44, A45, A61, MATRIXJ1:29; ::_thesis: verum end; supposeA62: ( k = j & m <> i ) ; ::_thesis: ((((B * P1) @) * P1) @) * (k,m) = 0. F_Real then P1 . m <> 1 by A9, A14, A21, A22, A42, FUNCT_1:def_4; then A63: P1 . m > 2 by A1, A10, A20, A50, A53, A62; P1 . k = 2 by A10, A12, A19, A62, FUNCT_1:34; hence ((((B * P1) @) * P1) @) * (k,m) = 0. F_Real by A16, A18, A27, A35, A44, A45, A63, MATRIXJ1:29; ::_thesis: verum end; supposeA64: ( m = i & k <> j ) ; ::_thesis: ((((B * P1) @) * P1) @) * (k,m) = 0. F_Real then P1 . k <> 2 by A10, A20, A21, A28, A38, FUNCT_1:def_4; then A65: P1 . k > 2 by A1, A9, A14, A41, A53, A64; P1 . m = 1 by A9, A12, A13, A64, FUNCT_1:34; hence ((((B * P1) @) * P1) @) * (k,m) = 0. F_Real by A16, A18, A27, A35, A44, A45, A65, MATRIXJ1:29; ::_thesis: verum end; supposeA66: ( m = j & k <> i ) ; ::_thesis: ((((B * P1) @) * P1) @) * (k,m) = 0. F_Real then P1 . k <> 1 by A9, A14, A21, A22, A38, FUNCT_1:def_4; then A67: P1 . k > 2 by A1, A10, A20, A41, A53, A66; P1 . m = 2 by A10, A12, A19, A66, FUNCT_1:34; hence ((((B * P1) @) * P1) @) * (k,m) = 0. F_Real by A16, A18, A27, A35, A44, A45, A67, MATRIXJ1:29; ::_thesis: verum end; end; end; end; begin definition let n be Nat; let r be real number ; let i, j be Nat; assume B1: ( 1 <= i & i < j & j <= n ) ; func Rotation (i,j,n,r) -> invertible Matrix of n,F_Real means :Def3: :: MATRTOP3:def 3 ( it * (i,i) = cos r & it * (j,j) = cos r & it * (i,j) = sin r & it * (j,i) = - (sin r) & ( for k, m being Nat st [k,m] in Indices it holds ( ( k = m & k <> i & k <> j implies it * (k,k) = 1. F_Real ) & ( k <> m & {k,m} <> {i,j} implies it * (k,m) = 0. F_Real ) ) ) ); existence ex b1 being invertible Matrix of n,F_Real st ( b1 * (i,i) = cos r & b1 * (j,j) = cos r & b1 * (i,j) = sin r & b1 * (j,i) = - (sin r) & ( for k, m being Nat st [k,m] in Indices b1 holds ( ( k = m & k <> i & k <> j implies b1 * (k,k) = 1. F_Real ) & ( k <> m & {k,m} <> {i,j} implies b1 * (k,m) = 0. F_Real ) ) ) ) proof consider A being Matrix of n,F_Real such that A1: Det A = 1. F_Real and A2: ( A * (i,i) = cos r & A * (j,j) = cos r & A * (i,j) = sin r & A * (j,i) = - (sin r) & ( for k, m being Nat st [k,m] in Indices A holds ( ( k = m & k <> i & k <> j implies A * (k,k) = 1. F_Real ) & ( k <> m & {k,m} <> {i,j} implies A * (k,m) = 0. F_Real ) ) ) ) by B1, Lm3; Det A <> 0. F_Real by A1; then A is invertible by LAPLACE:34; hence ex b1 being invertible Matrix of n,F_Real st ( b1 * (i,i) = cos r & b1 * (j,j) = cos r & b1 * (i,j) = sin r & b1 * (j,i) = - (sin r) & ( for k, m being Nat st [k,m] in Indices b1 holds ( ( k = m & k <> i & k <> j implies b1 * (k,k) = 1. F_Real ) & ( k <> m & {k,m} <> {i,j} implies b1 * (k,m) = 0. F_Real ) ) ) ) by A2; ::_thesis: verum end; uniqueness for b1, b2 being invertible Matrix of n,F_Real st b1 * (i,i) = cos r & b1 * (j,j) = cos r & b1 * (i,j) = sin r & b1 * (j,i) = - (sin r) & ( for k, m being Nat st [k,m] in Indices b1 holds ( ( k = m & k <> i & k <> j implies b1 * (k,k) = 1. F_Real ) & ( k <> m & {k,m} <> {i,j} implies b1 * (k,m) = 0. F_Real ) ) ) & b2 * (i,i) = cos r & b2 * (j,j) = cos r & b2 * (i,j) = sin r & b2 * (j,i) = - (sin r) & ( for k, m being Nat st [k,m] in Indices b2 holds ( ( k = m & k <> i & k <> j implies b2 * (k,k) = 1. F_Real ) & ( k <> m & {k,m} <> {i,j} implies b2 * (k,m) = 0. F_Real ) ) ) holds b1 = b2 proof let I1, I2 be invertible Matrix of n,F_Real; ::_thesis: ( I1 * (i,i) = cos r & I1 * (j,j) = cos r & I1 * (i,j) = sin r & I1 * (j,i) = - (sin r) & ( for k, m being Nat st [k,m] in Indices I1 holds ( ( k = m & k <> i & k <> j implies I1 * (k,k) = 1. F_Real ) & ( k <> m & {k,m} <> {i,j} implies I1 * (k,m) = 0. F_Real ) ) ) & I2 * (i,i) = cos r & I2 * (j,j) = cos r & I2 * (i,j) = sin r & I2 * (j,i) = - (sin r) & ( for k, m being Nat st [k,m] in Indices I2 holds ( ( k = m & k <> i & k <> j implies I2 * (k,k) = 1. F_Real ) & ( k <> m & {k,m} <> {i,j} implies I2 * (k,m) = 0. F_Real ) ) ) implies I1 = I2 ) assume that A3: ( I1 * (i,i) = cos r & I1 * (j,j) = cos r & I1 * (i,j) = sin r & I1 * (j,i) = - (sin r) ) and A4: for k, m being Nat st [k,m] in Indices I1 holds ( ( k = m & k <> i & k <> j implies I1 * (k,k) = 1. F_Real ) & ( k <> m & {k,m} <> {i,j} implies I1 * (k,m) = 0. F_Real ) ) and A5: ( I2 * (i,i) = cos r & I2 * (j,j) = cos r & I2 * (i,j) = sin r & I2 * (j,i) = - (sin r) ) and A6: for k, m being Nat st [k,m] in Indices I2 holds ( ( k = m & k <> i & k <> j implies I2 * (k,k) = 1. F_Real ) & ( k <> m & {k,m} <> {i,j} implies I2 * (k,m) = 0. F_Real ) ) ; ::_thesis: I1 = I2 for k, m being Nat st [k,m] in Indices I1 holds I1 * (k,m) = I2 * (k,m) proof let k, m be Nat; ::_thesis: ( [k,m] in Indices I1 implies I1 * (k,m) = I2 * (k,m) ) assume A7: [k,m] in Indices I1 ; ::_thesis: I1 * (k,m) = I2 * (k,m) then A8: [k,m] in Indices I2 by MATRIX_1:26; percases ( ( k = m & ( k = i or k = j ) ) or ( k <> m & ( ( k = i & m = j ) or ( k = j & m = i ) ) ) or ( k = m & k <> i & k <> j ) or ( k <> m & ( ( k = i & m <> j ) or ( k = j & m <> i ) or ( m = i & k <> j ) or ( m = j & k <> i ) or ( k <> i & k <> j & m <> i & m <> j ) ) ) ) ; suppose ( ( k = m & ( k = i or k = j ) ) or ( k <> m & ( ( k = i & m = j ) or ( k = j & m = i ) ) ) ) ; ::_thesis: I1 * (k,m) = I2 * (k,m) hence I1 * (k,m) = I2 * (k,m) by A3, A5; ::_thesis: verum end; supposeA9: ( k = m & k <> i & k <> j ) ; ::_thesis: I1 * (k,m) = I2 * (k,m) then I1 * (k,m) = 1. F_Real by A4, A7; hence I1 * (k,m) = I2 * (k,m) by A6, A8, A9; ::_thesis: verum end; supposeA10: ( k <> m & ( ( k = i & m <> j ) or ( k = j & m <> i ) or ( m = i & k <> j ) or ( m = j & k <> i ) or ( k <> i & k <> j & m <> i & m <> j ) ) ) ; ::_thesis: I1 * (k,m) = I2 * (k,m) then A11: {k,m} <> {i,j} by ZFMISC_1:6; then I1 * (k,m) = 0. F_Real by A4, A7, A10; hence I1 * (k,m) = I2 * (k,m) by A6, A8, A10, A11; ::_thesis: verum end; end; end; hence I1 = I2 by MATRIX_1:27; ::_thesis: verum end; end; :: deftheorem Def3 defines Rotation MATRTOP3:def_3_:_ for n being Nat for r being real number for i, j being Nat st 1 <= i & i < j & j <= n holds for b5 being invertible Matrix of n,F_Real holds ( b5 = Rotation (i,j,n,r) iff ( b5 * (i,i) = cos r & b5 * (j,j) = cos r & b5 * (i,j) = sin r & b5 * (j,i) = - (sin r) & ( for k, m being Nat st [k,m] in Indices b5 holds ( ( k = m & k <> i & k <> j implies b5 * (k,k) = 1. F_Real ) & ( k <> m & {k,m} <> {i,j} implies b5 * (k,m) = 0. F_Real ) ) ) ) ); theorem Th13: :: MATRTOP3:13 for r being real number for i, j, n being Nat st 1 <= i & i < j & j <= n holds Det (Rotation (i,j,n,r)) = 1. F_Real proof let r be real number ; ::_thesis: for i, j, n being Nat st 1 <= i & i < j & j <= n holds Det (Rotation (i,j,n,r)) = 1. F_Real let i, j, n be Nat; ::_thesis: ( 1 <= i & i < j & j <= n implies Det (Rotation (i,j,n,r)) = 1. F_Real ) assume A1: ( 1 <= i & i < j & j <= n ) ; ::_thesis: Det (Rotation (i,j,n,r)) = 1. F_Real then consider A being Matrix of n,F_Real such that A2: Det A = 1. F_Real and A3: ( A * (i,i) = cos r & A * (j,j) = cos r & A * (i,j) = sin r & A * (j,i) = - (sin r) & ( for k, m being Nat st [k,m] in Indices A holds ( ( k = m & k <> i & k <> j implies A * (k,k) = 1. F_Real ) & ( k <> m & {k,m} <> {i,j} implies A * (k,m) = 0. F_Real ) ) ) ) by Lm3; Det A <> 0. F_Real by A2; then A is invertible by LAPLACE:34; hence Det (Rotation (i,j,n,r)) = 1. F_Real by A1, A2, A3, Def3; ::_thesis: verum end; theorem Th14: :: MATRTOP3:14 for r being real number for i, j, n, k being Nat for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n & k in Seg n & k <> i & k <> j holds (@ p) "*" (Col ((Rotation (i,j,n,r)),k)) = p . k proof let r be real number ; ::_thesis: for i, j, n, k being Nat for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n & k in Seg n & k <> i & k <> j holds (@ p) "*" (Col ((Rotation (i,j,n,r)),k)) = p . k let i, j, n, k be Nat; ::_thesis: for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n & k in Seg n & k <> i & k <> j holds (@ p) "*" (Col ((Rotation (i,j,n,r)),k)) = p . k let p be Point of (TOP-REAL n); ::_thesis: ( 1 <= i & i < j & j <= n & k in Seg n & k <> i & k <> j implies (@ p) "*" (Col ((Rotation (i,j,n,r)),k)) = p . k ) set S = Seg n; assume that A1: ( 1 <= i & i < j & j <= n ) and A2: k in Seg n and A3: ( k <> i & k <> j ) ; ::_thesis: (@ p) "*" (Col ((Rotation (i,j,n,r)),k)) = p . k set O = Rotation (i,j,n,r); set C = Col ((Rotation (i,j,n,r)),k); A4: Indices (Rotation (i,j,n,r)) = [:(Seg n),(Seg n):] by MATRIX_1:24; then A5: [k,k] in Indices (Rotation (i,j,n,r)) by A2, ZFMISC_1:87; A6: len (Rotation (i,j,n,r)) = n by MATRIX_1:25; then A7: dom (Rotation (i,j,n,r)) = Seg n by FINSEQ_1:def_3; len (Col ((Rotation (i,j,n,r)),k)) = n by A6, MATRIX_1:def_8; then A8: dom (Col ((Rotation (i,j,n,r)),k)) = Seg n by FINSEQ_1:def_3; A9: now__::_thesis:_for_m_being_Nat_st_m_in_dom_(Col_((Rotation_(i,j,n,r)),k))_&_m_<>_k_holds_ (Col_((Rotation_(i,j,n,r)),k))_._m_=_0._F_Real let m be Nat; ::_thesis: ( m in dom (Col ((Rotation (i,j,n,r)),k)) & m <> k implies (Col ((Rotation (i,j,n,r)),k)) . m = 0. F_Real ) assume that A10: m in dom (Col ((Rotation (i,j,n,r)),k)) and A11: m <> k ; ::_thesis: (Col ((Rotation (i,j,n,r)),k)) . m = 0. F_Real A12: [m,k] in Indices (Rotation (i,j,n,r)) by A2, A4, A8, A10, ZFMISC_1:87; not k in {i,j} by A3, TARSKI:def_2; then A13: {m,k} <> {i,j} by TARSKI:def_2; thus (Col ((Rotation (i,j,n,r)),k)) . m = (Rotation (i,j,n,r)) * (m,k) by A7, A8, A10, MATRIX_1:def_8 .= 0. F_Real by A1, A11, A12, A13, Def3 ; ::_thesis: verum end; len p = n by CARD_1:def_7; then A14: dom p = Seg n by FINSEQ_1:def_3; (Col ((Rotation (i,j,n,r)),k)) . k = (Rotation (i,j,n,r)) * (k,k) by A2, A7, MATRIX_1:def_8 .= 1. F_Real by A1, A3, A5, Def3 ; hence p . k = Sum (mlt ((Col ((Rotation (i,j,n,r)),k)),(@ p))) by A2, A8, A9, A14, MATRIX_3:17 .= (@ p) "*" (Col ((Rotation (i,j,n,r)),k)) by FVSUM_1:64 ; ::_thesis: verum end; theorem Th15: :: MATRTOP3:15 for r being real number for i, j, n being Nat for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n holds (@ p) "*" (Col ((Rotation (i,j,n,r)),i)) = ((p . i) * (cos r)) + ((p . j) * (- (sin r))) proof let r be real number ; ::_thesis: for i, j, n being Nat for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n holds (@ p) "*" (Col ((Rotation (i,j,n,r)),i)) = ((p . i) * (cos r)) + ((p . j) * (- (sin r))) let i, j, n be Nat; ::_thesis: for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n holds (@ p) "*" (Col ((Rotation (i,j,n,r)),i)) = ((p . i) * (cos r)) + ((p . j) * (- (sin r))) let p be Point of (TOP-REAL n); ::_thesis: ( 1 <= i & i < j & j <= n implies (@ p) "*" (Col ((Rotation (i,j,n,r)),i)) = ((p . i) * (cos r)) + ((p . j) * (- (sin r))) ) assume that A1: 1 <= i and A2: i < j and A3: j <= n ; ::_thesis: (@ p) "*" (Col ((Rotation (i,j,n,r)),i)) = ((p . i) * (cos r)) + ((p . j) * (- (sin r))) set O = Rotation (i,j,n,r); set C = Col ((Rotation (i,j,n,r)),i); set S = Seg n; 1 <= j by A1, A2, XXREAL_0:2; then A4: j in Seg n by A3, FINSEQ_1:1; A5: len (Rotation (i,j,n,r)) = n by MATRIX_1:25; then A6: dom (Rotation (i,j,n,r)) = Seg n by FINSEQ_1:def_3; then A7: (Col ((Rotation (i,j,n,r)),i)) . j = (Rotation (i,j,n,r)) * (j,i) by A4, MATRIX_1:def_8; i <= n by A2, A3, XXREAL_0:2; then A8: i in Seg n by A1, FINSEQ_1:1; then A9: (Col ((Rotation (i,j,n,r)),i)) . i = (Rotation (i,j,n,r)) * (i,i) by A6, MATRIX_1:def_8; ( len p = n & len (Col ((Rotation (i,j,n,r)),i)) = n ) by A5, CARD_1:def_7; then A10: len (mlt ((@ p),(Col ((Rotation (i,j,n,r)),i)))) = n by MATRIX_3:6; then A11: dom (mlt ((@ p),(Col ((Rotation (i,j,n,r)),i)))) = Seg n by FINSEQ_1:def_3; A12: Indices (Rotation (i,j,n,r)) = [:(Seg n),(Seg n):] by MATRIX_1:24; for k being Nat st k in dom (mlt ((@ p),(Col ((Rotation (i,j,n,r)),i)))) & k <> i & k <> j holds (mlt ((@ p),(Col ((Rotation (i,j,n,r)),i)))) . k = 0. F_Real proof let k be Nat; ::_thesis: ( k in dom (mlt ((@ p),(Col ((Rotation (i,j,n,r)),i)))) & k <> i & k <> j implies (mlt ((@ p),(Col ((Rotation (i,j,n,r)),i)))) . k = 0. F_Real ) assume that A13: k in dom (mlt ((@ p),(Col ((Rotation (i,j,n,r)),i)))) and A14: k <> i and A15: k <> j ; ::_thesis: (mlt ((@ p),(Col ((Rotation (i,j,n,r)),i)))) . k = 0. F_Real not k in {i,j} by A14, A15, TARSKI:def_2; then A16: {k,i} <> {i,j} by TARSKI:def_2; reconsider pk = (@ p) . k as Element of F_Real ; A17: [k,i] in Indices (Rotation (i,j,n,r)) by A8, A11, A12, A13, ZFMISC_1:87; (Col ((Rotation (i,j,n,r)),i)) . k = (Rotation (i,j,n,r)) * (k,i) by A6, A11, A13, MATRIX_1:def_8; hence (mlt ((@ p),(Col ((Rotation (i,j,n,r)),i)))) . k = pk * ((Rotation (i,j,n,r)) * (k,i)) by A13, FVSUM_1:60 .= pk * (0. F_Real) by A1, A2, A3, A14, A16, A17, Def3 .= 0. F_Real ; ::_thesis: verum end; then A18: Sum (mlt ((@ p),(Col ((Rotation (i,j,n,r)),i)))) = ((mlt ((@ p),(Col ((Rotation (i,j,n,r)),i)))) /. i) + ((mlt ((@ p),(Col ((Rotation (i,j,n,r)),i)))) /. j) by A2, A4, A8, A11, MATRIX15:7; A19: i in dom (mlt ((@ p),(Col ((Rotation (i,j,n,r)),i)))) by A8, A10, FINSEQ_1:def_3; reconsider pii = (@ p) . i, pj = (@ p) . j as Element of F_Real ; A20: (mlt ((@ p),(Col ((Rotation (i,j,n,r)),i)))) /. i = (mlt ((@ p),(Col ((Rotation (i,j,n,r)),i)))) . i by A8, A11, PARTFUN1:def_6 .= pii * ((Rotation (i,j,n,r)) * (i,i)) by A9, A19, FVSUM_1:60 .= (p . i) * (cos r) by A1, A2, A3, Def3 ; (mlt ((@ p),(Col ((Rotation (i,j,n,r)),i)))) /. j = (mlt ((@ p),(Col ((Rotation (i,j,n,r)),i)))) . j by A4, A11, PARTFUN1:def_6 .= pj * ((Rotation (i,j,n,r)) * (j,i)) by A4, A7, A11, FVSUM_1:60 .= (p . j) * (- (sin r)) by A1, A2, A3, Def3 ; hence (@ p) "*" (Col ((Rotation (i,j,n,r)),i)) = ((p . i) * (cos r)) + ((p . j) * (- (sin r))) by A18, A20; ::_thesis: verum end; theorem Th16: :: MATRTOP3:16 for r being real number for i, j, n being Nat for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n holds (@ p) "*" (Col ((Rotation (i,j,n,r)),j)) = ((p . i) * (sin r)) + ((p . j) * (cos r)) proof let r be real number ; ::_thesis: for i, j, n being Nat for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n holds (@ p) "*" (Col ((Rotation (i,j,n,r)),j)) = ((p . i) * (sin r)) + ((p . j) * (cos r)) let i, j, n be Nat; ::_thesis: for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n holds (@ p) "*" (Col ((Rotation (i,j,n,r)),j)) = ((p . i) * (sin r)) + ((p . j) * (cos r)) let p be Point of (TOP-REAL n); ::_thesis: ( 1 <= i & i < j & j <= n implies (@ p) "*" (Col ((Rotation (i,j,n,r)),j)) = ((p . i) * (sin r)) + ((p . j) * (cos r)) ) assume that A1: 1 <= i and A2: i < j and A3: j <= n ; ::_thesis: (@ p) "*" (Col ((Rotation (i,j,n,r)),j)) = ((p . i) * (sin r)) + ((p . j) * (cos r)) set O = Rotation (i,j,n,r); set C = Col ((Rotation (i,j,n,r)),j); set S = Seg n; 1 <= j by A1, A2, XXREAL_0:2; then A4: j in Seg n by A3, FINSEQ_1:1; A5: len (Rotation (i,j,n,r)) = n by MATRIX_1:25; then A6: dom (Rotation (i,j,n,r)) = Seg n by FINSEQ_1:def_3; then A7: (Col ((Rotation (i,j,n,r)),j)) . j = (Rotation (i,j,n,r)) * (j,j) by A4, MATRIX_1:def_8; A8: i <= n by A2, A3, XXREAL_0:2; then A9: i in Seg n by A1, FINSEQ_1:1; then A10: (Col ((Rotation (i,j,n,r)),j)) . i = (Rotation (i,j,n,r)) * (i,j) by A6, MATRIX_1:def_8; ( len p = n & len (Col ((Rotation (i,j,n,r)),j)) = n ) by A5, CARD_1:def_7; then len (mlt ((@ p),(Col ((Rotation (i,j,n,r)),j)))) = n by MATRIX_3:6; then A11: dom (mlt ((@ p),(Col ((Rotation (i,j,n,r)),j)))) = Seg n by FINSEQ_1:def_3; then A12: i in dom (mlt ((@ p),(Col ((Rotation (i,j,n,r)),j)))) by A1, A8, FINSEQ_1:1; A13: Indices (Rotation (i,j,n,r)) = [:(Seg n),(Seg n):] by MATRIX_1:24; for k being Nat st k in dom (mlt ((@ p),(Col ((Rotation (i,j,n,r)),j)))) & k <> i & k <> j holds (mlt ((@ p),(Col ((Rotation (i,j,n,r)),j)))) . k = 0. F_Real proof let k be Nat; ::_thesis: ( k in dom (mlt ((@ p),(Col ((Rotation (i,j,n,r)),j)))) & k <> i & k <> j implies (mlt ((@ p),(Col ((Rotation (i,j,n,r)),j)))) . k = 0. F_Real ) assume that A14: k in dom (mlt ((@ p),(Col ((Rotation (i,j,n,r)),j)))) and A15: k <> i and A16: k <> j ; ::_thesis: (mlt ((@ p),(Col ((Rotation (i,j,n,r)),j)))) . k = 0. F_Real not k in {i,j} by A15, A16, TARSKI:def_2; then A17: {k,j} <> {i,j} by TARSKI:def_2; reconsider pk = (@ p) . k as Element of F_Real ; A18: [k,j] in Indices (Rotation (i,j,n,r)) by A4, A11, A13, A14, ZFMISC_1:87; (Col ((Rotation (i,j,n,r)),j)) . k = (Rotation (i,j,n,r)) * (k,j) by A6, A11, A14, MATRIX_1:def_8; hence (mlt ((@ p),(Col ((Rotation (i,j,n,r)),j)))) . k = pk * ((Rotation (i,j,n,r)) * (k,j)) by A14, FVSUM_1:60 .= pk * (0. F_Real) by A1, A2, A3, A16, A17, A18, Def3 .= 0. F_Real ; ::_thesis: verum end; then A19: Sum (mlt ((@ p),(Col ((Rotation (i,j,n,r)),j)))) = ((mlt ((@ p),(Col ((Rotation (i,j,n,r)),j)))) /. i) + ((mlt ((@ p),(Col ((Rotation (i,j,n,r)),j)))) /. j) by A2, A4, A9, A11, MATRIX15:7; reconsider pii = (@ p) . i, pj = (@ p) . j as Element of F_Real ; A20: (mlt ((@ p),(Col ((Rotation (i,j,n,r)),j)))) /. i = (mlt ((@ p),(Col ((Rotation (i,j,n,r)),j)))) . i by A9, A11, PARTFUN1:def_6 .= pii * ((Rotation (i,j,n,r)) * (i,j)) by A10, A12, FVSUM_1:60 .= (p . i) * (sin r) by A1, A2, A3, Def3 ; (mlt ((@ p),(Col ((Rotation (i,j,n,r)),j)))) /. j = (mlt ((@ p),(Col ((Rotation (i,j,n,r)),j)))) . j by A4, A11, PARTFUN1:def_6 .= pj * ((Rotation (i,j,n,r)) * (j,j)) by A4, A7, A11, FVSUM_1:60 .= (p . j) * (cos r) by A1, A2, A3, Def3 ; hence (@ p) "*" (Col ((Rotation (i,j,n,r)),j)) = ((p . i) * (sin r)) + ((p . j) * (cos r)) by A19, A20; ::_thesis: verum end; theorem Th17: :: MATRTOP3:17 for r1, r2 being real number for i, j, n being Nat st 1 <= i & i < j & j <= n holds (Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2)) = Rotation (i,j,n,(r1 + r2)) proof let r1, r2 be real number ; ::_thesis: for i, j, n being Nat st 1 <= i & i < j & j <= n holds (Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2)) = Rotation (i,j,n,(r1 + r2)) let i, j, n be Nat; ::_thesis: ( 1 <= i & i < j & j <= n implies (Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2)) = Rotation (i,j,n,(r1 + r2)) ) assume that A1: 1 <= i and A2: i < j and A3: j <= n ; ::_thesis: (Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2)) = Rotation (i,j,n,(r1 + r2)) set S = Seg n; 1 <= j by A1, A2, XXREAL_0:2; then A4: j in Seg n by A3, FINSEQ_1:1; set O1 = Rotation (i,j,n,r1); set O2 = Rotation (i,j,n,r2); set O = Rotation (i,j,n,(r1 + r2)); set O12 = (Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2)); A5: width (Rotation (i,j,n,r1)) = n by MATRIX_1:24; i <= n by A2, A3, XXREAL_0:2; then A6: i in Seg n by A1, FINSEQ_1:1; A7: Indices (Rotation (i,j,n,r1)) = [:(Seg n),(Seg n):] by MATRIX_1:24; A8: Indices ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) = [:(Seg n),(Seg n):] by MATRIX_1:24; A9: Indices (Rotation (i,j,n,(r1 + r2))) = [:(Seg n),(Seg n):] by MATRIX_1:24; A10: len (Rotation (i,j,n,r2)) = n by MATRIX_1:25; for k, m being Nat st [k,m] in Indices ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) holds ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = (Rotation (i,j,n,(r1 + r2))) * (k,m) proof let k, m be Nat; ::_thesis: ( [k,m] in Indices ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) implies ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = (Rotation (i,j,n,(r1 + r2))) * (k,m) ) assume A11: [k,m] in Indices ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) ; ::_thesis: ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = (Rotation (i,j,n,(r1 + r2))) * (k,m) then A12: k in Seg n by A8, ZFMISC_1:87; A13: ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = (Line ((Rotation (i,j,n,r1)),k)) "*" (Col ((Rotation (i,j,n,r2)),m)) by A5, A10, A11, MATRIX_3:def_4; len (@ (Line ((Rotation (i,j,n,r1)),k))) = n by A5, CARD_1:def_7; then reconsider L = @ (Line ((Rotation (i,j,n,r1)),k)) as Element of (TOP-REAL n) by TOPREAL3:46; A14: m in Seg n by A8, A11, ZFMISC_1:87; then A15: L . m = (Rotation (i,j,n,r1)) * (k,m) by A5, MATRIX_1:def_7; A16: @ L = Line ((Rotation (i,j,n,r1)),k) ; A17: L . i = (Rotation (i,j,n,r1)) * (k,i) by A5, A6, MATRIX_1:def_7; A18: L . j = (Rotation (i,j,n,r1)) * (k,j) by A4, A5, MATRIX_1:def_7; percases ( m = i or m = j or ( m <> i & m <> j ) ) ; supposeA19: m = i ; ::_thesis: ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = (Rotation (i,j,n,(r1 + r2))) * (k,m) then A20: (Line ((Rotation (i,j,n,r1)),k)) "*" (Col ((Rotation (i,j,n,r2)),m)) = ((L . i) * (cos r2)) + ((L . j) * (- (sin r2))) by A1, A2, A3, A16, Th15; percases ( k = i or k = j or ( k <> j & k <> i ) ) ; supposeA21: k = i ; ::_thesis: ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = (Rotation (i,j,n,(r1 + r2))) * (k,m) hence ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = ((cos r1) * (cos r2)) + ((L . j) * (- (sin r2))) by A1, A2, A3, A13, A17, A20, Def3 .= ((cos r1) * (cos r2)) + ((sin r1) * (- (sin r2))) by A1, A2, A3, A18, A21, Def3 .= ((cos r1) * (cos r2)) - ((sin r1) * (sin r2)) .= cos (r1 + r2) by SIN_COS:75 .= (Rotation (i,j,n,(r1 + r2))) * (k,m) by A1, A2, A3, A19, A21, Def3 ; ::_thesis: verum end; supposeA22: k = j ; ::_thesis: ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = (Rotation (i,j,n,(r1 + r2))) * (k,m) hence ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = ((- (sin r1)) * (cos r2)) + ((L . j) * (- (sin r2))) by A1, A2, A3, A13, A17, A20, Def3 .= ((- (sin r1)) * (cos r2)) + ((cos r1) * (- (sin r2))) by A1, A2, A3, A18, A22, Def3 .= - (((sin r1) * (cos r2)) + ((cos r1) * (sin r2))) .= - (sin (r1 + r2)) by SIN_COS:75 .= (Rotation (i,j,n,(r1 + r2))) * (k,m) by A1, A2, A3, A19, A22, Def3 ; ::_thesis: verum end; suppose ( k <> j & k <> i ) ; ::_thesis: ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = (Rotation (i,j,n,(r1 + r2))) * (k,m) then not k in {i,j} by TARSKI:def_2; then A23: ( {k,i} <> {i,j} & {k,j} <> {i,j} ) by TARSKI:def_2; A24: [k,j] in [:(Seg n),(Seg n):] by A4, A12, ZFMISC_1:87; A25: [k,i] in [:(Seg n),(Seg n):] by A6, A12, ZFMISC_1:87; then (Rotation (i,j,n,r1)) * (k,i) = 0. F_Real by A1, A2, A3, A7, A23, Def3; hence ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = (0 * (cos r2)) + (0 * (- (sin r2))) by A1, A2, A3, A7, A13, A17, A18, A20, A23, A24, Def3 .= (Rotation (i,j,n,(r1 + r2))) * (k,m) by A1, A2, A3, A9, A19, A23, A25, Def3 ; ::_thesis: verum end; end; end; supposeA26: m = j ; ::_thesis: ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = (Rotation (i,j,n,(r1 + r2))) * (k,m) then A27: (Line ((Rotation (i,j,n,r1)),k)) "*" (Col ((Rotation (i,j,n,r2)),m)) = ((L . i) * (sin r2)) + ((L . j) * (cos r2)) by A1, A2, A3, A16, Th16; percases ( k = i or k = j or ( k <> j & k <> i ) ) ; supposeA28: k = i ; ::_thesis: ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = (Rotation (i,j,n,(r1 + r2))) * (k,m) hence ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = ((cos r1) * (sin r2)) + ((L . j) * (cos r2)) by A1, A2, A3, A13, A17, A27, Def3 .= ((cos r1) * (sin r2)) + ((sin r1) * (cos r2)) by A1, A2, A3, A18, A28, Def3 .= sin (r1 + r2) by SIN_COS:75 .= (Rotation (i,j,n,(r1 + r2))) * (k,m) by A1, A2, A3, A26, A28, Def3 ; ::_thesis: verum end; supposeA29: k = j ; ::_thesis: ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = (Rotation (i,j,n,(r1 + r2))) * (k,m) hence ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = ((- (sin r1)) * (sin r2)) + ((L . j) * (cos r2)) by A1, A2, A3, A13, A17, A27, Def3 .= ((cos r1) * (cos r2)) - ((sin r1) * (sin r2)) by A1, A2, A3, A18, A29, Def3 .= cos (r1 + r2) by SIN_COS:75 .= (Rotation (i,j,n,(r1 + r2))) * (k,m) by A1, A2, A3, A26, A29, Def3 ; ::_thesis: verum end; suppose ( k <> j & k <> i ) ; ::_thesis: ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = (Rotation (i,j,n,(r1 + r2))) * (k,m) then not k in {i,j} by TARSKI:def_2; then A30: ( {k,i} <> {i,j} & {k,j} <> {i,j} ) by TARSKI:def_2; A31: [k,j] in [:(Seg n),(Seg n):] by A4, A12, ZFMISC_1:87; [k,i] in [:(Seg n),(Seg n):] by A6, A12, ZFMISC_1:87; then (Rotation (i,j,n,r1)) * (k,i) = 0. F_Real by A1, A2, A3, A7, A30, Def3; hence ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = (0 * (sin r2)) + (0 * (cos r2)) by A1, A2, A3, A7, A13, A17, A18, A27, A30, A31, Def3 .= (Rotation (i,j,n,(r1 + r2))) * (k,m) by A1, A2, A3, A9, A26, A30, A31, Def3 ; ::_thesis: verum end; end; end; supposeA32: ( m <> i & m <> j ) ; ::_thesis: ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = (Rotation (i,j,n,(r1 + r2))) * (k,m) then A33: ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = L . m by A1, A2, A3, A13, A14, A16, Th14; A34: [k,m] in [:(Seg n),(Seg n):] by A11, MATRIX_1:24; percases ( k = m or k <> m ) ; supposeA35: k = m ; ::_thesis: ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = (Rotation (i,j,n,(r1 + r2))) * (k,m) then (Rotation (i,j,n,r1)) * (k,m) = 1. F_Real by A1, A2, A3, A7, A8, A11, A32, Def3; hence ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = (Rotation (i,j,n,(r1 + r2))) * (k,m) by A1, A2, A3, A9, A15, A32, A33, A34, A35, Def3; ::_thesis: verum end; supposeA36: k <> m ; ::_thesis: ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = (Rotation (i,j,n,(r1 + r2))) * (k,m) not m in {i,j} by A32, TARSKI:def_2; then A37: {k,m} <> {i,j} by TARSKI:def_2; then (Rotation (i,j,n,r1)) * (k,m) = 0. F_Real by A1, A2, A3, A7, A8, A11, A36, Def3; hence ((Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2))) * (k,m) = (Rotation (i,j,n,(r1 + r2))) * (k,m) by A1, A2, A3, A9, A15, A33, A34, A36, A37, Def3; ::_thesis: verum end; end; end; end; end; hence (Rotation (i,j,n,r1)) * (Rotation (i,j,n,r2)) = Rotation (i,j,n,(r1 + r2)) by MATRIX_1:27; ::_thesis: verum end; Lm4: for r being real number for i, j, n being Nat st 1 <= i & i < j & j <= n holds (Rotation (i,j,n,r)) @ = Rotation (i,j,n,(- r)) proof let r be real number ; ::_thesis: for i, j, n being Nat st 1 <= i & i < j & j <= n holds (Rotation (i,j,n,r)) @ = Rotation (i,j,n,(- r)) let i, j, n be Nat; ::_thesis: ( 1 <= i & i < j & j <= n implies (Rotation (i,j,n,r)) @ = Rotation (i,j,n,(- r)) ) set O1 = Rotation (i,j,n,r); set O2 = Rotation (i,j,n,(- r)); set S = Seg n; assume A1: ( 1 <= i & i < j & j <= n ) ; ::_thesis: (Rotation (i,j,n,r)) @ = Rotation (i,j,n,(- r)) A2: Indices (Rotation (i,j,n,(- r))) = [:(Seg n),(Seg n):] by MATRIX_1:24; A3: Indices (Rotation (i,j,n,r)) = [:(Seg n),(Seg n):] by MATRIX_1:24; A4: Indices ((Rotation (i,j,n,r)) @) = [:(Seg n),(Seg n):] by MATRIX_1:24; for k, m being Nat st [k,m] in [:(Seg n),(Seg n):] holds ((Rotation (i,j,n,r)) @) * (k,m) = (Rotation (i,j,n,(- r))) * (k,m) proof A5: cos r = cos (- r) by SIN_COS:31; let k, m be Nat; ::_thesis: ( [k,m] in [:(Seg n),(Seg n):] implies ((Rotation (i,j,n,r)) @) * (k,m) = (Rotation (i,j,n,(- r))) * (k,m) ) A6: - (sin r) = sin (- r) by SIN_COS:31; assume A7: [k,m] in [:(Seg n),(Seg n):] ; ::_thesis: ((Rotation (i,j,n,r)) @) * (k,m) = (Rotation (i,j,n,(- r))) * (k,m) then A8: [m,k] in [:(Seg n),(Seg n):] by A3, A4, MATRIX_1:def_6; then A9: ((Rotation (i,j,n,r)) @) * (k,m) = (Rotation (i,j,n,r)) * (m,k) by A3, MATRIX_1:def_6; percases ( ( k = m & k = i ) or ( k = m & k = j ) or ( k = m & k <> i & k <> j ) or ( k <> m & k = i & m = j ) or ( k <> m & k = i & m <> j ) or ( k <> m & k = j & m = i ) or ( k <> m & k = j & m <> i ) or ( k <> m & k <> j & k <> i ) or ( k <> m & m <> j & m <> i ) ) ; supposeA10: ( k = m & k = i ) ; ::_thesis: ((Rotation (i,j,n,r)) @) * (k,m) = (Rotation (i,j,n,(- r))) * (k,m) then (Rotation (i,j,n,r)) * (m,k) = cos r by A1, Def3; hence ((Rotation (i,j,n,r)) @) * (k,m) = (Rotation (i,j,n,(- r))) * (k,m) by A1, A5, A9, A10, Def3; ::_thesis: verum end; supposeA11: ( k = m & k = j ) ; ::_thesis: ((Rotation (i,j,n,r)) @) * (k,m) = (Rotation (i,j,n,(- r))) * (k,m) then (Rotation (i,j,n,r)) * (m,k) = cos r by A1, Def3; hence ((Rotation (i,j,n,r)) @) * (k,m) = (Rotation (i,j,n,(- r))) * (k,m) by A1, A5, A9, A11, Def3; ::_thesis: verum end; supposeA12: ( k = m & k <> i & k <> j ) ; ::_thesis: ((Rotation (i,j,n,r)) @) * (k,m) = (Rotation (i,j,n,(- r))) * (k,m) then (Rotation (i,j,n,r)) * (m,k) = 1. F_Real by A1, A3, A7, Def3; hence ((Rotation (i,j,n,r)) @) * (k,m) = (Rotation (i,j,n,(- r))) * (k,m) by A1, A2, A7, A9, A12, Def3; ::_thesis: verum end; supposeA13: ( k <> m & k = i & m = j ) ; ::_thesis: ((Rotation (i,j,n,r)) @) * (k,m) = (Rotation (i,j,n,(- r))) * (k,m) then (Rotation (i,j,n,r)) * (m,k) = - (sin r) by A1, Def3; hence ((Rotation (i,j,n,r)) @) * (k,m) = (Rotation (i,j,n,(- r))) * (k,m) by A1, A6, A9, A13, Def3; ::_thesis: verum end; supposeA14: ( k <> m & k = i & m <> j ) ; ::_thesis: ((Rotation (i,j,n,r)) @) * (k,m) = (Rotation (i,j,n,(- r))) * (k,m) then not m in {i,j} by TARSKI:def_2; then A15: {m,k} <> {i,j} by TARSKI:def_2; then (Rotation (i,j,n,r)) * (m,k) = 0. F_Real by A1, A3, A8, A14, Def3; hence ((Rotation (i,j,n,r)) @) * (k,m) = (Rotation (i,j,n,(- r))) * (k,m) by A1, A2, A7, A9, A14, A15, Def3; ::_thesis: verum end; supposeA16: ( k <> m & k = j & m = i ) ; ::_thesis: ((Rotation (i,j,n,r)) @) * (k,m) = (Rotation (i,j,n,(- r))) * (k,m) then (Rotation (i,j,n,(- r))) * (k,m) = - (sin (- r)) by A1, Def3; hence ((Rotation (i,j,n,r)) @) * (k,m) = (Rotation (i,j,n,(- r))) * (k,m) by A1, A6, A9, A16, Def3; ::_thesis: verum end; supposeA17: ( k <> m & k = j & m <> i ) ; ::_thesis: ((Rotation (i,j,n,r)) @) * (k,m) = (Rotation (i,j,n,(- r))) * (k,m) then not m in {i,j} by TARSKI:def_2; then A18: {m,k} <> {i,j} by TARSKI:def_2; then (Rotation (i,j,n,r)) * (m,k) = 0. F_Real by A1, A3, A8, A17, Def3; hence ((Rotation (i,j,n,r)) @) * (k,m) = (Rotation (i,j,n,(- r))) * (k,m) by A1, A2, A7, A9, A17, A18, Def3; ::_thesis: verum end; supposeA19: ( k <> m & k <> j & k <> i ) ; ::_thesis: ((Rotation (i,j,n,r)) @) * (k,m) = (Rotation (i,j,n,(- r))) * (k,m) then not k in {i,j} by TARSKI:def_2; then A20: {m,k} <> {i,j} by TARSKI:def_2; then (Rotation (i,j,n,r)) * (m,k) = 0. F_Real by A1, A3, A8, A19, Def3; hence ((Rotation (i,j,n,r)) @) * (k,m) = (Rotation (i,j,n,(- r))) * (k,m) by A1, A2, A7, A9, A19, A20, Def3; ::_thesis: verum end; supposeA21: ( k <> m & m <> j & m <> i ) ; ::_thesis: ((Rotation (i,j,n,r)) @) * (k,m) = (Rotation (i,j,n,(- r))) * (k,m) then not m in {i,j} by TARSKI:def_2; then A22: {m,k} <> {i,j} by TARSKI:def_2; then (Rotation (i,j,n,r)) * (m,k) = 0. F_Real by A1, A3, A8, A21, Def3; hence ((Rotation (i,j,n,r)) @) * (k,m) = (Rotation (i,j,n,(- r))) * (k,m) by A1, A2, A7, A9, A21, A22, Def3; ::_thesis: verum end; end; end; hence (Rotation (i,j,n,r)) @ = Rotation (i,j,n,(- r)) by A4, MATRIX_1:27; ::_thesis: verum end; theorem Th18: :: MATRTOP3:18 for i, j, n being Nat st 1 <= i & i < j & j <= n holds Rotation (i,j,n,0) = 1. (F_Real,n) proof let i, j, n be Nat; ::_thesis: ( 1 <= i & i < j & j <= n implies Rotation (i,j,n,0) = 1. (F_Real,n) ) set O = Rotation (i,j,n,0); assume A1: ( 1 <= i & i < j & j <= n ) ; ::_thesis: Rotation (i,j,n,0) = 1. (F_Real,n) A2: for k, m being Nat st [k,m] in Indices (Rotation (i,j,n,0)) & k <> m holds (Rotation (i,j,n,0)) * (k,m) = 0. F_Real proof let k, m be Nat; ::_thesis: ( [k,m] in Indices (Rotation (i,j,n,0)) & k <> m implies (Rotation (i,j,n,0)) * (k,m) = 0. F_Real ) assume that A3: [k,m] in Indices (Rotation (i,j,n,0)) and A4: k <> m ; ::_thesis: (Rotation (i,j,n,0)) * (k,m) = 0. F_Real percases ( ( k = i & m = j ) or ( k = j & m = i ) or ( k = i & m <> j ) or ( k = j & m <> i ) or ( m = i & k <> j ) or ( m = j & k <> i ) or ( k <> i & k <> j & m <> i & m <> j ) ) ; suppose ( ( k = i & m = j ) or ( k = j & m = i ) ) ; ::_thesis: (Rotation (i,j,n,0)) * (k,m) = 0. F_Real then ( (Rotation (i,j,n,0)) * (k,m) = sin 0 or (Rotation (i,j,n,0)) * (k,m) = - (sin 0) ) by A1, Def3; hence (Rotation (i,j,n,0)) * (k,m) = 0. F_Real by SIN_COS:31; ::_thesis: verum end; suppose ( k = i & m <> j ) ; ::_thesis: (Rotation (i,j,n,0)) * (k,m) = 0. F_Real then not m in {i,j} by A4, TARSKI:def_2; then {k,m} <> {i,j} by TARSKI:def_2; hence (Rotation (i,j,n,0)) * (k,m) = 0. F_Real by A1, A3, A4, Def3; ::_thesis: verum end; suppose ( k = j & m <> i ) ; ::_thesis: (Rotation (i,j,n,0)) * (k,m) = 0. F_Real then not m in {i,j} by A4, TARSKI:def_2; then {k,m} <> {i,j} by TARSKI:def_2; hence (Rotation (i,j,n,0)) * (k,m) = 0. F_Real by A1, A3, A4, Def3; ::_thesis: verum end; suppose ( m = i & k <> j ) ; ::_thesis: (Rotation (i,j,n,0)) * (k,m) = 0. F_Real then not k in {i,j} by A4, TARSKI:def_2; then {k,m} <> {i,j} by TARSKI:def_2; hence (Rotation (i,j,n,0)) * (k,m) = 0. F_Real by A1, A3, A4, Def3; ::_thesis: verum end; suppose ( m = j & k <> i ) ; ::_thesis: (Rotation (i,j,n,0)) * (k,m) = 0. F_Real then not k in {i,j} by A4, TARSKI:def_2; then {k,m} <> {i,j} by TARSKI:def_2; hence (Rotation (i,j,n,0)) * (k,m) = 0. F_Real by A1, A3, A4, Def3; ::_thesis: verum end; suppose ( k <> i & k <> j & m <> i & m <> j ) ; ::_thesis: (Rotation (i,j,n,0)) * (k,m) = 0. F_Real then not m in {i,j} by TARSKI:def_2; then {k,m} <> {i,j} by TARSKI:def_2; hence (Rotation (i,j,n,0)) * (k,m) = 0. F_Real by A1, A3, A4, Def3; ::_thesis: verum end; end; end; for k being Nat st [k,k] in Indices (Rotation (i,j,n,0)) holds (Rotation (i,j,n,0)) * (k,k) = 1. F_Real proof let k be Nat; ::_thesis: ( [k,k] in Indices (Rotation (i,j,n,0)) implies (Rotation (i,j,n,0)) * (k,k) = 1. F_Real ) assume A5: [k,k] in Indices (Rotation (i,j,n,0)) ; ::_thesis: (Rotation (i,j,n,0)) * (k,k) = 1. F_Real ( k = i or k = j or ( k <> i & k <> j ) ) ; hence (Rotation (i,j,n,0)) * (k,k) = 1. F_Real by A1, A5, Def3, SIN_COS:31; ::_thesis: verum end; hence Rotation (i,j,n,0) = 1. (F_Real,n) by A2, MATRIX_1:def_11; ::_thesis: verum end; Lm5: for r being real number for i, j, n being Nat st 1 <= i & i < j & j <= n holds (Rotation (i,j,n,r)) ~ = Rotation (i,j,n,(- r)) proof let r be real number ; ::_thesis: for i, j, n being Nat st 1 <= i & i < j & j <= n holds (Rotation (i,j,n,r)) ~ = Rotation (i,j,n,(- r)) let i, j, n be Nat; ::_thesis: ( 1 <= i & i < j & j <= n implies (Rotation (i,j,n,r)) ~ = Rotation (i,j,n,(- r)) ) assume A1: ( 1 <= i & i < j & j <= n ) ; ::_thesis: (Rotation (i,j,n,r)) ~ = Rotation (i,j,n,(- r)) then A2: (Rotation (i,j,n,r)) * (Rotation (i,j,n,(- r))) = Rotation (i,j,n,(r + (- r))) by Th17 .= 1. (F_Real,n) by A1, Th18 ; (Rotation (i,j,n,(- r))) * (Rotation (i,j,n,r)) = Rotation (i,j,n,((- r) + r)) by A1, Th17 .= 1. (F_Real,n) by A1, Th18 ; then Rotation (i,j,n,r) is_reverse_of Rotation (i,j,n,(- r)) by A2, MATRIX_6:def_2; hence (Rotation (i,j,n,r)) ~ = Rotation (i,j,n,(- r)) by MATRIX_6:def_4; ::_thesis: verum end; theorem Th19: :: MATRTOP3:19 for r being real number for i, j, n being Nat st 1 <= i & i < j & j <= n holds ( Rotation (i,j,n,r) is Orthogonal & (Rotation (i,j,n,r)) ~ = Rotation (i,j,n,(- r)) ) proof let r be real number ; ::_thesis: for i, j, n being Nat st 1 <= i & i < j & j <= n holds ( Rotation (i,j,n,r) is Orthogonal & (Rotation (i,j,n,r)) ~ = Rotation (i,j,n,(- r)) ) let i, j, n be Nat; ::_thesis: ( 1 <= i & i < j & j <= n implies ( Rotation (i,j,n,r) is Orthogonal & (Rotation (i,j,n,r)) ~ = Rotation (i,j,n,(- r)) ) ) assume ( 1 <= i & i < j & j <= n ) ; ::_thesis: ( Rotation (i,j,n,r) is Orthogonal & (Rotation (i,j,n,r)) ~ = Rotation (i,j,n,(- r)) ) then ( (Rotation (i,j,n,r)) ~ = Rotation (i,j,n,(- r)) & (Rotation (i,j,n,r)) @ = Rotation (i,j,n,(- r)) ) by Lm4, Lm5; hence ( Rotation (i,j,n,r) is Orthogonal & (Rotation (i,j,n,r)) ~ = Rotation (i,j,n,(- r)) ) by MATRIX_6:def_7; ::_thesis: verum end; theorem Th20: :: MATRTOP3:20 for r being real number for i, j, n, k being Nat for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n & k <> i & k <> j holds ((Mx2Tran (Rotation (i,j,n,r))) . p) . k = p . k proof let r be real number ; ::_thesis: for i, j, n, k being Nat for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n & k <> i & k <> j holds ((Mx2Tran (Rotation (i,j,n,r))) . p) . k = p . k let i, j, n, k be Nat; ::_thesis: for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n & k <> i & k <> j holds ((Mx2Tran (Rotation (i,j,n,r))) . p) . k = p . k let p be Point of (TOP-REAL n); ::_thesis: ( 1 <= i & i < j & j <= n & k <> i & k <> j implies ((Mx2Tran (Rotation (i,j,n,r))) . p) . k = p . k ) set O = Rotation (i,j,n,r); set M = Mx2Tran (Rotation (i,j,n,r)); set Mp = (Mx2Tran (Rotation (i,j,n,r))) . p; set S = Seg n; assume A1: ( 1 <= i & i < j & j <= n & k <> i & k <> j ) ; ::_thesis: ((Mx2Tran (Rotation (i,j,n,r))) . p) . k = p . k len ((Mx2Tran (Rotation (i,j,n,r))) . p) = n by CARD_1:def_7; then A2: dom ((Mx2Tran (Rotation (i,j,n,r))) . p) = Seg n by FINSEQ_1:def_3; len p = n by CARD_1:def_7; then A3: dom p = Seg n by FINSEQ_1:def_3; percases ( k in Seg n or not k in Seg n ) ; supposeA4: k in Seg n ; ::_thesis: ((Mx2Tran (Rotation (i,j,n,r))) . p) . k = p . k then ( 1 <= k & k <= n ) by FINSEQ_1:1; hence ((Mx2Tran (Rotation (i,j,n,r))) . p) . k = (@ p) "*" (Col ((Rotation (i,j,n,r)),k)) by MATRTOP1:18 .= p . k by A1, A4, Th14 ; ::_thesis: verum end; supposeA5: not k in Seg n ; ::_thesis: ((Mx2Tran (Rotation (i,j,n,r))) . p) . k = p . k then ((Mx2Tran (Rotation (i,j,n,r))) . p) . k = {} by A2, FUNCT_1:def_2; hence ((Mx2Tran (Rotation (i,j,n,r))) . p) . k = p . k by A3, A5, FUNCT_1:def_2; ::_thesis: verum end; end; end; theorem Th21: :: MATRTOP3:21 for r being real number for i, j, n being Nat for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n holds ((Mx2Tran (Rotation (i,j,n,r))) . p) . i = ((p . i) * (cos r)) + ((p . j) * (- (sin r))) proof let r be real number ; ::_thesis: for i, j, n being Nat for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n holds ((Mx2Tran (Rotation (i,j,n,r))) . p) . i = ((p . i) * (cos r)) + ((p . j) * (- (sin r))) let i, j, n be Nat; ::_thesis: for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n holds ((Mx2Tran (Rotation (i,j,n,r))) . p) . i = ((p . i) * (cos r)) + ((p . j) * (- (sin r))) let p be Point of (TOP-REAL n); ::_thesis: ( 1 <= i & i < j & j <= n implies ((Mx2Tran (Rotation (i,j,n,r))) . p) . i = ((p . i) * (cos r)) + ((p . j) * (- (sin r))) ) set O = Rotation (i,j,n,r); set M = Mx2Tran (Rotation (i,j,n,r)); set Mp = (Mx2Tran (Rotation (i,j,n,r))) . p; set S = Seg n; assume that A1: 1 <= i and A2: ( i < j & j <= n ) ; ::_thesis: ((Mx2Tran (Rotation (i,j,n,r))) . p) . i = ((p . i) * (cos r)) + ((p . j) * (- (sin r))) i <= n by A2, XXREAL_0:2; hence ((Mx2Tran (Rotation (i,j,n,r))) . p) . i = (@ p) "*" (Col ((Rotation (i,j,n,r)),i)) by A1, MATRTOP1:18 .= ((p . i) * (cos r)) + ((p . j) * (- (sin r))) by A1, A2, Th15 ; ::_thesis: verum end; theorem Th22: :: MATRTOP3:22 for r being real number for i, j, n being Nat for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n holds ((Mx2Tran (Rotation (i,j,n,r))) . p) . j = ((p . i) * (sin r)) + ((p . j) * (cos r)) proof let r be real number ; ::_thesis: for i, j, n being Nat for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n holds ((Mx2Tran (Rotation (i,j,n,r))) . p) . j = ((p . i) * (sin r)) + ((p . j) * (cos r)) let i, j, n be Nat; ::_thesis: for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n holds ((Mx2Tran (Rotation (i,j,n,r))) . p) . j = ((p . i) * (sin r)) + ((p . j) * (cos r)) let p be Point of (TOP-REAL n); ::_thesis: ( 1 <= i & i < j & j <= n implies ((Mx2Tran (Rotation (i,j,n,r))) . p) . j = ((p . i) * (sin r)) + ((p . j) * (cos r)) ) set O = Rotation (i,j,n,r); set M = Mx2Tran (Rotation (i,j,n,r)); set Mp = (Mx2Tran (Rotation (i,j,n,r))) . p; set S = Seg n; assume that A1: ( 1 <= i & i < j ) and A2: j <= n ; ::_thesis: ((Mx2Tran (Rotation (i,j,n,r))) . p) . j = ((p . i) * (sin r)) + ((p . j) * (cos r)) 1 <= j by A1, XXREAL_0:2; hence ((Mx2Tran (Rotation (i,j,n,r))) . p) . j = (@ p) "*" (Col ((Rotation (i,j,n,r)),j)) by A2, MATRTOP1:18 .= ((p . i) * (sin r)) + ((p . j) * (cos r)) by A1, A2, Th16 ; ::_thesis: verum end; theorem Th23: :: MATRTOP3:23 for r being real number for i, j, n being Nat for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n holds (Mx2Tran (Rotation (i,j,n,r))) . p = ((((p | (i -' 1)) ^ <*(((p . i) * (cos r)) + ((p . j) * (- (sin r))))*>) ^ ((p /^ i) | ((j -' i) -' 1))) ^ <*(((p . i) * (sin r)) + ((p . j) * (cos r)))*>) ^ (p /^ j) proof let r be real number ; ::_thesis: for i, j, n being Nat for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n holds (Mx2Tran (Rotation (i,j,n,r))) . p = ((((p | (i -' 1)) ^ <*(((p . i) * (cos r)) + ((p . j) * (- (sin r))))*>) ^ ((p /^ i) | ((j -' i) -' 1))) ^ <*(((p . i) * (sin r)) + ((p . j) * (cos r)))*>) ^ (p /^ j) let i, j, n be Nat; ::_thesis: for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n holds (Mx2Tran (Rotation (i,j,n,r))) . p = ((((p | (i -' 1)) ^ <*(((p . i) * (cos r)) + ((p . j) * (- (sin r))))*>) ^ ((p /^ i) | ((j -' i) -' 1))) ^ <*(((p . i) * (sin r)) + ((p . j) * (cos r)))*>) ^ (p /^ j) let p be Point of (TOP-REAL n); ::_thesis: ( 1 <= i & i < j & j <= n implies (Mx2Tran (Rotation (i,j,n,r))) . p = ((((p | (i -' 1)) ^ <*(((p . i) * (cos r)) + ((p . j) * (- (sin r))))*>) ^ ((p /^ i) | ((j -' i) -' 1))) ^ <*(((p . i) * (sin r)) + ((p . j) * (cos r)))*>) ^ (p /^ j) ) assume that A1: 1 <= i and A2: i < j and A3: j <= n ; ::_thesis: (Mx2Tran (Rotation (i,j,n,r))) . p = ((((p | (i -' 1)) ^ <*(((p . i) * (cos r)) + ((p . j) * (- (sin r))))*>) ^ ((p /^ i) | ((j -' i) -' 1))) ^ <*(((p . i) * (sin r)) + ((p . j) * (cos r)))*>) ^ (p /^ j) set M = Mx2Tran (Rotation (i,j,n,r)); set Mp = (Mx2Tran (Rotation (i,j,n,r))) . p; set i1 = i -' 1; set ji = j -' i; A4: i < n by A2, A3, XXREAL_0:2; A5: ( i -' 1 < (i -' 1) + 1 & i -' 1 = i - 1 ) by A1, NAT_1:13, XREAL_1:233; A6: now__::_thesis:_for_k_being_Nat_st_1_<=_k_&_k_<=_i_-'_1_holds_ (((Mx2Tran_(Rotation_(i,j,n,r)))_._p)_|_(i_-'_1))_._k_=_(p_|_(i_-'_1))_._k let k be Nat; ::_thesis: ( 1 <= k & k <= i -' 1 implies (((Mx2Tran (Rotation (i,j,n,r))) . p) | (i -' 1)) . k = (p | (i -' 1)) . k ) assume that 1 <= k and A7: k <= i -' 1 ; ::_thesis: (((Mx2Tran (Rotation (i,j,n,r))) . p) | (i -' 1)) . k = (p | (i -' 1)) . k A8: ( (((Mx2Tran (Rotation (i,j,n,r))) . p) | (i -' 1)) . k = ((Mx2Tran (Rotation (i,j,n,r))) . p) . k & (p | (i -' 1)) . k = p . k ) by A7, FINSEQ_3:112; k < i by A5, A7, XXREAL_0:2; hence (((Mx2Tran (Rotation (i,j,n,r))) . p) | (i -' 1)) . k = (p | (i -' 1)) . k by A1, A2, A3, A8, Th20; ::_thesis: verum end; A9: len ((Mx2Tran (Rotation (i,j,n,r))) . p) = n by CARD_1:def_7; then A10: len (((Mx2Tran (Rotation (i,j,n,r))) . p) | (i -' 1)) = i -' 1 by A5, A4, FINSEQ_1:59, XXREAL_0:2; A11: len p = n by CARD_1:def_7; then A12: len (p /^ i) = n - i by A4, RFINSEQ:def_1; A13: j -' i = j - i by A2, XREAL_1:233; then A14: ( (j -' i) -' 1 < ((j -' i) -' 1) + 1 & j -' i <= n - i ) by A3, NAT_1:13, XREAL_1:6; j - i > i - i by A2, XREAL_1:8; then A15: (j -' i) -' 1 = (j -' i) - 1 by A13, NAT_1:14, XREAL_1:233; A16: len (p /^ j) = n - j by A3, A11, RFINSEQ:def_1; A17: len (((Mx2Tran (Rotation (i,j,n,r))) . p) /^ i) = n - i by A9, A4, RFINSEQ:def_1; then A18: len ((((Mx2Tran (Rotation (i,j,n,r))) . p) /^ i) | ((j -' i) -' 1)) = (j -' i) -' 1 by A14, A15, FINSEQ_1:59, XXREAL_0:2; A19: (j -' i) -' 1 < n - i by A14, A15, XXREAL_0:2; A20: now__::_thesis:_for_k_being_Nat_st_1_<=_k_&_k_<=_(j_-'_i)_-'_1_holds_ ((((Mx2Tran_(Rotation_(i,j,n,r)))_._p)_/^_i)_|_((j_-'_i)_-'_1))_._k_=_((p_/^_i)_|_((j_-'_i)_-'_1))_._k let k be Nat; ::_thesis: ( 1 <= k & k <= (j -' i) -' 1 implies ((((Mx2Tran (Rotation (i,j,n,r))) . p) /^ i) | ((j -' i) -' 1)) . k = ((p /^ i) | ((j -' i) -' 1)) . k ) assume that A21: 1 <= k and A22: k <= (j -' i) -' 1 ; ::_thesis: ((((Mx2Tran (Rotation (i,j,n,r))) . p) /^ i) | ((j -' i) -' 1)) . k = ((p /^ i) | ((j -' i) -' 1)) . k A23: ((p /^ i) | ((j -' i) -' 1)) . k = (p /^ i) . k by A22, FINSEQ_3:112; A24: k <= n - i by A19, A22, XXREAL_0:2; then k in dom (((Mx2Tran (Rotation (i,j,n,r))) . p) /^ i) by A17, A21, FINSEQ_3:25; then A25: (((Mx2Tran (Rotation (i,j,n,r))) . p) /^ i) . k = ((Mx2Tran (Rotation (i,j,n,r))) . p) . (i + k) by A9, A4, RFINSEQ:def_1; k < ((j -' i) -' 1) + 1 by A22, NAT_1:13; then A26: i + k < i + (j -' i) by A15, XREAL_1:8; k in dom (p /^ i) by A12, A21, A24, FINSEQ_3:25; then A27: (p /^ i) . k = p . (i + k) by A11, A4, RFINSEQ:def_1; ( k + i <> i & ((((Mx2Tran (Rotation (i,j,n,r))) . p) /^ i) | ((j -' i) -' 1)) . k = (((Mx2Tran (Rotation (i,j,n,r))) . p) /^ i) . k ) by A21, A22, FINSEQ_3:112, NAT_1:14; hence ((((Mx2Tran (Rotation (i,j,n,r))) . p) /^ i) | ((j -' i) -' 1)) . k = ((p /^ i) | ((j -' i) -' 1)) . k by A1, A2, A3, A13, A25, A26, A27, A23, Th20; ::_thesis: verum end; len ((p /^ i) | ((j -' i) -' 1)) = (j -' i) -' 1 by A14, A15, A12, FINSEQ_1:59, XXREAL_0:2; then A28: (((Mx2Tran (Rotation (i,j,n,r))) . p) /^ i) | ((j -' i) -' 1) = (p /^ i) | ((j -' i) -' 1) by A20, A18, FINSEQ_1:14; A29: len (((Mx2Tran (Rotation (i,j,n,r))) . p) /^ j) = n - j by A3, A9, RFINSEQ:def_1; now__::_thesis:_for_k_being_Nat_st_1_<=_k_&_k_<=_n_-_j_holds_ (((Mx2Tran_(Rotation_(i,j,n,r)))_._p)_/^_j)_._k_=_(p_/^_j)_._k let k be Nat; ::_thesis: ( 1 <= k & k <= n - j implies (((Mx2Tran (Rotation (i,j,n,r))) . p) /^ j) . k = (p /^ j) . k ) assume that A30: 1 <= k and A31: k <= n - j ; ::_thesis: (((Mx2Tran (Rotation (i,j,n,r))) . p) /^ j) . k = (p /^ j) . k k in dom (((Mx2Tran (Rotation (i,j,n,r))) . p) /^ j) by A29, A30, A31, FINSEQ_3:25; then A32: (((Mx2Tran (Rotation (i,j,n,r))) . p) /^ j) . k = ((Mx2Tran (Rotation (i,j,n,r))) . p) . (j + k) by A3, A9, RFINSEQ:def_1; k in dom (p /^ j) by A16, A30, A31, FINSEQ_3:25; then A33: (p /^ j) . k = p . (j + k) by A3, A11, RFINSEQ:def_1; ( j + k >= j & j + k <> j ) by A30, NAT_1:11, NAT_1:14; hence (((Mx2Tran (Rotation (i,j,n,r))) . p) /^ j) . k = (p /^ j) . k by A1, A2, A3, A32, A33, Th20; ::_thesis: verum end; then A34: ((Mx2Tran (Rotation (i,j,n,r))) . p) /^ j = p /^ j by A16, A29, FINSEQ_1:14; len (p | (i -' 1)) = i -' 1 by A5, A11, A4, FINSEQ_1:59, XXREAL_0:2; then A35: ((Mx2Tran (Rotation (i,j,n,r))) . p) | (i -' 1) = p | (i -' 1) by A6, A10, FINSEQ_1:14; A36: ((Mx2Tran (Rotation (i,j,n,r))) . p) . i = ((p . i) * (cos r)) + ((p . j) * (- (sin r))) by A1, A2, A3, Th21; (Mx2Tran (Rotation (i,j,n,r))) . p = @ (@ ((Mx2Tran (Rotation (i,j,n,r))) . p)) ; then (Mx2Tran (Rotation (i,j,n,r))) . p = ((((((Mx2Tran (Rotation (i,j,n,r))) . p) | (i -' 1)) ^ <*(((Mx2Tran (Rotation (i,j,n,r))) . p) . i)*>) ^ ((((Mx2Tran (Rotation (i,j,n,r))) . p) /^ i) | ((j -' i) -' 1))) ^ <*(((Mx2Tran (Rotation (i,j,n,r))) . p) . j)*>) ^ (((Mx2Tran (Rotation (i,j,n,r))) . p) /^ j) by A1, A2, A3, A9, FINSEQ_7:1; hence (Mx2Tran (Rotation (i,j,n,r))) . p = ((((p | (i -' 1)) ^ <*(((p . i) * (cos r)) + ((p . j) * (- (sin r))))*>) ^ ((p /^ i) | ((j -' i) -' 1))) ^ <*(((p . i) * (sin r)) + ((p . j) * (cos r)))*>) ^ (p /^ j) by A1, A2, A3, A34, A28, A35, A36, Th22; ::_thesis: verum end; theorem Th24: :: MATRTOP3:24 for s being real number for i, j, n being Nat for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n & s ^2 <= ((p . i) ^2) + ((p . j) ^2) holds ex r being real number st ((Mx2Tran (Rotation (i,j,n,r))) . p) . i = s proof let s be real number ; ::_thesis: for i, j, n being Nat for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n & s ^2 <= ((p . i) ^2) + ((p . j) ^2) holds ex r being real number st ((Mx2Tran (Rotation (i,j,n,r))) . p) . i = s let i, j, n be Nat; ::_thesis: for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n & s ^2 <= ((p . i) ^2) + ((p . j) ^2) holds ex r being real number st ((Mx2Tran (Rotation (i,j,n,r))) . p) . i = s let p be Point of (TOP-REAL n); ::_thesis: ( 1 <= i & i < j & j <= n & s ^2 <= ((p . i) ^2) + ((p . j) ^2) implies ex r being real number st ((Mx2Tran (Rotation (i,j,n,r))) . p) . i = s ) set pk = p . i; set pj = p . j; set pkj = ((p . i) ^2) + ((p . j) ^2); set P = sqrt (((p . i) ^2) + ((p . j) ^2)); assume that A1: ( 1 <= i & i < j & j <= n ) and A2: s ^2 <= ((p . i) ^2) + ((p . j) ^2) ; ::_thesis: ex r being real number st ((Mx2Tran (Rotation (i,j,n,r))) . p) . i = s A3: ( 0 <= (p . i) * (p . i) & 0 <= (p . j) * (p . j) ) by XREAL_1:63; percases ( p . i <> 0 or p . j <> 0 or ( p . i = 0 & p . j = 0 ) ) ; supposeA4: ( p . i <> 0 or p . j <> 0 ) ; ::_thesis: ex r being real number st ((Mx2Tran (Rotation (i,j,n,r))) . p) . i = s A5: 0 <= (p . i) * (p . i) by XREAL_1:63; A6: sqrt (((p . i) ^2) + ((p . j) ^2)) > 0 by A3, A4, SQUARE_1:25; then A7: (s / (sqrt (((p . i) ^2) + ((p . j) ^2)))) * (sqrt (((p . i) ^2) + ((p . j) ^2))) = s by XCMPLX_1:87; A8: 0 <= (p . j) * (p . j) by XREAL_1:63; then ((p . i) ^2) + 0 <= ((p . i) ^2) + ((p . j) ^2) by XREAL_1:6; then A9: sqrt ((p . i) ^2) <= sqrt (((p . i) ^2) + ((p . j) ^2)) by A5, SQUARE_1:26; now__::_thesis:_(p_._i)_/_(sqrt_(((p_._i)_^2)_+_((p_._j)_^2)))_in_[.(-_1),1.] percases ( p . i >= 0 or p . i <= 0 ) ; supposeA10: p . i >= 0 ; ::_thesis: (p . i) / (sqrt (((p . i) ^2) + ((p . j) ^2))) in [.(- 1),1.] then A11: p . i <= sqrt (((p . i) ^2) + ((p . j) ^2)) by A9, SQUARE_1:22; then (p . i) / (sqrt (((p . i) ^2) + ((p . j) ^2))) <= 1 by A6, XREAL_1:185; hence (p . i) / (sqrt (((p . i) ^2) + ((p . j) ^2))) in [.(- 1),1.] by A10, A11, XXREAL_1:1; ::_thesis: verum end; supposeA12: p . i <= 0 ; ::_thesis: (p . i) / (sqrt (((p . i) ^2) + ((p . j) ^2))) in [.(- 1),1.] then - (p . i) <= sqrt (((p . i) ^2) + ((p . j) ^2)) by A9, SQUARE_1:23; then - 1 <= (p . i) / (sqrt (((p . i) ^2) + ((p . j) ^2))) by A6, XREAL_1:194; hence (p . i) / (sqrt (((p . i) ^2) + ((p . j) ^2))) in [.(- 1),1.] by A6, A12, XXREAL_1:1; ::_thesis: verum end; end; end; then consider x being set such that A13: x in dom sin and x in [.(- (PI / 2)),(PI / 2).] and A14: sin . x = (p . i) / (sqrt (((p . i) ^2) + ((p . j) ^2))) by FUNCT_1:def_6, SIN_COS6:45; A15: (sqrt (((p . i) ^2) + ((p . j) ^2))) * (sqrt (((p . i) ^2) + ((p . j) ^2))) = (sqrt (((p . i) ^2) + ((p . j) ^2))) ^2 .= ((p . i) ^2) + ((p . j) ^2) by A5, A8, SQUARE_1:def_2 ; 0 <= s * s by XREAL_1:63; then A16: sqrt (s ^2) <= sqrt (((p . i) ^2) + ((p . j) ^2)) by A2, SQUARE_1:26; now__::_thesis:_s_/_(sqrt_(((p_._i)_^2)_+_((p_._j)_^2)))_in_[.(-_1),1.] percases ( s >= 0 or s <= 0 ) ; supposeA17: s >= 0 ; ::_thesis: s / (sqrt (((p . i) ^2) + ((p . j) ^2))) in [.(- 1),1.] then A18: s <= sqrt (((p . i) ^2) + ((p . j) ^2)) by A16, SQUARE_1:22; then s / (sqrt (((p . i) ^2) + ((p . j) ^2))) <= 1 by A6, XREAL_1:185; hence s / (sqrt (((p . i) ^2) + ((p . j) ^2))) in [.(- 1),1.] by A17, A18, XXREAL_1:1; ::_thesis: verum end; supposeA19: s <= 0 ; ::_thesis: s / (sqrt (((p . i) ^2) + ((p . j) ^2))) in [.(- 1),1.] then - s <= sqrt (((p . i) ^2) + ((p . j) ^2)) by A16, SQUARE_1:23; then - 1 <= s / (sqrt (((p . i) ^2) + ((p . j) ^2))) by A6, XREAL_1:194; hence s / (sqrt (((p . i) ^2) + ((p . j) ^2))) in [.(- 1),1.] by A6, A19, XXREAL_1:1; ::_thesis: verum end; end; end; then consider y being set such that y in dom sin and A20: y in [.(- (PI / 2)),(PI / 2).] and A21: sin . y = s / (sqrt (((p . i) ^2) + ((p . j) ^2))) by FUNCT_1:def_6, SIN_COS6:45; reconsider y = y as Real by A20; A22: ((sqrt (((p . i) ^2) + ((p . j) ^2))) * (p . i)) / (sqrt (((p . i) ^2) + ((p . j) ^2))) = p . i by A6, XCMPLX_1:89; reconsider x = x as Real by A13, FUNCT_2:def_1; A23: sin . x = sin x by SIN_COS:def_17; ((p . i) / (sqrt (((p . i) ^2) + ((p . j) ^2)))) * ((p . i) / (sqrt (((p . i) ^2) + ((p . j) ^2)))) = ((p . i) * (p . i)) / (((p . i) ^2) + ((p . j) ^2)) by A15, XCMPLX_1:76; then ((cos x) * (cos x)) + (((p . i) * (p . i)) / (((p . i) ^2) + ((p . j) ^2))) = 1 by A14, A23, SIN_COS:29; then (cos x) * (cos x) = 1 - (((p . i) * (p . i)) / (((p . i) ^2) + ((p . j) ^2))) .= ((((p . i) ^2) + ((p . j) ^2)) / (((p . i) ^2) + ((p . j) ^2))) - (((p . i) * (p . i)) / (((p . i) ^2) + ((p . j) ^2))) by A3, A4, XCMPLX_1:60 .= ((p . j) * (p . j)) / (((p . i) ^2) + ((p . j) ^2)) .= ((p . j) / (sqrt (((p . i) ^2) + ((p . j) ^2)))) ^2 by A15, XCMPLX_1:76 ; then A24: (cos x) ^2 = ((p . j) / (sqrt (((p . i) ^2) + ((p . j) ^2)))) ^2 ; percases ( cos x = (p . j) / (sqrt (((p . i) ^2) + ((p . j) ^2))) or cos x = - ((p . j) / (sqrt (((p . i) ^2) + ((p . j) ^2)))) ) by A24, SQUARE_1:40; supposeA25: cos x = (p . j) / (sqrt (((p . i) ^2) + ((p . j) ^2))) ; ::_thesis: ex r being real number st ((Mx2Tran (Rotation (i,j,n,r))) . p) . i = s take r = x - y; ::_thesis: ((Mx2Tran (Rotation (i,j,n,r))) . p) . i = s - (sin y) = sin ((- x) + r) by SIN_COS:31 .= ((sin (- x)) * (cos r)) + ((cos (- x)) * (sin r)) by SIN_COS:75 .= ((- (sin x)) * (cos r)) + ((cos (- x)) * (sin r)) by SIN_COS:31 .= (- ((sin x) * (cos r))) + ((cos x) * (sin r)) by SIN_COS:31 .= (- ((sin x) * (cos r))) + (- ((cos x) * (- (sin r)))) ; then sin y = ((sin x) * (cos r)) + ((cos x) * (- (sin r))) ; hence s = (sqrt (((p . i) ^2) + ((p . j) ^2))) * ((((p . i) / (sqrt (((p . i) ^2) + ((p . j) ^2)))) * (cos r)) + (((p . j) / (sqrt (((p . i) ^2) + ((p . j) ^2)))) * (- (sin r)))) by A7, A14, A21, A23, A25, SIN_COS:def_17 .= ((((sqrt (((p . i) ^2) + ((p . j) ^2))) * (p . i)) / (sqrt (((p . i) ^2) + ((p . j) ^2)))) * (cos r)) + ((((sqrt (((p . i) ^2) + ((p . j) ^2))) * (p . j)) / (sqrt (((p . i) ^2) + ((p . j) ^2)))) * (- (sin r))) .= ((p . i) * (cos r)) + ((p . j) * (- (sin r))) by A6, A22, XCMPLX_1:89 .= ((Mx2Tran (Rotation (i,j,n,r))) . p) . i by A1, Th21 ; ::_thesis: verum end; supposeA26: cos x = - ((p . j) / (sqrt (((p . i) ^2) + ((p . j) ^2)))) ; ::_thesis: ex r being real number st ((Mx2Tran (Rotation (i,j,n,r))) . p) . i = s take r = y - x; ::_thesis: ((Mx2Tran (Rotation (i,j,n,r))) . p) . i = s sin y = sin (x + r) .= ((sin x) * (cos r)) + ((cos x) * (sin r)) by SIN_COS:75 .= (((p . i) / (sqrt (((p . i) ^2) + ((p . j) ^2)))) * (cos r)) + (((p . j) / (sqrt (((p . i) ^2) + ((p . j) ^2)))) * (- (sin r))) by A14, A26, SIN_COS:def_17 ; hence s = (sqrt (((p . i) ^2) + ((p . j) ^2))) * ((((p . i) / (sqrt (((p . i) ^2) + ((p . j) ^2)))) * (cos r)) + (((p . j) / (sqrt (((p . i) ^2) + ((p . j) ^2)))) * (- (sin r)))) by A7, A21, SIN_COS:def_17 .= ((((sqrt (((p . i) ^2) + ((p . j) ^2))) * (p . i)) / (sqrt (((p . i) ^2) + ((p . j) ^2)))) * (cos r)) + ((((sqrt (((p . i) ^2) + ((p . j) ^2))) * (p . j)) / (sqrt (((p . i) ^2) + ((p . j) ^2)))) * (- (sin r))) .= ((p . i) * (cos r)) + ((p . j) * (- (sin r))) by A6, A22, XCMPLX_1:89 .= ((Mx2Tran (Rotation (i,j,n,r))) . p) . i by A1, Th21 ; ::_thesis: verum end; end; end; supposeA27: ( p . i = 0 & p . j = 0 ) ; ::_thesis: ex r being real number st ((Mx2Tran (Rotation (i,j,n,r))) . p) . i = s take r = 0 ; ::_thesis: ((Mx2Tran (Rotation (i,j,n,r))) . p) . i = s set M = Mx2Tran (Rotation (i,j,n,r)); Mx2Tran (Rotation (i,j,n,r)) = Mx2Tran (1. (F_Real,n)) by A1, Th18; then A28: Mx2Tran (Rotation (i,j,n,r)) = id (TOP-REAL n) by MATRTOP1:33; s = 0 by A2, A27, XREAL_1:63; hence ((Mx2Tran (Rotation (i,j,n,r))) . p) . i = s by A27, A28, FUNCT_1:17; ::_thesis: verum end; end; end; Lm6: for r being real number for i, j, n being Nat for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n holds ((((Mx2Tran (Rotation (i,j,n,r))) . p) . i) * (((Mx2Tran (Rotation (i,j,n,r))) . p) . i)) + ((((Mx2Tran (Rotation (i,j,n,r))) . p) . j) * (((Mx2Tran (Rotation (i,j,n,r))) . p) . j)) = ((p . i) * (p . i)) + ((p . j) * (p . j)) proof let r be real number ; ::_thesis: for i, j, n being Nat for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n holds ((((Mx2Tran (Rotation (i,j,n,r))) . p) . i) * (((Mx2Tran (Rotation (i,j,n,r))) . p) . i)) + ((((Mx2Tran (Rotation (i,j,n,r))) . p) . j) * (((Mx2Tran (Rotation (i,j,n,r))) . p) . j)) = ((p . i) * (p . i)) + ((p . j) * (p . j)) let i, j, n be Nat; ::_thesis: for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n holds ((((Mx2Tran (Rotation (i,j,n,r))) . p) . i) * (((Mx2Tran (Rotation (i,j,n,r))) . p) . i)) + ((((Mx2Tran (Rotation (i,j,n,r))) . p) . j) * (((Mx2Tran (Rotation (i,j,n,r))) . p) . j)) = ((p . i) * (p . i)) + ((p . j) * (p . j)) let p be Point of (TOP-REAL n); ::_thesis: ( 1 <= i & i < j & j <= n implies ((((Mx2Tran (Rotation (i,j,n,r))) . p) . i) * (((Mx2Tran (Rotation (i,j,n,r))) . p) . i)) + ((((Mx2Tran (Rotation (i,j,n,r))) . p) . j) * (((Mx2Tran (Rotation (i,j,n,r))) . p) . j)) = ((p . i) * (p . i)) + ((p . j) * (p . j)) ) set pk = p . i; set pj = p . j; set pkj = ((p . i) ^2) + ((p . j) ^2); set M = Mx2Tran (Rotation (i,j,n,r)); set S = sin r; set C = cos r; set Mp = (Mx2Tran (Rotation (i,j,n,r))) . p; A1: ((cos r) * (cos r)) + ((sin r) * (sin r)) = 1 by SIN_COS:29; assume A2: ( 1 <= i & i < j & j <= n ) ; ::_thesis: ((((Mx2Tran (Rotation (i,j,n,r))) . p) . i) * (((Mx2Tran (Rotation (i,j,n,r))) . p) . i)) + ((((Mx2Tran (Rotation (i,j,n,r))) . p) . j) * (((Mx2Tran (Rotation (i,j,n,r))) . p) . j)) = ((p . i) * (p . i)) + ((p . j) * (p . j)) then ((Mx2Tran (Rotation (i,j,n,r))) . p) . i = ((p . i) * (cos r)) + ((p . j) * (- (sin r))) by Th21; then A3: (((Mx2Tran (Rotation (i,j,n,r))) . p) . i) * (((Mx2Tran (Rotation (i,j,n,r))) . p) . i) = (((((p . i) * (p . i)) * (cos r)) * (cos r)) - ((((2 * (p . i)) * (p . j)) * (cos r)) * (sin r))) + ((((p . j) * (p . j)) * (sin r)) * (sin r)) ; ((Mx2Tran (Rotation (i,j,n,r))) . p) . j = ((p . i) * (sin r)) + ((p . j) * (cos r)) by A2, Th22; then (((Mx2Tran (Rotation (i,j,n,r))) . p) . j) * (((Mx2Tran (Rotation (i,j,n,r))) . p) . j) = (((((p . i) * (p . i)) * (sin r)) * (sin r)) + ((2 * ((p . i) * (p . j))) * ((cos r) * (sin r)))) + (((p . j) * (p . j)) * ((cos r) * (cos r))) ; hence ((((Mx2Tran (Rotation (i,j,n,r))) . p) . i) * (((Mx2Tran (Rotation (i,j,n,r))) . p) . i)) + ((((Mx2Tran (Rotation (i,j,n,r))) . p) . j) * (((Mx2Tran (Rotation (i,j,n,r))) . p) . j)) = (((p . i) * (p . i)) * (((cos r) * (cos r)) + ((sin r) * (sin r)))) + (((p . j) * (p . j)) * (((cos r) * (cos r)) + ((sin r) * (sin r)))) by A3 .= ((p . i) * (p . i)) + ((p . j) * (p . j)) by A1 ; ::_thesis: verum end; theorem Th25: :: MATRTOP3:25 for s being real number for i, j, n being Nat for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n & s ^2 <= ((p . i) ^2) + ((p . j) ^2) holds ex r being real number st ((Mx2Tran (Rotation (i,j,n,r))) . p) . j = s proof let s be real number ; ::_thesis: for i, j, n being Nat for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n & s ^2 <= ((p . i) ^2) + ((p . j) ^2) holds ex r being real number st ((Mx2Tran (Rotation (i,j,n,r))) . p) . j = s let i, j, n be Nat; ::_thesis: for p being Point of (TOP-REAL n) st 1 <= i & i < j & j <= n & s ^2 <= ((p . i) ^2) + ((p . j) ^2) holds ex r being real number st ((Mx2Tran (Rotation (i,j,n,r))) . p) . j = s let p be Point of (TOP-REAL n); ::_thesis: ( 1 <= i & i < j & j <= n & s ^2 <= ((p . i) ^2) + ((p . j) ^2) implies ex r being real number st ((Mx2Tran (Rotation (i,j,n,r))) . p) . j = s ) set pk = p . i; set pj = p . j; set pkj = ((p . i) ^2) + ((p . j) ^2); set ps = (((p . i) ^2) + ((p . j) ^2)) - (s ^2); assume that A1: ( 1 <= i & i < j & j <= n ) and A2: s ^2 <= ((p . i) ^2) + ((p . j) ^2) ; ::_thesis: ex r being real number st ((Mx2Tran (Rotation (i,j,n,r))) . p) . j = s 0 <= s * s by XREAL_1:63; then A3: (((p . i) ^2) + ((p . j) ^2)) - (s ^2) <= (((p . i) ^2) + ((p . j) ^2)) - 0 by XREAL_1:6; A4: (s ^2) - (s ^2) <= (((p . i) ^2) + ((p . j) ^2)) - (s ^2) by A2, XREAL_1:6; then (sqrt ((((p . i) ^2) + ((p . j) ^2)) - (s ^2))) ^2 = (((p . i) ^2) + ((p . j) ^2)) - (s ^2) by SQUARE_1:def_2; then consider r being real number such that A5: ((Mx2Tran (Rotation (i,j,n,r))) . p) . i = sqrt ((((p . i) ^2) + ((p . j) ^2)) - (s ^2)) by A1, A3, Th24; set M = Mx2Tran (Rotation (i,j,n,r)); set Mp = (Mx2Tran (Rotation (i,j,n,r))) . p; ((p . i) ^2) + ((p . j) ^2) = ((sqrt ((((p . i) ^2) + ((p . j) ^2)) - (s ^2))) ^2) + ((((Mx2Tran (Rotation (i,j,n,r))) . p) . j) * (((Mx2Tran (Rotation (i,j,n,r))) . p) . j)) by A5, A1, Lm6 .= ((((p . i) ^2) + ((p . j) ^2)) - (s ^2)) + ((((Mx2Tran (Rotation (i,j,n,r))) . p) . j) * (((Mx2Tran (Rotation (i,j,n,r))) . p) . j)) by A4, SQUARE_1:def_2 ; then A6: s ^2 = (((Mx2Tran (Rotation (i,j,n,r))) . p) . j) ^2 ; percases ( ((Mx2Tran (Rotation (i,j,n,r))) . p) . j = s or ((Mx2Tran (Rotation (i,j,n,r))) . p) . j = - s ) by A6, SQUARE_1:40; suppose ((Mx2Tran (Rotation (i,j,n,r))) . p) . j = s ; ::_thesis: ex r being real number st ((Mx2Tran (Rotation (i,j,n,r))) . p) . j = s hence ex r being real number st ((Mx2Tran (Rotation (i,j,n,r))) . p) . j = s ; ::_thesis: verum end; supposeA7: ((Mx2Tran (Rotation (i,j,n,r))) . p) . j = - s ; ::_thesis: ex r being real number st ((Mx2Tran (Rotation (i,j,n,r))) . p) . j = s take R = r + PI; ::_thesis: ((Mx2Tran (Rotation (i,j,n,R))) . p) . j = s thus ((Mx2Tran (Rotation (i,j,n,R))) . p) . j = ((p . i) * (sin R)) + ((p . j) * (cos R)) by A1, Th22 .= ((p . i) * (- (sin r))) + ((p . j) * (cos R)) by SIN_COS:79 .= ((p . i) * (- (sin r))) + ((p . j) * (- (cos r))) by SIN_COS:79 .= - (((p . i) * (sin r)) + ((p . j) * (cos r))) .= - (((Mx2Tran (Rotation (i,j,n,r))) . p) . j) by A1, Th22 .= s by A7 ; ::_thesis: verum end; end; end; theorem Th26: :: MATRTOP3:26 for r being real number for i, j, n being Nat st 1 <= i & i < j & j <= n holds Mx2Tran (Rotation (i,j,n,r)) is {i,j} -support-yielding proof let r be real number ; ::_thesis: for i, j, n being Nat st 1 <= i & i < j & j <= n holds Mx2Tran (Rotation (i,j,n,r)) is {i,j} -support-yielding let i, j, n be Nat; ::_thesis: ( 1 <= i & i < j & j <= n implies Mx2Tran (Rotation (i,j,n,r)) is {i,j} -support-yielding ) set M = Mx2Tran (Rotation (i,j,n,r)); assume A1: ( 1 <= i & i < j & j <= n ) ; ::_thesis: Mx2Tran (Rotation (i,j,n,r)) is {i,j} -support-yielding let f be Function; :: according to MATRTOP3:def_1 ::_thesis: for x being set st f in dom (Mx2Tran (Rotation (i,j,n,r))) & ((Mx2Tran (Rotation (i,j,n,r))) . f) . x <> f . x holds x in {i,j} let x be set ; ::_thesis: ( f in dom (Mx2Tran (Rotation (i,j,n,r))) & ((Mx2Tran (Rotation (i,j,n,r))) . f) . x <> f . x implies x in {i,j} ) assume that A2: f in dom (Mx2Tran (Rotation (i,j,n,r))) and A3: ((Mx2Tran (Rotation (i,j,n,r))) . f) . x <> f . x ; ::_thesis: x in {i,j} reconsider p = f as Point of (TOP-REAL n) by A2, FUNCT_2:52; len p = n by CARD_1:def_7; then A4: dom p = Seg n by FINSEQ_1:def_3; len ((Mx2Tran (Rotation (i,j,n,r))) . p) = n by CARD_1:def_7; then A5: dom ((Mx2Tran (Rotation (i,j,n,r))) . p) = Seg n by FINSEQ_1:def_3; percases ( not x in Seg n or x in Seg n ) ; supposeA6: not x in Seg n ; ::_thesis: x in {i,j} then ((Mx2Tran (Rotation (i,j,n,r))) . p) . x = {} by A5, FUNCT_1:def_2; hence x in {i,j} by A3, A4, A6, FUNCT_1:def_2; ::_thesis: verum end; supposeA7: x in Seg n ; ::_thesis: x in {i,j} ((Mx2Tran (Rotation (i,j,n,r))) . p) . x <> p . x by A3; then ( x = i or x = j ) by A1, A7, Th20; hence x in {i,j} by TARSKI:def_2; ::_thesis: verum end; end; end; begin definition let n be Nat; let f be Function of (TOP-REAL n),(TOP-REAL n); attrf is rotation means :Def4: :: MATRTOP3:def 4 for p being Point of (TOP-REAL n) holds |.p.| = |.(f . p).|; end; :: deftheorem Def4 defines rotation MATRTOP3:def_4_:_ for n being Nat for f being Function of (TOP-REAL n),(TOP-REAL n) holds ( f is rotation iff for p being Point of (TOP-REAL n) holds |.p.| = |.(f . p).| ); theorem Th27: :: MATRTOP3:27 for i, n being Nat st i in Seg n holds Mx2Tran (AxialSymmetry (i,n)) is rotation proof let i, n be Nat; ::_thesis: ( i in Seg n implies Mx2Tran (AxialSymmetry (i,n)) is rotation ) set S = Seg n; set M = Mx2Tran (AxialSymmetry (i,n)); assume A1: i in Seg n ; ::_thesis: Mx2Tran (AxialSymmetry (i,n)) is rotation let p be Point of (TOP-REAL n); :: according to MATRTOP3:def_4 ::_thesis: |.p.| = |.((Mx2Tran (AxialSymmetry (i,n))) . p).| len p = n by CARD_1:def_7; then A2: i in dom p by A1, FINSEQ_1:def_3; thus |.((Mx2Tran (AxialSymmetry (i,n))) . p).| = sqrt (Sum (sqr (p +* (i,(- (p . i)))))) by A1, Th10 .= sqrt (((Sum (sqr p)) - ((p . i) ^2)) + ((- (p . i)) ^2)) by A2, Th3 .= |.p.| ; ::_thesis: verum end; definition let n be Nat; let f be Function of (TOP-REAL n),(TOP-REAL n); attrf is base_rotation means :Def5: :: MATRTOP3:def 5 ex F being FinSequence of (GFuncs the carrier of (TOP-REAL n)) st ( f = Product F & ( for k being Nat st k in dom F holds ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & F . k = Mx2Tran (Rotation (i,j,n,r)) ) ) ); end; :: deftheorem Def5 defines base_rotation MATRTOP3:def_5_:_ for n being Nat for f being Function of (TOP-REAL n),(TOP-REAL n) holds ( f is base_rotation iff ex F being FinSequence of (GFuncs the carrier of (TOP-REAL n)) st ( f = Product F & ( for k being Nat st k in dom F holds ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & F . k = Mx2Tran (Rotation (i,j,n,r)) ) ) ) ); registration let n be Nat; cluster id (TOP-REAL n) -> base_rotation ; coherence id (TOP-REAL n) is base_rotation proof set S = the carrier of (TOP-REAL n); set G = GFuncs the carrier of (TOP-REAL n); take F = <*> the carrier of (GFuncs the carrier of (TOP-REAL n)); :: according to MATRTOP3:def_5 ::_thesis: ( id (TOP-REAL n) = Product F & ( for k being Nat st k in dom F holds ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & F . k = Mx2Tran (Rotation (i,j,n,r)) ) ) ) thus Product F = 1_ (GFuncs the carrier of (TOP-REAL n)) by GROUP_4:8 .= the_unity_wrt the multF of (GFuncs the carrier of (TOP-REAL n)) by GROUP_1:22 .= id (TOP-REAL n) by MONOID_0:75 ; ::_thesis: for k being Nat st k in dom F holds ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & F . k = Mx2Tran (Rotation (i,j,n,r)) ) let k be Nat; ::_thesis: ( k in dom F implies ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & F . k = Mx2Tran (Rotation (i,j,n,r)) ) ) assume k in dom F ; ::_thesis: ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & F . k = Mx2Tran (Rotation (i,j,n,r)) ) hence ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & F . k = Mx2Tran (Rotation (i,j,n,r)) ) ; ::_thesis: verum end; end; registration let n be Nat; cluster non empty Relation-like the carrier of (TOP-REAL n) -defined the carrier of (TOP-REAL n) -valued Function-like total quasi_total FinSequence-yielding Function-yielding V235() base_rotation for Element of bool [: the carrier of (TOP-REAL n), the carrier of (TOP-REAL n):]; existence ex b1 being Function of (TOP-REAL n),(TOP-REAL n) st b1 is base_rotation proof id (TOP-REAL n) is base_rotation ; hence ex b1 being Function of (TOP-REAL n),(TOP-REAL n) st b1 is base_rotation ; ::_thesis: verum end; end; registration let n be Nat; let f, g be base_rotation Function of (TOP-REAL n),(TOP-REAL n); clusterg (#) f -> base_rotation for Function of (TOP-REAL n),(TOP-REAL n); coherence for b1 being Function of (TOP-REAL n),(TOP-REAL n) st b1 = f * g holds b1 is base_rotation proof consider F being FinSequence of (GFuncs the carrier of (TOP-REAL n)) such that A1: f = Product F and A2: for k being Nat st k in dom F holds ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & F . k = Mx2Tran (Rotation (i,j,n,r)) ) by Def5; consider G being FinSequence of (GFuncs the carrier of (TOP-REAL n)) such that A3: g = Product G and A4: for k being Nat st k in dom G holds ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & G . k = Mx2Tran (Rotation (i,j,n,r)) ) by Def5; f * g is base_rotation proof take GF = G ^ F; :: according to MATRTOP3:def_5 ::_thesis: ( f * g = Product GF & ( for k being Nat st k in dom GF holds ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & GF . k = Mx2Tran (Rotation (i,j,n,r)) ) ) ) thus Product GF = (Product G) * (Product F) by GROUP_4:5 .= f * g by A3, A1, MONOID_0:70 ; ::_thesis: for k being Nat st k in dom GF holds ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & GF . k = Mx2Tran (Rotation (i,j,n,r)) ) let k be Nat; ::_thesis: ( k in dom GF implies ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & GF . k = Mx2Tran (Rotation (i,j,n,r)) ) ) assume A5: k in dom GF ; ::_thesis: ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & GF . k = Mx2Tran (Rotation (i,j,n,r)) ) percases ( k in dom G or ex m being Nat st ( m in dom F & k = (len G) + m ) ) by A5, FINSEQ_1:25; supposeA6: k in dom G ; ::_thesis: ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & GF . k = Mx2Tran (Rotation (i,j,n,r)) ) then G . k = GF . k by FINSEQ_1:def_7; hence ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & GF . k = Mx2Tran (Rotation (i,j,n,r)) ) by A4, A6; ::_thesis: verum end; suppose ex m being Nat st ( m in dom F & k = (len G) + m ) ; ::_thesis: ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & GF . k = Mx2Tran (Rotation (i,j,n,r)) ) then consider m being Nat such that A7: m in dom F and A8: k = (len G) + m ; GF . k = F . m by A7, A8, FINSEQ_1:def_7; hence ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & GF . k = Mx2Tran (Rotation (i,j,n,r)) ) by A2, A7; ::_thesis: verum end; end; end; hence for b1 being Function of (TOP-REAL n),(TOP-REAL n) st b1 = f * g holds b1 is base_rotation ; ::_thesis: verum end; end; Lm7: for r being real number for i, j, n being Nat st 1 <= i & i < j & j <= n holds Mx2Tran (Rotation (i,j,n,r)) is rotation proof let r be real number ; ::_thesis: for i, j, n being Nat st 1 <= i & i < j & j <= n holds Mx2Tran (Rotation (i,j,n,r)) is rotation let i, j, n be Nat; ::_thesis: ( 1 <= i & i < j & j <= n implies Mx2Tran (Rotation (i,j,n,r)) is rotation ) assume A1: ( 1 <= i & i < j & j <= n ) ; ::_thesis: Mx2Tran (Rotation (i,j,n,r)) is rotation let p be Point of (TOP-REAL n); :: according to MATRTOP3:def_4 ::_thesis: |.p.| = |.((Mx2Tran (Rotation (i,j,n,r))) . p).| p = @ (@ p) ; then reconsider P = p as FinSequence of REAL ; set M = Mx2Tran (Rotation (i,j,n,r)); set Mp = (Mx2Tran (Rotation (i,j,n,r))) . p; set p1 = P | (i -' 1); set p2 = (P /^ i) | ((j -' i) -' 1); set p3 = P /^ j; set s = sin r; set c = cos r; reconsider pk = P . i, pj = P . j as Element of REAL ; A2: sqr <*pk*> = <*(pk ^2)*> by RVSUM_1:55; set pij1 = (pk * (cos r)) + (pj * (- (sin r))); set pij2 = (pk * (sin r)) + (pj * (cos r)); A3: sqr <*((pk * (cos r)) + (pj * (- (sin r))))*> = <*(((pk * (cos r)) + (pj * (- (sin r)))) ^2)*> by RVSUM_1:55; A4: sqr <*((pk * (sin r)) + (pj * (cos r)))*> = <*(((pk * (sin r)) + (pj * (cos r))) ^2)*> by RVSUM_1:55; A5: sqr <*pj*> = <*(pj ^2)*> by RVSUM_1:55; len p = n by CARD_1:def_7; then p = ((((P | (i -' 1)) ^ <*pk*>) ^ ((P /^ i) | ((j -' i) -' 1))) ^ <*pj*>) ^ (P /^ j) by A1, FINSEQ_7:1; then sqr p = (sqr ((((P | (i -' 1)) ^ <*pk*>) ^ ((P /^ i) | ((j -' i) -' 1))) ^ <*pj*>)) ^ (sqr (P /^ j)) by RVSUM_1:144 .= ((sqr (((P | (i -' 1)) ^ <*pk*>) ^ ((P /^ i) | ((j -' i) -' 1)))) ^ (sqr <*pj*>)) ^ (sqr (P /^ j)) by RVSUM_1:144 .= (((sqr ((P | (i -' 1)) ^ <*pk*>)) ^ (sqr ((P /^ i) | ((j -' i) -' 1)))) ^ (sqr <*pj*>)) ^ (sqr (P /^ j)) by RVSUM_1:144 .= ((((sqr (P | (i -' 1))) ^ (sqr <*pk*>)) ^ (sqr ((P /^ i) | ((j -' i) -' 1)))) ^ (sqr <*pj*>)) ^ (sqr (P /^ j)) by RVSUM_1:144 ; then A6: Sum (sqr p) = (Sum ((((sqr (P | (i -' 1))) ^ (sqr <*pk*>)) ^ (sqr ((P /^ i) | ((j -' i) -' 1)))) ^ (sqr <*pj*>))) + (Sum (sqr (P /^ j))) by RVSUM_1:75 .= ((Sum (((sqr (P | (i -' 1))) ^ (sqr <*pk*>)) ^ (sqr ((P /^ i) | ((j -' i) -' 1))))) + (pj ^2)) + (Sum (sqr (P /^ j))) by A5, RVSUM_1:74 .= (((Sum ((sqr (P | (i -' 1))) ^ (sqr <*pk*>))) + (Sum (sqr ((P /^ i) | ((j -' i) -' 1))))) + (pj ^2)) + (Sum (sqr (P /^ j))) by RVSUM_1:75 .= ((((Sum (sqr (P | (i -' 1)))) + (pk ^2)) + (Sum (sqr ((P /^ i) | ((j -' i) -' 1))))) + (pj ^2)) + (Sum (sqr (P /^ j))) by A2, RVSUM_1:74 ; A7: ((cos r) * (cos r)) + ((sin r) * (sin r)) = 1 by SIN_COS:29; A8: (((pk * (cos r)) + (pj * (- (sin r)))) ^2) + (((pk * (sin r)) + (pj * (cos r))) ^2) = ((pk * pk) * (((cos r) * (cos r)) + ((sin r) * (sin r)))) + ((pj * pj) * (((cos r) * (cos r)) + ((sin r) * (sin r)))) .= (pk ^2) + (pj ^2) by A7 ; (Mx2Tran (Rotation (i,j,n,r))) . p = ((((P | (i -' 1)) ^ <*((pk * (cos r)) + (pj * (- (sin r))))*>) ^ ((P /^ i) | ((j -' i) -' 1))) ^ <*((pk * (sin r)) + (pj * (cos r)))*>) ^ (P /^ j) by A1, Th23; then sqr ((Mx2Tran (Rotation (i,j,n,r))) . p) = (sqr ((((P | (i -' 1)) ^ <*((pk * (cos r)) + (pj * (- (sin r))))*>) ^ ((P /^ i) | ((j -' i) -' 1))) ^ <*((pk * (sin r)) + (pj * (cos r)))*>)) ^ (sqr (P /^ j)) by RVSUM_1:144 .= ((sqr (((P | (i -' 1)) ^ <*((pk * (cos r)) + (pj * (- (sin r))))*>) ^ ((P /^ i) | ((j -' i) -' 1)))) ^ (sqr <*((pk * (sin r)) + (pj * (cos r)))*>)) ^ (sqr (P /^ j)) by RVSUM_1:144 .= (((sqr ((P | (i -' 1)) ^ <*((pk * (cos r)) + (pj * (- (sin r))))*>)) ^ (sqr ((P /^ i) | ((j -' i) -' 1)))) ^ (sqr <*((pk * (sin r)) + (pj * (cos r)))*>)) ^ (sqr (P /^ j)) by RVSUM_1:144 .= ((((sqr (P | (i -' 1))) ^ (sqr <*((pk * (cos r)) + (pj * (- (sin r))))*>)) ^ (sqr ((P /^ i) | ((j -' i) -' 1)))) ^ (sqr <*((pk * (sin r)) + (pj * (cos r)))*>)) ^ (sqr (P /^ j)) by RVSUM_1:144 ; then Sum (sqr ((Mx2Tran (Rotation (i,j,n,r))) . p)) = (Sum ((((sqr (P | (i -' 1))) ^ (sqr <*((pk * (cos r)) + (pj * (- (sin r))))*>)) ^ (sqr ((P /^ i) | ((j -' i) -' 1)))) ^ (sqr <*((pk * (sin r)) + (pj * (cos r)))*>))) + (Sum (sqr (P /^ j))) by RVSUM_1:75 .= ((Sum (((sqr (P | (i -' 1))) ^ (sqr <*((pk * (cos r)) + (pj * (- (sin r))))*>)) ^ (sqr ((P /^ i) | ((j -' i) -' 1))))) + (((pk * (sin r)) + (pj * (cos r))) ^2)) + (Sum (sqr (P /^ j))) by A4, RVSUM_1:74 .= (((Sum ((sqr (P | (i -' 1))) ^ (sqr <*((pk * (cos r)) + (pj * (- (sin r))))*>))) + (Sum (sqr ((P /^ i) | ((j -' i) -' 1))))) + (((pk * (sin r)) + (pj * (cos r))) ^2)) + (Sum (sqr (P /^ j))) by RVSUM_1:75 .= ((((Sum (sqr (P | (i -' 1)))) + (((pk * (cos r)) + (pj * (- (sin r)))) ^2)) + (Sum (sqr ((P /^ i) | ((j -' i) -' 1))))) + (((pk * (sin r)) + (pj * (cos r))) ^2)) + (Sum (sqr (P /^ j))) by A3, RVSUM_1:74 ; hence |.p.| = |.((Mx2Tran (Rotation (i,j,n,r))) . p).| by A6, A8; ::_thesis: verum end; theorem Th28: :: MATRTOP3:28 for r being real number for i, j, n being Nat st 1 <= i & i < j & j <= n holds Mx2Tran (Rotation (i,j,n,r)) is base_rotation proof let r be real number ; ::_thesis: for i, j, n being Nat st 1 <= i & i < j & j <= n holds Mx2Tran (Rotation (i,j,n,r)) is base_rotation let i, j, n be Nat; ::_thesis: ( 1 <= i & i < j & j <= n implies Mx2Tran (Rotation (i,j,n,r)) is base_rotation ) assume A1: ( 1 <= i & i < j & j <= n ) ; ::_thesis: Mx2Tran (Rotation (i,j,n,r)) is base_rotation set S = the carrier of (TOP-REAL n); set G = GFuncs the carrier of (TOP-REAL n); reconsider M = Mx2Tran (Rotation (i,j,n,r)) as Element of (GFuncs the carrier of (TOP-REAL n)) by MONOID_0:73; take F = <*M*>; :: according to MATRTOP3:def_5 ::_thesis: ( Mx2Tran (Rotation (i,j,n,r)) = Product F & ( for k being Nat st k in dom F holds ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & F . k = Mx2Tran (Rotation (i,j,n,r)) ) ) ) thus Product F = Mx2Tran (Rotation (i,j,n,r)) by GROUP_4:9; ::_thesis: for k being Nat st k in dom F holds ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & F . k = Mx2Tran (Rotation (i,j,n,r)) ) let k be Nat; ::_thesis: ( k in dom F implies ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & F . k = Mx2Tran (Rotation (i,j,n,r)) ) ) assume k in dom F ; ::_thesis: ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & F . k = Mx2Tran (Rotation (i,j,n,r)) ) then k in {1} by FINSEQ_1:2, FINSEQ_1:38; then A2: k = 1 by TARSKI:def_1; F . 1 = M by FINSEQ_1:40; hence ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & F . k = Mx2Tran (Rotation (i,j,n,r)) ) by A1, A2; ::_thesis: verum end; Lm8: for n being Nat for f, g being Function of (TOP-REAL n),(TOP-REAL n) st f is rotation & g is rotation holds f * g is rotation proof let n be Nat; ::_thesis: for f, g being Function of (TOP-REAL n),(TOP-REAL n) st f is rotation & g is rotation holds f * g is rotation set TR = TOP-REAL n; let f, g be Function of (TOP-REAL n),(TOP-REAL n); ::_thesis: ( f is rotation & g is rotation implies f * g is rotation ) assume that A1: f is rotation and A2: g is rotation ; ::_thesis: f * g is rotation let p be Point of (TOP-REAL n); :: according to MATRTOP3:def_4 ::_thesis: |.p.| = |.((f * g) . p).| dom (f * g) = the carrier of (TOP-REAL n) by FUNCT_2:52; hence |.((f * g) . p).| = |.(f . (g . p)).| by FUNCT_1:12 .= |.(g . p).| by A1, Def4 .= |.p.| by A2, Def4 ; ::_thesis: verum end; registration let n be Nat; cluster Function-like quasi_total base_rotation -> additive homogeneous being_homeomorphism rotation for Element of bool [: the carrier of (TOP-REAL n), the carrier of (TOP-REAL n):]; coherence for b1 being Function of (TOP-REAL n),(TOP-REAL n) st b1 is base_rotation holds ( b1 is homogeneous & b1 is additive & b1 is rotation & b1 is being_homeomorphism ) proof set TR = TOP-REAL n; set G = GFuncs the carrier of (TOP-REAL n); let f be Function of (TOP-REAL n),(TOP-REAL n); ::_thesis: ( f is base_rotation implies ( f is homogeneous & f is additive & f is rotation & f is being_homeomorphism ) ) assume f is base_rotation ; ::_thesis: ( f is homogeneous & f is additive & f is rotation & f is being_homeomorphism ) then consider F being FinSequence of (GFuncs the carrier of (TOP-REAL n)) such that A1: f = Product F and A2: for k being Nat st k in dom F holds ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & F . k = Mx2Tran (Rotation (i,j,n,r)) ) by Def5; defpred S1[ Nat] means ( n <= len F implies for g being Function of (TOP-REAL n),(TOP-REAL n) st g = Product (F | n) holds ( g is homogeneous & g is additive & g is being_homeomorphism & g is rotation ) ); A3: for m being Nat st S1[m] holds S1[m + 1] proof let m be Nat; ::_thesis: ( S1[m] implies S1[m + 1] ) assume A4: S1[m] ; ::_thesis: S1[m + 1] set m1 = m + 1; reconsider P = Product (F | m) as Function of (TOP-REAL n),(TOP-REAL n) by MONOID_0:73; assume A5: m + 1 <= len F ; ::_thesis: for g being Function of (TOP-REAL n),(TOP-REAL n) st g = Product (F | (m + 1)) holds ( g is homogeneous & g is additive & g is being_homeomorphism & g is rotation ) 1 <= m + 1 by NAT_1:11; then A6: m + 1 in dom F by A5, FINSEQ_3:25; then consider i, j being Nat, r being real number such that A7: ( 1 <= i & i < j & j <= n ) and A8: F . (m + 1) = Mx2Tran (Rotation (i,j,n,r)) by A2; reconsider M = Mx2Tran (Rotation (i,j,n,r)) as Element of (GFuncs the carrier of (TOP-REAL n)) by MONOID_0:73; A9: F | (m + 1) = (F | m) ^ <*M*> by A6, A8, FINSEQ_5:10; let g be Function of (TOP-REAL n),(TOP-REAL n); ::_thesis: ( g = Product (F | (m + 1)) implies ( g is homogeneous & g is additive & g is being_homeomorphism & g is rotation ) ) assume A10: g = Product (F | (m + 1)) ; ::_thesis: ( g is homogeneous & g is additive & g is being_homeomorphism & g is rotation ) A11: g = (Product (F | m)) * M by A9, A10, GROUP_4:6 .= (Mx2Tran (Rotation (i,j,n,r))) * P by MONOID_0:70 ; A12: Mx2Tran (Rotation (i,j,n,r)) is rotation by A7, Lm7; ( P is homogeneous & P is additive & P is being_homeomorphism & P is rotation ) by A4, A5, NAT_1:13; hence ( g is homogeneous & g is additive & g is being_homeomorphism & g is rotation ) by A11, A12, Lm8, TOPS_2:57; ::_thesis: verum end; A13: F | (len F) = F by FINSEQ_1:58; A14: S1[ 0 ] proof A15: id (TOP-REAL n) is rotation proof let p be Point of (TOP-REAL n); :: according to MATRTOP3:def_4 ::_thesis: |.p.| = |.((id (TOP-REAL n)) . p).| thus |.p.| = |.((id (TOP-REAL n)) . p).| by FUNCT_1:17; ::_thesis: verum end; assume 0 <= len F ; ::_thesis: for g being Function of (TOP-REAL n),(TOP-REAL n) st g = Product (F | 0) holds ( g is homogeneous & g is additive & g is being_homeomorphism & g is rotation ) let g be Function of (TOP-REAL n),(TOP-REAL n); ::_thesis: ( g = Product (F | 0) implies ( g is homogeneous & g is additive & g is being_homeomorphism & g is rotation ) ) A16: id (TOP-REAL n) is homogeneous proof let r be real number ; :: according to TOPREAL9:def_4 ::_thesis: for b1 being Element of the carrier of (TOP-REAL n) holds (id (TOP-REAL n)) . (r * b1) = r * ((id (TOP-REAL n)) . b1) let p be Point of (TOP-REAL n); ::_thesis: (id (TOP-REAL n)) . (r * p) = r * ((id (TOP-REAL n)) . p) thus (id (TOP-REAL n)) . (r * p) = r * p by FUNCT_1:17 .= r * ((id (TOP-REAL n)) . p) by FUNCT_1:17 ; ::_thesis: verum end; assume A17: g = Product (F | 0) ; ::_thesis: ( g is homogeneous & g is additive & g is being_homeomorphism & g is rotation ) F | 0 = <*> the carrier of (GFuncs the carrier of (TOP-REAL n)) ; then g = 1_ (GFuncs the carrier of (TOP-REAL n)) by A17, GROUP_4:8 .= the_unity_wrt the multF of (GFuncs the carrier of (TOP-REAL n)) by GROUP_1:22 .= id (TOP-REAL n) by MONOID_0:75 ; hence ( g is homogeneous & g is additive & g is being_homeomorphism & g is rotation ) by A16, A15; ::_thesis: verum end; for m being Nat holds S1[m] from NAT_1:sch_2(A14, A3); hence ( f is homogeneous & f is additive & f is rotation & f is being_homeomorphism ) by A1, A13; ::_thesis: verum end; end; registration let n be Nat; let f be base_rotation Function of (TOP-REAL n),(TOP-REAL n); clusterf /" -> base_rotation ; coherence f " is base_rotation proof set TR = TOP-REAL n; set G = GFuncs the carrier of (TOP-REAL n); defpred S1[ Nat] means for F being FinSequence of (GFuncs the carrier of (TOP-REAL n)) for f being Function of (TOP-REAL n),(TOP-REAL n) st len F = n & Product F = f & ( for k being Nat st k in dom F holds ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & F . k = Mx2Tran (Rotation (i,j,n,r)) ) ) holds f " is base_rotation ; consider F being FinSequence of (GFuncs the carrier of (TOP-REAL n)) such that A1: ( f = Product F & ( for k being Nat st k in dom F holds ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & F . k = Mx2Tran (Rotation (i,j,n,r)) ) ) ) by Def5; A2: for i being Nat st S1[i] holds S1[i + 1] proof let z be Nat; ::_thesis: ( S1[z] implies S1[z + 1] ) assume A3: S1[z] ; ::_thesis: S1[z + 1] set z1 = z + 1; let F be FinSequence of (GFuncs the carrier of (TOP-REAL n)); ::_thesis: for f being Function of (TOP-REAL n),(TOP-REAL n) st len F = z + 1 & Product F = f & ( for k being Nat st k in dom F holds ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & F . k = Mx2Tran (Rotation (i,j,n,r)) ) ) holds f " is base_rotation let f be Function of (TOP-REAL n),(TOP-REAL n); ::_thesis: ( len F = z + 1 & Product F = f & ( for k being Nat st k in dom F holds ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & F . k = Mx2Tran (Rotation (i,j,n,r)) ) ) implies f " is base_rotation ) assume that A4: len F = z + 1 and A5: Product F = f and A6: for k being Nat st k in dom F holds ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & F . k = Mx2Tran (Rotation (i,j,n,r)) ) ; ::_thesis: f " is base_rotation set Fz = F | z; reconsider fz = Product (F | z) as Function of (TOP-REAL n),(TOP-REAL n) by MONOID_0:73; 1 <= z + 1 by NAT_1:11; then z + 1 in dom F by A4, FINSEQ_3:25; then consider i, j being Nat, r being real number such that A7: ( 1 <= i & i < j & j <= n ) and A8: F . (z + 1) = Mx2Tran (Rotation (i,j,n,r)) by A6; set m = Mx2Tran (Rotation (i,j,n,r)); reconsider M = Mx2Tran (Rotation (i,j,n,r)) as Element of (GFuncs the carrier of (TOP-REAL n)) by MONOID_0:73; F = (F | z) ^ <*M*> by A4, A8, FINSEQ_3:55; then A9: f = (Product (F | z)) * M by A5, GROUP_4:6 .= (Mx2Tran (Rotation (i,j,n,r))) * fz by MONOID_0:70 ; A10: dom (F | z) c= dom F by RELAT_1:60; A11: now__::_thesis:_for_k_being_Nat_st_k_in_dom_(F_|_z)_holds_ ex_i,_j_being_Nat_ex_r_being_real_number_st_ (_1_<=_i_&_i_<_j_&_j_<=_n_&_(F_|_z)_._k_=_Mx2Tran_(Rotation_(i,j,n,r))_) let k be Nat; ::_thesis: ( k in dom (F | z) implies ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & (F | z) . k = Mx2Tran (Rotation (i,j,n,r)) ) ) assume A12: k in dom (F | z) ; ::_thesis: ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & (F | z) . k = Mx2Tran (Rotation (i,j,n,r)) ) then (F | z) . k = F . k by FUNCT_1:47; hence ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & (F | z) . k = Mx2Tran (Rotation (i,j,n,r)) ) by A6, A10, A12; ::_thesis: verum end; then A13: fz is base_rotation by Def5; ( Det (Rotation (i,j,n,r)) <> 0. F_Real & (Rotation (i,j,n,r)) ~ = Rotation (i,j,n,(- r)) ) by A7, Lm5, Th13; then A14: (Mx2Tran (Rotation (i,j,n,r))) " = Mx2Tran (Rotation (i,j,n,(- r))) by MATRTOP1:43; A15: ( rng (Mx2Tran (Rotation (i,j,n,r))) = [#] (TOP-REAL n) & Mx2Tran (Rotation (i,j,n,r)) is one-to-one & dom (Mx2Tran (Rotation (i,j,n,r))) = [#] (TOP-REAL n) ) by TOPS_2:def_5; then Mx2Tran (Rotation (i,j,n,r)) is onto by FUNCT_2:def_3; then A16: (Mx2Tran (Rotation (i,j,n,r))) " = (Mx2Tran (Rotation (i,j,n,r))) " by A15, TOPS_2:def_4; len (F | z) = z by A4, FINSEQ_1:59, NAT_1:11; then A17: fz " is base_rotation by A3, A11; ( fz is one-to-one & dom fz = [#] (TOP-REAL n) & rng fz = [#] (TOP-REAL n) ) by A13, TOPS_2:def_5; then A18: f " = (fz ") * ((Mx2Tran (Rotation (i,j,n,r))) ") by A9, A15, TOPS_2:53; Mx2Tran (Rotation (i,j,n,(- r))) is base_rotation by A7, Th28; hence f " is base_rotation by A14, A16, A17, A18; ::_thesis: verum end; A19: S1[ 0 ] proof let F be FinSequence of (GFuncs the carrier of (TOP-REAL n)); ::_thesis: for f being Function of (TOP-REAL n),(TOP-REAL n) st len F = 0 & Product F = f & ( for k being Nat st k in dom F holds ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & F . k = Mx2Tran (Rotation (i,j,n,r)) ) ) holds f " is base_rotation let f be Function of (TOP-REAL n),(TOP-REAL n); ::_thesis: ( len F = 0 & Product F = f & ( for k being Nat st k in dom F holds ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & F . k = Mx2Tran (Rotation (i,j,n,r)) ) ) implies f " is base_rotation ) assume that A20: len F = 0 and A21: Product F = f ; ::_thesis: ( ex k being Nat st ( k in dom F & ( for i, j being Nat for r being real number holds ( not 1 <= i or not i < j or not j <= n or not F . k = Mx2Tran (Rotation (i,j,n,r)) ) ) ) or f " is base_rotation ) F = <*> the carrier of (GFuncs the carrier of (TOP-REAL n)) by A20; then A22: f = 1_ (GFuncs the carrier of (TOP-REAL n)) by A21, GROUP_4:8 .= the_unity_wrt the multF of (GFuncs the carrier of (TOP-REAL n)) by GROUP_1:22 .= id (TOP-REAL n) by MONOID_0:75 ; then rng f = [#] (TOP-REAL n) by TOPS_2:def_5; then f is onto by FUNCT_2:def_3; then f /" = f " by A22, TOPS_2:def_4; hence ( ex k being Nat st ( k in dom F & ( for i, j being Nat for r being real number holds ( not 1 <= i or not i < j or not j <= n or not F . k = Mx2Tran (Rotation (i,j,n,r)) ) ) ) or f " is base_rotation ) by A22, FUNCT_1:45; ::_thesis: verum end; for i being Nat holds S1[i] from NAT_1:sch_2(A19, A2); then S1[ len F] ; hence f " is base_rotation by A1; ::_thesis: verum end; end; registration let n be Nat; let f, g be rotation Function of (TOP-REAL n),(TOP-REAL n); clusterg (#) f -> rotation for Function of (TOP-REAL n),(TOP-REAL n); coherence for b1 being Function of (TOP-REAL n),(TOP-REAL n) st b1 = f * g holds b1 is rotation by Lm8; end; definition let n be Nat; let f be additive homogeneous Function of (TOP-REAL n),(TOP-REAL n); func AutMt f -> Matrix of n,F_Real means :Def6: :: MATRTOP3:def 6 f = Mx2Tran it; existence ex b1 being Matrix of n,F_Real st f = Mx2Tran b1 proof set T = n -VectSp_over F_Real; set TR = TOP-REAL n; reconsider B = MX2FinS (1. (F_Real,n)) as OrdBasis of n -VectSp_over F_Real by MATRLIN2:45; A1: the carrier of (n -VectSp_over F_Real) = REAL n by MATRIX13:102 .= the carrier of (TOP-REAL n) by EUCLID:22 ; then reconsider F = f as Function of (n -VectSp_over F_Real),(n -VectSp_over F_Real) ; now__::_thesis:_for_v1,_v2_being_Vector_of_(n_-VectSp_over_F_Real)_holds_F_._(v1_+_v2)_=_(F_._v1)_+_(F_._v2) let v1, v2 be Vector of (n -VectSp_over F_Real); ::_thesis: F . (v1 + v2) = (F . v1) + (F . v2) reconsider P1 = v1, P2 = v2, FP1 = F . v1, FP2 = F . v2 as Element of n -tuples_on the carrier of F_Real by MATRIX13:102; A2: ( @ (@ FP1) = FP1 & @ (@ FP2) = FP2 ) ; reconsider p1 = v1, p2 = v2 as Point of (TOP-REAL n) by A1; A3: ( @ (@ P1) = P1 & @ (@ P2) = P2 ) ; v1 + v2 = P1 + P2 by MATRIX13:102 .= p1 + p2 by A3, MATRTOP1:1 ; hence F . (v1 + v2) = (f . p1) + (f . p2) by VECTSP_1:def_20 .= FP1 + FP2 by A2, MATRTOP1:1 .= (F . v1) + (F . v2) by MATRIX13:102 ; ::_thesis: verum end; then A4: F is additive by VECTSP_1:def_20; len B = n by MATRTOP1:19; then reconsider A = AutMt (F,B,B) as Matrix of n,F_Real ; take A ; ::_thesis: f = Mx2Tran A now__::_thesis:_for_r_being_Scalar_of_F_Real for_v_being_Vector_of_(n_-VectSp_over_F_Real)_holds_F_._(r_*_v)_=_r_*_(F_._v) let r be Scalar of F_Real; ::_thesis: for v being Vector of (n -VectSp_over F_Real) holds F . (r * v) = r * (F . v) let v be Vector of (n -VectSp_over F_Real); ::_thesis: F . (r * v) = r * (F . v) reconsider p = v as Point of (TOP-REAL n) by A1; reconsider P = v, FP = F . v as Element of n -tuples_on the carrier of F_Real by MATRIX13:102; r * v = r * P by MATRIX13:102 .= r * p by MATRIXR1:17 ; hence F . (r * v) = r * (f . p) by TOPREAL9:def_4 .= r * FP by MATRIXR1:17 .= r * (F . v) by MATRIX13:102 ; ::_thesis: verum end; then A5: F is homogeneous by MOD_2:def_2; Mx2Tran A = Mx2Tran ((AutMt (F,B,B)),B,B) by MATRTOP1:20 .= f by A5, A4, MATRLIN2:34 ; hence f = Mx2Tran A ; ::_thesis: verum end; uniqueness for b1, b2 being Matrix of n,F_Real st f = Mx2Tran b1 & f = Mx2Tran b2 holds b1 = b2 by MATRTOP1:34; end; :: deftheorem Def6 defines AutMt MATRTOP3:def_6_:_ for n being Nat for f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) for b3 being Matrix of n,F_Real holds ( b3 = AutMt f iff f = Mx2Tran b3 ); theorem Th29: :: MATRTOP3:29 for n being Nat for f1, f2 being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) holds AutMt (f1 * f2) = (AutMt f2) * (AutMt f1) proof let n be Nat; ::_thesis: for f1, f2 being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) holds AutMt (f1 * f2) = (AutMt f2) * (AutMt f1) let f1, f2 be additive homogeneous Function of (TOP-REAL n),(TOP-REAL n); ::_thesis: AutMt (f1 * f2) = (AutMt f2) * (AutMt f1) set A1 = AutMt f1; set A2 = AutMt f2; A1: ( width (AutMt f1) = n & width (AutMt f2) = n & len (AutMt f2) = n ) by MATRIX_1:24; ( n = 0 implies n = 0 ) ; then Mx2Tran ((AutMt f2) * (AutMt f1)) = (Mx2Tran (AutMt f1)) * (Mx2Tran (AutMt f2)) by A1, MATRTOP1:32 .= f1 * (Mx2Tran (AutMt f2)) by Def6 .= f1 * f2 by Def6 ; hence AutMt (f1 * f2) = (AutMt f2) * (AutMt f1) by Def6; ::_thesis: verum end; theorem Th30: :: MATRTOP3:30 for X being set for k, n being Nat for p being Point of (TOP-REAL n) st k in X & k in Seg n holds ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is X -support-yielding & f is base_rotation & ( card (X /\ (Seg n)) > 1 implies (f . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg n) & i <> k holds (f . p) . i = 0 ) ) proof let X be set ; ::_thesis: for k, n being Nat for p being Point of (TOP-REAL n) st k in X & k in Seg n holds ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is X -support-yielding & f is base_rotation & ( card (X /\ (Seg n)) > 1 implies (f . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg n) & i <> k holds (f . p) . i = 0 ) ) let k, n be Nat; ::_thesis: for p being Point of (TOP-REAL n) st k in X & k in Seg n holds ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is X -support-yielding & f is base_rotation & ( card (X /\ (Seg n)) > 1 implies (f . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg n) & i <> k holds (f . p) . i = 0 ) ) let p be Point of (TOP-REAL n); ::_thesis: ( k in X & k in Seg n implies ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is X -support-yielding & f is base_rotation & ( card (X /\ (Seg n)) > 1 implies (f . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg n) & i <> k holds (f . p) . i = 0 ) ) ) assume that A1: k in X and A2: k in Seg n ; ::_thesis: ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is X -support-yielding & f is base_rotation & ( card (X /\ (Seg n)) > 1 implies (f . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg n) & i <> k holds (f . p) . i = 0 ) ) set TR = TOP-REAL n; defpred S1[ Nat] means ( $1 <= n implies ex f being base_rotation Function of (TOP-REAL n),(TOP-REAL n) st ( f is (X /\ (Seg $1)) \/ {k} -support-yielding & ( card ((X /\ (Seg $1)) \/ {k}) > 1 implies (f . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg $1) & i <> k holds (f . p) . i = 0 ) ) ); A3: for z being Nat st S1[z] holds S1[z + 1] proof let z be Nat; ::_thesis: ( S1[z] implies S1[z + 1] ) set z1 = z + 1; assume A4: S1[z] ; ::_thesis: S1[z + 1] A5: Seg (z + 1) = (Seg z) \/ {(z + 1)} by FINSEQ_1:9; A6: Seg (z + 1) = (Seg z) \/ {(z + 1)} by FINSEQ_1:9; A7: ( z + 1 in X implies ((X /\ (Seg z)) \/ {k}) \/ {(z + 1),k} = (X /\ (Seg (z + 1))) \/ {k} ) proof assume z + 1 in X ; ::_thesis: ((X /\ (Seg z)) \/ {k}) \/ {(z + 1),k} = (X /\ (Seg (z + 1))) \/ {k} then A8: X \/ {(z + 1)} = X by ZFMISC_1:40; {(z + 1),k} = {(z + 1)} \/ {k} by ENUMSET1:1; hence ((X /\ (Seg z)) \/ {k}) \/ {(z + 1),k} = (X /\ (Seg z)) \/ ({k} \/ ({k} \/ {(z + 1)})) by XBOOLE_1:4 .= (X /\ (Seg z)) \/ (({k} \/ {k}) \/ {(z + 1)}) by XBOOLE_1:4 .= ((X /\ (Seg z)) \/ {(z + 1)}) \/ {k} by XBOOLE_1:4 .= (X /\ (Seg (z + 1))) \/ {k} by A6, A8, XBOOLE_1:24 ; ::_thesis: verum end; assume A9: z + 1 <= n ; ::_thesis: ex f being base_rotation Function of (TOP-REAL n),(TOP-REAL n) st ( f is (X /\ (Seg (z + 1))) \/ {k} -support-yielding & ( card ((X /\ (Seg (z + 1))) \/ {k}) > 1 implies (f . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg (z + 1)) & i <> k holds (f . p) . i = 0 ) ) then consider f being base_rotation Function of (TOP-REAL n),(TOP-REAL n) such that A10: f is (X /\ (Seg z)) \/ {k} -support-yielding and A11: ( card ((X /\ (Seg z)) \/ {k}) > 1 implies (f . p) . k >= 0 ) and A12: for m being Nat st m in X /\ (Seg z) & m <> k holds (f . p) . m = 0 by A4, NAT_1:13; set z1 = z + 1; percases ( z + 1 = k or not z + 1 in X or ( z + 1 < k & z + 1 in X ) or ( z + 1 > k & z + 1 in X ) ) by XXREAL_0:1; supposeA13: z + 1 = k ; ::_thesis: ex f being base_rotation Function of (TOP-REAL n),(TOP-REAL n) st ( f is (X /\ (Seg (z + 1))) \/ {k} -support-yielding & ( card ((X /\ (Seg (z + 1))) \/ {k}) > 1 implies (f . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg (z + 1)) & i <> k holds (f . p) . i = 0 ) ) take f ; ::_thesis: ( f is (X /\ (Seg (z + 1))) \/ {k} -support-yielding & ( card ((X /\ (Seg (z + 1))) \/ {k}) > 1 implies (f . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg (z + 1)) & i <> k holds (f . p) . i = 0 ) ) (Seg (z + 1)) \/ {(z + 1)} = (Seg z) \/ ({(z + 1)} \/ {(z + 1)}) by A5, XBOOLE_1:4 .= (Seg z) \/ {(z + 1)} ; then A14: (X /\ (Seg (z + 1))) \/ {k} = (X \/ {k}) /\ ((Seg z) \/ {k}) by A13, XBOOLE_1:24 .= (X /\ (Seg z)) \/ {k} by XBOOLE_1:24 ; hence f is (X /\ (Seg (z + 1))) \/ {k} -support-yielding by A10; ::_thesis: ( ( card ((X /\ (Seg (z + 1))) \/ {k}) > 1 implies (f . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg (z + 1)) & i <> k holds (f . p) . i = 0 ) ) thus ( card ((X /\ (Seg (z + 1))) \/ {k}) > 1 implies (f . p) . k >= 0 ) by A11, A14; ::_thesis: for i being Nat st i in X /\ (Seg (z + 1)) & i <> k holds (f . p) . i = 0 let m be Nat; ::_thesis: ( m in X /\ (Seg (z + 1)) & m <> k implies (f . p) . m = 0 ) assume that A15: m in X /\ (Seg (z + 1)) and A16: m <> k ; ::_thesis: (f . p) . m = 0 A17: m in Seg (z + 1) by A15, XBOOLE_0:def_4; A18: m in X by A15, XBOOLE_0:def_4; not m in {(z + 1)} by A13, A16, TARSKI:def_1; then m in Seg z by A5, A17, XBOOLE_0:def_3; then m in X /\ (Seg z) by A18, XBOOLE_0:def_4; hence (f . p) . m = 0 by A12, A16; ::_thesis: verum end; supposeA19: not z + 1 in X ; ::_thesis: ex f being base_rotation Function of (TOP-REAL n),(TOP-REAL n) st ( f is (X /\ (Seg (z + 1))) \/ {k} -support-yielding & ( card ((X /\ (Seg (z + 1))) \/ {k}) > 1 implies (f . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg (z + 1)) & i <> k holds (f . p) . i = 0 ) ) take f ; ::_thesis: ( f is (X /\ (Seg (z + 1))) \/ {k} -support-yielding & ( card ((X /\ (Seg (z + 1))) \/ {k}) > 1 implies (f . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg (z + 1)) & i <> k holds (f . p) . i = 0 ) ) A20: {(z + 1)} misses X by A19, ZFMISC_1:50; A21: X /\ (Seg (z + 1)) = (X /\ (Seg z)) \/ (X /\ {(z + 1)}) by A5, XBOOLE_1:23 .= (X /\ (Seg z)) \/ {} by A20, XBOOLE_0:def_7 .= X /\ (Seg z) ; hence f is (X /\ (Seg (z + 1))) \/ {k} -support-yielding by A10; ::_thesis: ( ( card ((X /\ (Seg (z + 1))) \/ {k}) > 1 implies (f . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg (z + 1)) & i <> k holds (f . p) . i = 0 ) ) thus ( card ((X /\ (Seg (z + 1))) \/ {k}) > 1 implies (f . p) . k >= 0 ) by A11, A21; ::_thesis: for i being Nat st i in X /\ (Seg (z + 1)) & i <> k holds (f . p) . i = 0 let m be Nat; ::_thesis: ( m in X /\ (Seg (z + 1)) & m <> k implies (f . p) . m = 0 ) assume that A22: m in X /\ (Seg (z + 1)) and A23: m <> k ; ::_thesis: (f . p) . m = 0 A24: m in Seg (z + 1) by A22, XBOOLE_0:def_4; A25: m in X by A22, XBOOLE_0:def_4; then not m in {(z + 1)} by A19, TARSKI:def_1; then m in Seg z by A5, A24, XBOOLE_0:def_3; then m in X /\ (Seg z) by A25, XBOOLE_0:def_4; hence (f . p) . m = 0 by A12, A23; ::_thesis: verum end; supposeA26: ( z + 1 < k & z + 1 in X ) ; ::_thesis: ex f being base_rotation Function of (TOP-REAL n),(TOP-REAL n) st ( f is (X /\ (Seg (z + 1))) \/ {k} -support-yielding & ( card ((X /\ (Seg (z + 1))) \/ {k}) > 1 implies (f . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg (z + 1)) & i <> k holds (f . p) . i = 0 ) ) set fp = f . p; set S = (((f . p) . (z + 1)) ^2) + (((f . p) . k) ^2); A27: ( z + 1 >= 1 & k <= n ) by A2, FINSEQ_1:1, NAT_1:11; A28: ( ((f . p) . k) ^2 >= 0 & ((f . p) . (z + 1)) ^2 >= 0 ) by XREAL_1:63; then A29: (sqrt ((((f . p) . (z + 1)) ^2) + (((f . p) . k) ^2))) ^2 = (((f . p) . (z + 1)) ^2) + (((f . p) . k) ^2) by SQUARE_1:def_2; then consider r being real number such that A30: ((Mx2Tran (Rotation ((z + 1),k,n,r))) . (f . p)) . k = sqrt ((((f . p) . (z + 1)) ^2) + (((f . p) . k) ^2)) by A26, A27, Th25; reconsider M = Mx2Tran (Rotation ((z + 1),k,n,r)) as base_rotation Function of (TOP-REAL n),(TOP-REAL n) by A26, A27, Th28; take Mf = M * f; ::_thesis: ( Mf is (X /\ (Seg (z + 1))) \/ {k} -support-yielding & ( card ((X /\ (Seg (z + 1))) \/ {k}) > 1 implies (Mf . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg (z + 1)) & i <> k holds (Mf . p) . i = 0 ) ) A31: M is {(z + 1),k} -support-yielding by A26, A27, Th26; hence Mf is (X /\ (Seg (z + 1))) \/ {k} -support-yielding by A7, A10, A26; ::_thesis: ( ( card ((X /\ (Seg (z + 1))) \/ {k}) > 1 implies (Mf . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg (z + 1)) & i <> k holds (Mf . p) . i = 0 ) ) A32: dom Mf = the carrier of (TOP-REAL n) by FUNCT_2:52; then A33: f . p in dom M by FUNCT_1:11; A34: Mf . p = M . (f . p) by A32, FUNCT_1:12; hence ( card ((X /\ (Seg (z + 1))) \/ {k}) > 1 implies (Mf . p) . k >= 0 ) by A28, A30, SQUARE_1:def_2; ::_thesis: for i being Nat st i in X /\ (Seg (z + 1)) & i <> k holds (Mf . p) . i = 0 let i be Nat; ::_thesis: ( i in X /\ (Seg (z + 1)) & i <> k implies (Mf . p) . i = 0 ) assume that A35: i in X /\ (Seg (z + 1)) and A36: i <> k ; ::_thesis: (Mf . p) . i = 0 A37: i in X by A35, XBOOLE_0:def_4; i in Seg (z + 1) by A35, XBOOLE_0:def_4; then A38: ( i in Seg z or i in {(z + 1)} ) by A5, XBOOLE_0:def_3; percases ( i in Seg z or i = z + 1 ) by A38, TARSKI:def_1; supposeA39: i in Seg z ; ::_thesis: (Mf . p) . i = 0 then A40: i in X /\ (Seg z) by A37, XBOOLE_0:def_4; i <= z by A39, FINSEQ_1:1; then i < z + 1 by NAT_1:13; then not i in {(z + 1),k} by A36, TARSKI:def_2; hence (Mf . p) . i = (f . p) . i by A31, A33, A34, Def1 .= 0 by A12, A36, A40 ; ::_thesis: verum end; suppose i = z + 1 ; ::_thesis: (Mf . p) . i = 0 then A41: (((M . (f . p)) . i) * ((M . (f . p)) . i)) + ((((f . p) . (z + 1)) ^2) + (((f . p) . k) ^2)) = (((f . p) . (z + 1)) ^2) + (((f . p) . k) ^2) by A26, A27, A29, A30, Lm6; thus (Mf . p) . i = (M . (f . p)) . i by A32, FUNCT_1:12 .= 0 by A41 ; ::_thesis: verum end; end; end; supposeA42: ( z + 1 > k & z + 1 in X ) ; ::_thesis: ex f being base_rotation Function of (TOP-REAL n),(TOP-REAL n) st ( f is (X /\ (Seg (z + 1))) \/ {k} -support-yielding & ( card ((X /\ (Seg (z + 1))) \/ {k}) > 1 implies (f . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg (z + 1)) & i <> k holds (f . p) . i = 0 ) ) set fp = f . p; set S = (((f . p) . (z + 1)) ^2) + (((f . p) . k) ^2); A43: 1 <= k by A2, FINSEQ_1:1; A44: ( ((f . p) . k) ^2 >= 0 & ((f . p) . (z + 1)) ^2 >= 0 ) by XREAL_1:63; then A45: (sqrt ((((f . p) . (z + 1)) ^2) + (((f . p) . k) ^2))) ^2 = (((f . p) . (z + 1)) ^2) + (((f . p) . k) ^2) by SQUARE_1:def_2; then consider r being real number such that A46: ((Mx2Tran (Rotation (k,(z + 1),n,r))) . (f . p)) . k = sqrt ((((f . p) . (z + 1)) ^2) + (((f . p) . k) ^2)) by A9, A42, A43, Th24; reconsider M = Mx2Tran (Rotation (k,(z + 1),n,r)) as base_rotation Function of (TOP-REAL n),(TOP-REAL n) by A9, A42, A43, Th28; take Mf = M * f; ::_thesis: ( Mf is (X /\ (Seg (z + 1))) \/ {k} -support-yielding & ( card ((X /\ (Seg (z + 1))) \/ {k}) > 1 implies (Mf . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg (z + 1)) & i <> k holds (Mf . p) . i = 0 ) ) A47: M is {k,(z + 1)} -support-yielding by A9, A42, A43, Th26; hence Mf is (X /\ (Seg (z + 1))) \/ {k} -support-yielding by A7, A10, A42; ::_thesis: ( ( card ((X /\ (Seg (z + 1))) \/ {k}) > 1 implies (Mf . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg (z + 1)) & i <> k holds (Mf . p) . i = 0 ) ) A48: dom Mf = the carrier of (TOP-REAL n) by FUNCT_2:52; then A49: Mf . p = M . (f . p) by FUNCT_1:12; hence ( card ((X /\ (Seg (z + 1))) \/ {k}) > 1 implies (Mf . p) . k >= 0 ) by A44, A46, SQUARE_1:def_2; ::_thesis: for i being Nat st i in X /\ (Seg (z + 1)) & i <> k holds (Mf . p) . i = 0 let i be Nat; ::_thesis: ( i in X /\ (Seg (z + 1)) & i <> k implies (Mf . p) . i = 0 ) assume that A50: i in X /\ (Seg (z + 1)) and A51: i <> k ; ::_thesis: (Mf . p) . i = 0 A52: i in X by A50, XBOOLE_0:def_4; i in Seg (z + 1) by A50, XBOOLE_0:def_4; then A53: ( i in Seg z or i in {(z + 1)} ) by A5, XBOOLE_0:def_3; percases ( i in Seg z or i = z + 1 ) by A53, TARSKI:def_1; supposeA54: i in Seg z ; ::_thesis: (Mf . p) . i = 0 then i <= z by FINSEQ_1:1; then i < z + 1 by NAT_1:13; then A55: not i in {(z + 1),k} by A51, TARSKI:def_2; A56: i in X /\ (Seg z) by A52, A54, XBOOLE_0:def_4; f . p in dom M by A48, FUNCT_1:11; hence (Mf . p) . i = (f . p) . i by A47, A49, A55, Def1 .= 0 by A12, A51, A56 ; ::_thesis: verum end; suppose i = z + 1 ; ::_thesis: (Mf . p) . i = 0 then A57: (((M . (f . p)) . i) * ((M . (f . p)) . i)) + ((((f . p) . (z + 1)) ^2) + (((f . p) . k) ^2)) = (((f . p) . (z + 1)) ^2) + (((f . p) . k) ^2) by A9, A42, A43, A45, A46, Lm6; thus (Mf . p) . i = (M . (f . p)) . i by A48, FUNCT_1:12 .= 0 by A57 ; ::_thesis: verum end; end; end; end; end; A58: S1[ 0 ] proof assume 0 <= n ; ::_thesis: ex f being base_rotation Function of (TOP-REAL n),(TOP-REAL n) st ( f is (X /\ (Seg 0)) \/ {k} -support-yielding & ( card ((X /\ (Seg 0)) \/ {k}) > 1 implies (f . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg 0) & i <> k holds (f . p) . i = 0 ) ) take f = id (TOP-REAL n); ::_thesis: ( f is (X /\ (Seg 0)) \/ {k} -support-yielding & ( card ((X /\ (Seg 0)) \/ {k}) > 1 implies (f . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg 0) & i <> k holds (f . p) . i = 0 ) ) A59: f is {} -support-yielding proof let F be Function; :: according to MATRTOP3:def_1 ::_thesis: for x being set st F in dom f & (f . F) . x <> F . x holds x in {} let x be set ; ::_thesis: ( F in dom f & (f . F) . x <> F . x implies x in {} ) assume that A60: F in dom f and A61: (f . F) . x <> F . x ; ::_thesis: x in {} F in the carrier of (TOP-REAL n) by A60, FUNCT_1:17; hence x in {} by A61, FUNCT_1:17; ::_thesis: verum end; {} c= (X /\ (Seg 0)) \/ {k} by XBOOLE_1:2; hence f is (X /\ (Seg 0)) \/ {k} -support-yielding by A59; ::_thesis: ( ( card ((X /\ (Seg 0)) \/ {k}) > 1 implies (f . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg 0) & i <> k holds (f . p) . i = 0 ) ) thus ( card ((X /\ (Seg 0)) \/ {k}) > 1 implies (f . p) . k >= 0 ) by CARD_2:42; ::_thesis: for i being Nat st i in X /\ (Seg 0) & i <> k holds (f . p) . i = 0 let m be Nat; ::_thesis: ( m in X /\ (Seg 0) & m <> k implies (f . p) . m = 0 ) assume m in X /\ (Seg 0) ; ::_thesis: ( not m <> k or (f . p) . m = 0 ) hence ( not m <> k or (f . p) . m = 0 ) ; ::_thesis: verum end; for z being Nat holds S1[z] from NAT_1:sch_2(A58, A3); then consider f being base_rotation Function of (TOP-REAL n),(TOP-REAL n) such that A62: ( f is (X /\ (Seg n)) \/ {k} -support-yielding & ( card ((X /\ (Seg n)) \/ {k}) > 1 implies (f . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg n) & i <> k holds (f . p) . i = 0 ) ) ; take f ; ::_thesis: ( f is X -support-yielding & f is base_rotation & ( card (X /\ (Seg n)) > 1 implies (f . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg n) & i <> k holds (f . p) . i = 0 ) ) A63: X /\ (Seg n) c= X by XBOOLE_1:17; ( {k} c= X & {k} c= Seg n ) by A1, A2, ZFMISC_1:31; then (X /\ (Seg n)) \/ {k} = X /\ (Seg n) by XBOOLE_1:12, XBOOLE_1:19; hence ( f is X -support-yielding & f is base_rotation & ( card (X /\ (Seg n)) > 1 implies (f . p) . k >= 0 ) & ( for i being Nat st i in X /\ (Seg n) & i <> k holds (f . p) . i = 0 ) ) by A62, A63; ::_thesis: verum end; theorem Th31: :: MATRTOP3:31 for n being Nat for f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) for A being Subset of (TOP-REAL n) st f | A = id A holds f | (Lin A) = id (Lin A) proof let n be Nat; ::_thesis: for f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) for A being Subset of (TOP-REAL n) st f | A = id A holds f | (Lin A) = id (Lin A) let f be additive homogeneous Function of (TOP-REAL n),(TOP-REAL n); ::_thesis: for A being Subset of (TOP-REAL n) st f | A = id A holds f | (Lin A) = id (Lin A) set TR = TOP-REAL n; let A be Subset of (TOP-REAL n); ::_thesis: ( f | A = id A implies f | (Lin A) = id (Lin A) ) assume A1: f | A = id A ; ::_thesis: f | (Lin A) = id (Lin A) defpred S1[ Nat] means for L being Linear_Combination of A st card (Carrier L) <= $1 holds f . (Sum L) = Sum L; A2: for i being Nat st S1[i] holds S1[i + 1] proof let i be Nat; ::_thesis: ( S1[i] implies S1[i + 1] ) assume A3: S1[i] ; ::_thesis: S1[i + 1] set i1 = i + 1; let L be Linear_Combination of A; ::_thesis: ( card (Carrier L) <= i + 1 implies f . (Sum L) = Sum L ) assume A4: card (Carrier L) <= i + 1 ; ::_thesis: f . (Sum L) = Sum L percases ( card (Carrier L) <= i or card (Carrier L) = i + 1 ) by A4, NAT_1:8; suppose card (Carrier L) <= i ; ::_thesis: f . (Sum L) = Sum L hence f . (Sum L) = Sum L by A3; ::_thesis: verum end; supposeA5: card (Carrier L) = i + 1 ; ::_thesis: f . (Sum L) = Sum L then not Carrier L is empty ; then consider x being set such that A6: x in Carrier L by XBOOLE_0:def_1; reconsider x = x as Point of (TOP-REAL n) by A6; reconsider X = {x} as Subset of (TOP-REAL n) by ZFMISC_1:31; consider K being Linear_Combination of X such that A7: K . x = L . x by RLVECT_4:37; L . x <> 0 by A6, RLVECT_2:19; then ( Carrier K c= {x} & x in Carrier K ) by A7, RLVECT_2:19, RLVECT_2:def_6; then A8: Carrier K = {x} by ZFMISC_1:33; {x} \/ (Carrier L) = Carrier L by A6, ZFMISC_1:40; then A9: Carrier (L - K) c= Carrier L by A8, RLVECT_2:55; (L - K) . x = (L . x) - (K . x) by RLVECT_2:54 .= 0 by A7 ; then not x in Carrier (L - K) by RLVECT_2:19; then Carrier (L - K) c< Carrier L by A6, A9, XBOOLE_0:def_8; then card (Carrier (L - K)) < i + 1 by A5, CARD_2:48; then A10: card (Carrier (L - K)) <= i by NAT_1:13; ZeroLC (TOP-REAL n) = (- K) - (- K) by RLVECT_2:57 .= (- K) + (- (- K)) by RLVECT_2:def_13 .= (- K) + K by RLVECT_2:53 ; then L = L + ((- K) + K) by RLVECT_2:41 .= (L + (- K)) + K by RLVECT_2:40 .= (L - K) + K by RLVECT_2:def_13 ; then A11: Sum L = (Sum (L - K)) + (Sum K) by RLVECT_3:1; A12: Carrier L c= A by RLVECT_2:def_6; then (f | A) . x = f . x by A6, FUNCT_1:49; then A13: f . x = x by A1, A6, A12, FUNCT_1:17; Carrier (L - K) c= A by A9, A12, XBOOLE_1:1; then L - K is Linear_Combination of A by RLVECT_2:def_6; then A14: f . (Sum (L - K)) = Sum (L - K) by A3, A10; Sum K = (L . x) * x by A7, RLVECT_2:32; then f . (Sum K) = Sum K by A13, TOPREAL9:def_4; hence f . (Sum L) = Sum L by A11, A14, VECTSP_1:def_20; ::_thesis: verum end; end; end; set L = Lin A; set cL = the carrier of (Lin A); the carrier of (Lin A) c= the carrier of (TOP-REAL n) by RLSUB_1:def_2; then A15: f | (Lin A) = f | the carrier of (Lin A) by TMAP_1:def_3; A16: S1[ 0 ] proof let L be Linear_Combination of A; ::_thesis: ( card (Carrier L) <= 0 implies f . (Sum L) = Sum L ) assume card (Carrier L) <= 0 ; ::_thesis: f . (Sum L) = Sum L then Carrier L = {} ; then L is Linear_Combination of {} the carrier of (TOP-REAL n) by RLVECT_2:def_6; then A17: Sum L = 0. (TOP-REAL n) by RLVECT_2:31; f . (0. (TOP-REAL n)) = f . (0 * (0. (TOP-REAL n))) by RLVECT_1:10 .= 0 * (f . (0. (TOP-REAL n))) by TOPREAL9:def_4 .= 0. (TOP-REAL n) by RLVECT_1:10 ; hence f . (Sum L) = Sum L by A17; ::_thesis: verum end; A18: for i being Nat holds S1[i] from NAT_1:sch_2(A16, A2); A19: for x being set st x in the carrier of (Lin A) holds (f | (Lin A)) . x = (id (Lin A)) . x proof let x be set ; ::_thesis: ( x in the carrier of (Lin A) implies (f | (Lin A)) . x = (id (Lin A)) . x ) assume A20: x in the carrier of (Lin A) ; ::_thesis: (f | (Lin A)) . x = (id (Lin A)) . x then x in Lin A by STRUCT_0:def_5; then consider K being Linear_Combination of A such that A21: Sum K = x by RLVECT_3:14; S1[ card (Carrier K)] by A18; then A22: f . x = x by A21; (f | (Lin A)) . x = f . x by A15, A20, FUNCT_1:49; hence (f | (Lin A)) . x = (id (Lin A)) . x by A20, A22, FUNCT_1:17; ::_thesis: verum end; ( dom (f | (Lin A)) = the carrier of (Lin A) & dom (id (Lin A)) = the carrier of (Lin A) ) by FUNCT_2:def_1; hence f | (Lin A) = id (Lin A) by A19, FUNCT_1:2; ::_thesis: verum end; theorem Th32: :: MATRTOP3:32 for n being Nat for p being Point of (TOP-REAL n) for f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) for A being Subset of (TOP-REAL n) st f is rotation & f | A = id A holds for i being Nat st i in Seg n & Base_FinSeq (n,i) in Lin A holds (f . p) . i = p . i proof let n be Nat; ::_thesis: for p being Point of (TOP-REAL n) for f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) for A being Subset of (TOP-REAL n) st f is rotation & f | A = id A holds for i being Nat st i in Seg n & Base_FinSeq (n,i) in Lin A holds (f . p) . i = p . i let p be Point of (TOP-REAL n); ::_thesis: for f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) for A being Subset of (TOP-REAL n) st f is rotation & f | A = id A holds for i being Nat st i in Seg n & Base_FinSeq (n,i) in Lin A holds (f . p) . i = p . i let f be additive homogeneous Function of (TOP-REAL n),(TOP-REAL n); ::_thesis: for A being Subset of (TOP-REAL n) st f is rotation & f | A = id A holds for i being Nat st i in Seg n & Base_FinSeq (n,i) in Lin A holds (f . p) . i = p . i set TR = TOP-REAL n; let A be Subset of (TOP-REAL n); ::_thesis: ( f is rotation & f | A = id A implies for i being Nat st i in Seg n & Base_FinSeq (n,i) in Lin A holds (f . p) . i = p . i ) assume that A1: f is rotation and A2: f | A = id A ; ::_thesis: for i being Nat st i in Seg n & Base_FinSeq (n,i) in Lin A holds (f . p) . i = p . i set L = Lin A; set n0 = 0* n; A3: f | (Lin A) = id (Lin A) by A2, Th31; A4: len (0* n) = n by CARD_1:def_7; then A5: dom (0* n) = Seg n by FINSEQ_1:def_3; let k be Nat; ::_thesis: ( k in Seg n & Base_FinSeq (n,k) in Lin A implies (f . p) . k = p . k ) assume that A6: k in Seg n and A7: Base_FinSeq (n,k) in Lin A ; ::_thesis: (f . p) . k = p . k set n0k = (0* n) +* (k,1); A8: (0* n) +* (k,1) = Base_FinSeq (n,k) by MATRIXR2:def_4; set pk = p . k; the carrier of (Lin A) c= the carrier of (TOP-REAL n) by RLSUB_1:def_2; then A9: f | (Lin A) = f | the carrier of (Lin A) by TMAP_1:def_3; dom ((0* n) +* (k,1)) = dom (0* n) by FUNCT_7:30; then A10: len ((0* n) +* (k,1)) = n by A4, FINSEQ_3:29; A11: (0* n) +* (k,1) = @ (@ ((0* n) +* (k,1))) ; then reconsider N0k = (0* n) +* (k,1) as Point of (TOP-REAL n) by A10, TOPREAL3:46; A12: (0* n) +* (k,1) is Element of n -tuples_on REAL by A10, A11, FINSEQ_2:92; A13: for p being Point of (TOP-REAL n) st p . k <> 0 holds (f . p) . k = p . k proof let p be Point of (TOP-REAL n); ::_thesis: ( p . k <> 0 implies (f . p) . k = p . k ) assume A14: p . k <> 0 ; ::_thesis: (f . p) . k = p . k set fp = f . p; set pk = p . k; set pN = (p . k) * N0k; set ppN = p - ((p . k) * N0k); A15: (f . (p - ((p . k) * N0k))) + (f . ((p . k) * N0k)) = f . ((p - ((p . k) * N0k)) + ((p . k) * N0k)) by VECTSP_1:def_20; ( len (f . (p - ((p . k) * N0k))) = n & len (f . ((p . k) * N0k)) = n ) by CARD_1:def_7; then A16: |.(f . ((p - ((p . k) * N0k)) + ((p . k) * N0k))).| ^2 = ((|.(f . (p - ((p . k) * N0k))).| ^2) + (2 * |((f . ((p . k) * N0k)),(f . (p - ((p . k) * N0k))))|)) + (|.(f . ((p . k) * N0k)).| ^2) by A15, EUCLID_2:11; A17: (n |-> (p . k)) . k = p . k by A6, FINSEQ_2:57; A18: (p . k) * ((0* n) +* (k,1)) = mlt ((n |-> (p . k)),((0* n) +* (k,1))) by A12, RVSUM_1:63 .= (0* n) +* (k,((p . k) * 1)) by A17, TOPREALC:15 .= (0* n) +* (k,(p . k)) ; len (f . p) = n by CARD_1:def_7; then A19: dom (f . p) = Seg n by FINSEQ_1:def_3; A20: len (p - ((p . k) * N0k)) = n by CARD_1:def_7; then dom (p - ((p . k) * N0k)) = Seg n by FINSEQ_1:def_3; then A21: (p - ((p . k) * N0k)) . k = (p . k) - (((p . k) * N0k) . k) by A6, VALUED_1:13; len ((p . k) * N0k) = n by CARD_1:def_7; then A22: |.((p - ((p . k) * N0k)) + ((p . k) * N0k)).| ^2 = ((|.(p - ((p . k) * N0k)).| ^2) + (2 * |(((p . k) * N0k),(p - ((p . k) * N0k)))|)) + (|.((p . k) * N0k).| ^2) by A20, EUCLID_2:11; (p . k) * N0k in Lin A by A7, A8, RLSUB_1:21; then A23: (p . k) * N0k in the carrier of (Lin A) by STRUCT_0:def_5; then A24: (p . k) * N0k = (f | (Lin A)) . ((p . k) * N0k) by A3, FUNCT_1:17 .= f . ((p . k) * N0k) by A9, A23, FUNCT_1:49 ; ( |.((p - ((p . k) * N0k)) + ((p . k) * N0k)).| = |.(f . ((p - ((p . k) * N0k)) + ((p . k) * N0k))).| & |.(p - ((p . k) * N0k)).| = |.(f . (p - ((p . k) * N0k))).| ) by A1, Def4; then |(((p . k) * N0k),(p - ((p . k) * N0k)))| = (p . k) * ((f . (p - ((p . k) * N0k))) . k) by A18, A16, A22, A24, TOPREALC:16; then A25: (p . k) * ((p - ((p . k) * N0k)) . k) = (p . k) * ((f . (p - ((p . k) * N0k))) . k) by A18, TOPREALC:16; ((p . k) * N0k) . k = p . k by A6, A5, A18, FUNCT_7:31; then A26: (f . (p - ((p . k) * N0k))) . k = 0 by A14, A21, A25; (p - ((p . k) * N0k)) + ((p . k) * N0k) = p - (((p . k) * N0k) - ((p . k) * N0k)) by EUCLID:47 .= p - (0. (TOP-REAL n)) by EUCLID:42 .= p by RLVECT_1:13 ; then (f . p) . k = ((f . (p - ((p . k) * N0k))) . k) + ((f . ((p . k) * N0k)) . k) by A6, A15, A19, VALUED_1:def_1 .= ((f . (p - ((p . k) * N0k))) . k) + (p . k) by A6, A5, A18, A24, FUNCT_7:31 ; hence (f . p) . k = p . k by A26; ::_thesis: verum end; percases ( p . k <> 0 or p . k = 0 ) ; suppose p . k <> 0 ; ::_thesis: (f . p) . k = p . k hence (f . p) . k = p . k by A13; ::_thesis: verum end; suppose p . k = 0 ; ::_thesis: (f . p) . k = p . k len (f . p) = n by CARD_1:def_7; then A27: |.((f . p) + N0k).| ^2 = ((|.(f . p).| ^2) + (2 * |(N0k,(f . p))|)) + (|.N0k.| ^2) by A10, EUCLID_2:11; len p = n by CARD_1:def_7; then A28: |.(p + N0k).| ^2 = ((|.p.| ^2) + (2 * |(N0k,p)|)) + (|.N0k.| ^2) by A10, EUCLID_2:11; A29: N0k in the carrier of (Lin A) by A7, A8, STRUCT_0:def_5; then N0k = (f | (Lin A)) . N0k by A3, FUNCT_1:17 .= f . N0k by A9, A29, FUNCT_1:49 ; then A30: f . (p + N0k) = (f . p) + N0k by VECTSP_1:def_20; ( |.(p + N0k).| = |.(f . (p + N0k)).| & |.(f . p).| = |.p.| ) by A1, Def4; then |(N0k,(f . p))| = 1 * (p . k) by A27, A28, A30, TOPREALC:16; then p . k = 1 * ((f . p) . k) by TOPREALC:16; hence (f . p) . k = p . k ; ::_thesis: verum end; end; end; theorem Th33: :: MATRTOP3:33 for X being set for n being Nat for p being Point of (TOP-REAL n) for f being rotation Function of (TOP-REAL n),(TOP-REAL n) st f is X -support-yielding & ( for i being Nat st i in X /\ (Seg n) holds p . i = 0 ) holds f . p = p proof let X be set ; ::_thesis: for n being Nat for p being Point of (TOP-REAL n) for f being rotation Function of (TOP-REAL n),(TOP-REAL n) st f is X -support-yielding & ( for i being Nat st i in X /\ (Seg n) holds p . i = 0 ) holds f . p = p let n be Nat; ::_thesis: for p being Point of (TOP-REAL n) for f being rotation Function of (TOP-REAL n),(TOP-REAL n) st f is X -support-yielding & ( for i being Nat st i in X /\ (Seg n) holds p . i = 0 ) holds f . p = p let p be Point of (TOP-REAL n); ::_thesis: for f being rotation Function of (TOP-REAL n),(TOP-REAL n) st f is X -support-yielding & ( for i being Nat st i in X /\ (Seg n) holds p . i = 0 ) holds f . p = p set TR = TOP-REAL n; let f be rotation Function of (TOP-REAL n),(TOP-REAL n); ::_thesis: ( f is X -support-yielding & ( for i being Nat st i in X /\ (Seg n) holds p . i = 0 ) implies f . p = p ) assume that A1: f is X -support-yielding and A2: for i being Nat st i in X /\ (Seg n) holds p . i = 0 ; ::_thesis: f . p = p set sp = sqr p; set sfp = sqr (f . p); A3: Sum (sqr p) >= 0 by RVSUM_1:86; ( Sum (sqr (f . p)) >= 0 & |.(f . p).| = |.p.| ) by Def4, RVSUM_1:86; then A4: Sum (sqr p) = Sum (sqr (f . p)) by A3, SQUARE_1:28; A5: len p = n by CARD_1:def_7; p = @ (@ p) ; then A6: len (sqr p) = n by A5, RVSUM_1:143; A7: len (f . p) = n by CARD_1:def_7; @ (@ (f . p)) = f . p ; then len (sqr (f . p)) = n by A7, RVSUM_1:143; then reconsider sp = sqr p, sfp = sqr (f . p) as Element of n -tuples_on REAL by A6, FINSEQ_2:92; A8: dom f = the carrier of (TOP-REAL n) by FUNCT_2:52; A9: for i being Nat st i in Seg n holds sp . i <= sfp . i proof let i be Nat; ::_thesis: ( i in Seg n implies sp . i <= sfp . i ) A10: ( sp . i = (p . i) ^2 & sfp . i = ((f . p) . i) ^2 ) by VALUED_1:11; assume A11: i in Seg n ; ::_thesis: sp . i <= sfp . i percases ( i in X or not i in X ) ; suppose i in X ; ::_thesis: sp . i <= sfp . i then i in X /\ (Seg n) by A11, XBOOLE_0:def_4; then p . i = 0 by A2; hence sp . i <= sfp . i by A10, XREAL_1:63; ::_thesis: verum end; suppose not i in X ; ::_thesis: sp . i <= sfp . i hence sp . i <= sfp . i by A1, A8, A10, Def1; ::_thesis: verum end; end; end; for i being Nat st 1 <= i & i <= n holds p . i = (f . p) . i proof let i be Nat; ::_thesis: ( 1 <= i & i <= n implies p . i = (f . p) . i ) A12: sp . i = (p . i) ^2 by VALUED_1:11; assume ( 1 <= i & i <= n ) ; ::_thesis: p . i = (f . p) . i then A13: i in Seg n by FINSEQ_1:1; then A14: ( sp . i >= sfp . i & sp . i <= sfp . i ) by A4, A9, RVSUM_1:83; percases ( i in X or not i in X ) ; suppose i in X ; ::_thesis: p . i = (f . p) . i then A15: i in X /\ (Seg n) by A13, XBOOLE_0:def_4; then p . i = 0 by A2; then ((f . p) . i) ^2 = 0 by A12, A14, VALUED_1:11; then (f . p) . i = 0 ; hence p . i = (f . p) . i by A2, A15; ::_thesis: verum end; suppose not i in X ; ::_thesis: p . i = (f . p) . i hence p . i = (f . p) . i by A1, A8, Def1; ::_thesis: verum end; end; end; hence f . p = p by A5, A7, FINSEQ_1:14; ::_thesis: verum end; theorem Th34: :: MATRTOP3:34 for i, n being Nat for p being Point of (TOP-REAL n) st i in Seg n & n >= 2 holds ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is base_rotation & f . p = p +* (i,(- (p . i))) ) proof let i, n be Nat; ::_thesis: for p being Point of (TOP-REAL n) st i in Seg n & n >= 2 holds ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is base_rotation & f . p = p +* (i,(- (p . i))) ) let p be Point of (TOP-REAL n); ::_thesis: ( i in Seg n & n >= 2 implies ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is base_rotation & f . p = p +* (i,(- (p . i))) ) ) set TR = TOP-REAL n; assume that A1: i in Seg n and A2: n >= 2 ; ::_thesis: ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is base_rotation & f . p = p +* (i,(- (p . i))) ) A3: {i} c= Seg n by A1, ZFMISC_1:31; A4: 1 <= i by A1, FINSEQ_1:1; ( card (Seg n) = n & card {i} = 1 ) by CARD_2:42, FINSEQ_1:57; then {i} <> Seg n by A2; then {i} c< Seg n by A3, XBOOLE_0:def_8; then consider j being set such that A5: j in Seg n and A6: not j in {i} by XBOOLE_0:6; reconsider j = j as Nat by A5; A7: j <> i by A6, TARSKI:def_1; A8: 1 <= j by A5, FINSEQ_1:1; set p0 = p +* (i,(- (p . i))); A9: len (p +* (i,(- (p . i)))) = len p by FUNCT_7:97; A10: len p = n by CARD_1:def_7; then A11: dom p = Seg n by FINSEQ_1:def_3; A12: i <= n by A1, FINSEQ_1:1; ( (p . i) * (p . i) >= 0 & (p . j) * (p . j) >= 0 ) by XREAL_1:63; then A13: ( 0 ^2 = 0* 0 & 0 <= ((p . i) ^2) + ((p . j) ^2) ) ; A14: j <= n by A5, FINSEQ_1:1; percases ( i < j or j < i ) by A7, XXREAL_0:1; supposeA15: i < j ; ::_thesis: ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is base_rotation & f . p = p +* (i,(- (p . i))) ) then consider r being real number such that A16: ((Mx2Tran (Rotation (i,j,n,r))) . p) . i = 0 by A4, A13, A14, Th24; set s = sin r; set c = cos r; A17: 0 = ((p . i) * (cos r)) + ((p . j) * (- (sin r))) by A4, A14, A15, A16, Th21; reconsider M = Mx2Tran (Rotation (i,j,n,(r + r))) as base_rotation Function of (TOP-REAL n),(TOP-REAL n) by A4, A14, A15, Th28; set Mp = M . p; A18: ( cos (r + r) = ((cos r) * (cos r)) - ((sin r) * (sin r)) & sin (r + r) = ((sin r) * (cos r)) + ((sin r) * (cos r)) ) by SIN_COS:75; A19: M is {i,j} -support-yielding by A4, A14, A15, Th26; A20: for k being Nat st 1 <= k & k <= n holds (p +* (i,(- (p . i)))) . k = (M . p) . k proof let k be Nat; ::_thesis: ( 1 <= k & k <= n implies (p +* (i,(- (p . i)))) . k = (M . p) . k ) assume ( 1 <= k & k <= n ) ; ::_thesis: (p +* (i,(- (p . i)))) . k = (M . p) . k then A21: k in Seg n by FINSEQ_1:1; percases ( i = k or j = k or ( i <> k & j <> k ) ) ; supposeA22: i = k ; ::_thesis: (p +* (i,(- (p . i)))) . k = (M . p) . k hence (p +* (i,(- (p . i)))) . k = - ((p . i) * 1) by A11, A21, FUNCT_7:31 .= - ((p . i) * (((sin r) * (sin r)) + ((cos r) * (cos r)))) by SIN_COS:29 .= (((p . i) * (((cos r) * (cos r)) - ((sin r) * (sin r)))) - (((p . i) * (cos r)) * (cos r))) - (((p . i) * (cos r)) * (cos r)) .= (((p . i) * (((cos r) * (cos r)) - ((sin r) * (sin r)))) - (((p . j) * (sin r)) * (cos r))) - (((p . j) * (sin r)) * (cos r)) by A17 .= ((p . i) * (((cos r) * (cos r)) - ((sin r) * (sin r)))) + ((p . j) * (- (((sin r) * (cos r)) + ((sin r) * (cos r))))) .= (M . p) . k by A4, A14, A15, A18, A22, Th21 ; ::_thesis: verum end; supposeA23: j = k ; ::_thesis: (p +* (i,(- (p . i)))) . k = (M . p) . k hence (p +* (i,(- (p . i)))) . k = (p . j) * 1 by A15, FUNCT_7:32 .= (p . j) * (((sin r) * (sin r)) + ((cos r) * (cos r))) by SIN_COS:29 .= ((((p . j) * (sin r)) * (sin r)) + (((p . j) * (sin r)) * (sin r))) + ((p . j) * (((cos r) * (cos r)) - ((sin r) * (sin r)))) .= ((((p . i) * (cos r)) * (sin r)) + (((p . i) * (cos r)) * (sin r))) + ((p . j) * (((cos r) * (cos r)) - ((sin r) * (sin r)))) by A17 .= ((p . i) * (((sin r) * (cos r)) + ((sin r) * (cos r)))) + ((p . j) * (((cos r) * (cos r)) - ((sin r) * (sin r)))) .= (M . p) . k by A4, A14, A15, A18, A23, Th22 ; ::_thesis: verum end; supposeA24: ( i <> k & j <> k ) ; ::_thesis: (p +* (i,(- (p . i)))) . k = (M . p) . k A25: dom M = the carrier of (TOP-REAL n) by FUNCT_2:52; ( (p +* (i,(- (p . i)))) . k = p . k & not k in {i,j} ) by A24, FUNCT_7:32, TARSKI:def_2; hence (p +* (i,(- (p . i)))) . k = (M . p) . k by A19, A25, Def1; ::_thesis: verum end; end; end; take M ; ::_thesis: ( M is base_rotation & M . p = p +* (i,(- (p . i))) ) len (M . p) = n by CARD_1:def_7; hence ( M is base_rotation & M . p = p +* (i,(- (p . i))) ) by A9, A10, A20, FINSEQ_1:14; ::_thesis: verum end; supposeA26: j < i ; ::_thesis: ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is base_rotation & f . p = p +* (i,(- (p . i))) ) then consider r being real number such that A27: ((Mx2Tran (Rotation (j,i,n,r))) . p) . i = 0 by A8, A12, A13, Th25; set s = sin r; set c = cos r; A28: 0 = ((p . j) * (sin r)) + ((p . i) * (cos r)) by A8, A12, A26, A27, Th22; reconsider M = Mx2Tran (Rotation (j,i,n,(r + r))) as base_rotation Function of (TOP-REAL n),(TOP-REAL n) by A8, A12, A26, Th28; set Mp = M . p; A29: ( cos (r + r) = ((cos r) * (cos r)) - ((sin r) * (sin r)) & sin (r + r) = ((sin r) * (cos r)) + ((sin r) * (cos r)) ) by SIN_COS:75; A30: M is {i,j} -support-yielding by A8, A12, A26, Th26; A31: for k being Nat st 1 <= k & k <= n holds (p +* (i,(- (p . i)))) . k = (M . p) . k proof let k be Nat; ::_thesis: ( 1 <= k & k <= n implies (p +* (i,(- (p . i)))) . k = (M . p) . k ) assume ( 1 <= k & k <= n ) ; ::_thesis: (p +* (i,(- (p . i)))) . k = (M . p) . k then A32: k in Seg n by FINSEQ_1:1; percases ( i = k or j = k or ( i <> k & j <> k ) ) ; supposeA33: i = k ; ::_thesis: (p +* (i,(- (p . i)))) . k = (M . p) . k hence (p +* (i,(- (p . i)))) . k = - ((p . i) * 1) by A11, A32, FUNCT_7:31 .= - ((p . i) * (((sin r) * (sin r)) + ((cos r) * (cos r)))) by SIN_COS:29 .= ((- (((p . i) * (cos r)) * (cos r))) + ((- ((p . i) * (cos r))) * (cos r))) + ((p . i) * (((cos r) * (cos r)) - ((sin r) * (sin r)))) .= ((((p . j) * (sin r)) * (cos r)) + (((p . j) * (sin r)) * (cos r))) + ((p . i) * (((cos r) * (cos r)) - ((sin r) * (sin r)))) by A28 .= ((p . j) * (((sin r) * (cos r)) + ((sin r) * (cos r)))) + ((p . i) * (((cos r) * (cos r)) - ((sin r) * (sin r)))) .= (M . p) . k by A8, A12, A26, A29, A33, Th22 ; ::_thesis: verum end; supposeA34: j = k ; ::_thesis: (p +* (i,(- (p . i)))) . k = (M . p) . k hence (p +* (i,(- (p . i)))) . k = (p . j) * 1 by A26, FUNCT_7:32 .= (p . j) * (((sin r) * (sin r)) + ((cos r) * (cos r))) by SIN_COS:29 .= (((p . j) * (((cos r) * (cos r)) - ((sin r) * (sin r)))) + (((p . j) * (sin r)) * (sin r))) + (((p . j) * (sin r)) * (sin r)) .= (((p . j) * (((cos r) * (cos r)) - ((sin r) * (sin r)))) + ((- ((p . i) * (cos r))) * (sin r))) + ((- ((p . i) * (cos r))) * (sin r)) by A28 .= ((p . j) * (((cos r) * (cos r)) - ((sin r) * (sin r)))) + ((p . i) * (- (((sin r) * (cos r)) + ((sin r) * (cos r))))) .= (M . p) . k by A8, A12, A26, A29, A34, Th21 ; ::_thesis: verum end; supposeA35: ( i <> k & j <> k ) ; ::_thesis: (p +* (i,(- (p . i)))) . k = (M . p) . k A36: dom M = the carrier of (TOP-REAL n) by FUNCT_2:52; ( (p +* (i,(- (p . i)))) . k = p . k & not k in {i,j} ) by A35, FUNCT_7:32, TARSKI:def_2; hence (p +* (i,(- (p . i)))) . k = (M . p) . k by A30, A36, Def1; ::_thesis: verum end; end; end; take M ; ::_thesis: ( M is base_rotation & M . p = p +* (i,(- (p . i))) ) len (M . p) = n by CARD_1:def_7; hence ( M is base_rotation & M . p = p +* (i,(- (p . i))) ) by A9, A10, A31, FINSEQ_1:14; ::_thesis: verum end; end; end; Lm9: for X being set for n, i being Nat for f being additive homogeneous rotation Function of (TOP-REAL n),(TOP-REAL n) st f is X -support-yielding & not i in X holds f . (Base_FinSeq (n,i)) = Base_FinSeq (n,i) proof let X be set ; ::_thesis: for n, i being Nat for f being additive homogeneous rotation Function of (TOP-REAL n),(TOP-REAL n) st f is X -support-yielding & not i in X holds f . (Base_FinSeq (n,i)) = Base_FinSeq (n,i) let n, i be Nat; ::_thesis: for f being additive homogeneous rotation Function of (TOP-REAL n),(TOP-REAL n) st f is X -support-yielding & not i in X holds f . (Base_FinSeq (n,i)) = Base_FinSeq (n,i) let f be additive homogeneous rotation Function of (TOP-REAL n),(TOP-REAL n); ::_thesis: ( f is X -support-yielding & not i in X implies f . (Base_FinSeq (n,i)) = Base_FinSeq (n,i) ) set B = Base_FinSeq (n,i); assume that A1: f is X -support-yielding and A2: not i in X ; ::_thesis: f . (Base_FinSeq (n,i)) = Base_FinSeq (n,i) A3: now__::_thesis:_for_j_being_Nat_st_j_in_X_/\_(Seg_n)_holds_ (Base_FinSeq_(n,i))_._j_=_0 let j be Nat; ::_thesis: ( j in X /\ (Seg n) implies (Base_FinSeq (n,i)) . j = 0 ) assume A4: j in X /\ (Seg n) ; ::_thesis: (Base_FinSeq (n,i)) . j = 0 then j in Seg n by XBOOLE_0:def_4; then A5: ( 1 <= j & j <= n ) by FINSEQ_1:1; j <> i by A2, A4, XBOOLE_0:def_4; hence (Base_FinSeq (n,i)) . j = 0 by A5, MATRIXR2:76; ::_thesis: verum end; len (Base_FinSeq (n,i)) = n by MATRIXR2:74; then Base_FinSeq (n,i) is Point of (TOP-REAL n) by TOPREAL3:46; hence f . (Base_FinSeq (n,i)) = Base_FinSeq (n,i) by A1, A3, Th33; ::_thesis: verum end; theorem Th35: :: MATRTOP3:35 for i, n being Nat for f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st f is {i} -support-yielding & f is rotation & not AutMt f = AxialSymmetry (i,n) holds AutMt f = 1. (F_Real,n) proof let i, n be Nat; ::_thesis: for f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st f is {i} -support-yielding & f is rotation & not AutMt f = AxialSymmetry (i,n) holds AutMt f = 1. (F_Real,n) let f be additive homogeneous Function of (TOP-REAL n),(TOP-REAL n); ::_thesis: ( f is {i} -support-yielding & f is rotation & not AutMt f = AxialSymmetry (i,n) implies AutMt f = 1. (F_Real,n) ) set TR = TOP-REAL n; set S = { (Base_FinSeq (n,j)) where j is Element of NAT : ( 1 <= j & j <= n ) } ; { (Base_FinSeq (n,j)) where j is Element of NAT : ( 1 <= j & j <= n ) } c= the carrier of (TOP-REAL n) proof let x be set ; :: according to TARSKI:def_3 ::_thesis: ( not x in { (Base_FinSeq (n,j)) where j is Element of NAT : ( 1 <= j & j <= n ) } or x in the carrier of (TOP-REAL n) ) assume x in { (Base_FinSeq (n,j)) where j is Element of NAT : ( 1 <= j & j <= n ) } ; ::_thesis: x in the carrier of (TOP-REAL n) then consider j being Element of NAT such that A1: x = Base_FinSeq (n,j) and ( 1 <= j & j <= n ) ; len (Base_FinSeq (n,j)) = n by MATRIXR2:74; hence x in the carrier of (TOP-REAL n) by A1, TOPREAL3:46; ::_thesis: verum end; then reconsider S = { (Base_FinSeq (n,j)) where j is Element of NAT : ( 1 <= j & j <= n ) } as Subset of (TOP-REAL n) ; set M = Mx2Tran (AxialSymmetry (n,i)); assume A2: ( f is {i} -support-yielding & f is rotation ) ; ::_thesis: ( AutMt f = AxialSymmetry (i,n) or AutMt f = 1. (F_Real,n) ) A3: id (TOP-REAL n) = Mx2Tran (1. (F_Real,n)) by MATRTOP1:33; then A4: AutMt (id (TOP-REAL n)) = 1. (F_Real,n) by Def6; A5: dom f = the carrier of (TOP-REAL n) by FUNCT_2:52; percases ( not i in Seg n or i in Seg n ) ; supposeA6: not i in Seg n ; ::_thesis: ( AutMt f = AxialSymmetry (i,n) or AutMt f = 1. (F_Real,n) ) now__::_thesis:_for_p_being_Point_of_(TOP-REAL_n)_holds_f_._p_=_p let p be Point of (TOP-REAL n); ::_thesis: f . p = p A7: now__::_thesis:_for_j_being_Nat_st_1_<=_j_&_j_<=_n_holds_ (f_._p)_._j_=_p_._j let j be Nat; ::_thesis: ( 1 <= j & j <= n implies (f . p) . j = p . j ) assume ( 1 <= j & j <= n ) ; ::_thesis: (f . p) . j = p . j then j <> i by A6, FINSEQ_1:1; then not j in {i} by TARSKI:def_1; hence (f . p) . j = p . j by A2, A5, Def1; ::_thesis: verum end; ( len (f . p) = n & len p = n ) by CARD_1:def_7; hence f . p = p by A7, FINSEQ_1:14; ::_thesis: verum end; then f = id (TOP-REAL n) by FUNCT_2:124; hence ( AutMt f = AxialSymmetry (i,n) or AutMt f = 1. (F_Real,n) ) by A3, Def6; ::_thesis: verum end; supposeA8: i in Seg n ; ::_thesis: ( AutMt f = AxialSymmetry (i,n) or AutMt f = 1. (F_Real,n) ) then ( 1 <= i & i <= n ) by FINSEQ_1:1; then Base_FinSeq (n,i) in S by A8; then reconsider B = Base_FinSeq (n,i) as Point of (TOP-REAL n) ; B = (0* n) +* (i,1) by MATRIXR2:def_4; then A9: |.B.| = abs 1 by A8, TOPREALC:13 .= 1 by ABSVALUE:def_1 ; set B0 = (0* n) +* (i,((f . B) . i)); A10: len (0* n) = n by CARD_1:def_7; A11: for j being Nat st 1 <= j & j <= n holds (f . B) . j = ((0* n) +* (i,((f . B) . i))) . j proof let j be Nat; ::_thesis: ( 1 <= j & j <= n implies (f . B) . j = ((0* n) +* (i,((f . B) . i))) . j ) assume A12: ( 1 <= j & j <= n ) ; ::_thesis: (f . B) . j = ((0* n) +* (i,((f . B) . i))) . j A13: j in dom (0* n) by A10, A12, FINSEQ_3:25; percases ( j = i or j <> i ) ; suppose j = i ; ::_thesis: (f . B) . j = ((0* n) +* (i,((f . B) . i))) . j hence ((0* n) +* (i,((f . B) . i))) . j = (f . B) . j by A13, FUNCT_7:31; ::_thesis: verum end; supposeA14: j <> i ; ::_thesis: (f . B) . j = ((0* n) +* (i,((f . B) . i))) . j then A15: not j in {i} by TARSKI:def_1; thus ((0* n) +* (i,((f . B) . i))) . j = (0* n) . j by A14, FUNCT_7:32 .= 0 .= B . j by A12, A14, MATRIXR2:76 .= (f . B) . j by A2, A5, A15, Def1 ; ::_thesis: verum end; end; end; ( len ((0* n) +* (i,((f . B) . i))) = len (0* n) & len (f . B) = n ) by CARD_1:def_7, FUNCT_7:97; then A16: f . B = (0* n) +* (i,((f . B) . i)) by A10, A11, FINSEQ_1:14; then A17: |.B.| = |.((0* n) +* (i,((f . B) . i))).| by A2, Def4 .= abs ((f . B) . i) by A8, TOPREALC:13 ; A18: for h being additive homogeneous rotation Function of (TOP-REAL n),(TOP-REAL n) st h | S = id S holds h = id (TOP-REAL n) proof let h be additive homogeneous rotation Function of (TOP-REAL n),(TOP-REAL n); ::_thesis: ( h | S = id S implies h = id (TOP-REAL n) ) A19: dom (id (TOP-REAL n)) = the carrier of (TOP-REAL n) by FUNCT_2:def_1; assume A20: h | S = id S ; ::_thesis: h = id (TOP-REAL n) A21: for x being set st x in dom (id (TOP-REAL n)) holds (id (TOP-REAL n)) . x = h . x proof let x be set ; ::_thesis: ( x in dom (id (TOP-REAL n)) implies (id (TOP-REAL n)) . x = h . x ) assume A22: x in dom (id (TOP-REAL n)) ; ::_thesis: (id (TOP-REAL n)) . x = h . x then reconsider p = x as Point of (TOP-REAL n) by FUNCT_2:def_1; set hp = h . p; A23: ( len p = n & len (h . p) = n ) by CARD_1:def_7; A24: now__::_thesis:_for_j_being_Nat_st_1_<=_j_&_j_<=_n_holds_ p_._j_=_(h_._p)_._j let j be Nat; ::_thesis: ( 1 <= j & j <= n implies p . j = (h . p) . j ) assume A25: ( 1 <= j & j <= n ) ; ::_thesis: p . j = (h . p) . j then A26: j in Seg n by FINSEQ_1:1; then Base_FinSeq (n,j) in S by A25; then Base_FinSeq (n,j) in Lin S by RLVECT_3:15; hence p . j = (h . p) . j by A20, A26, Th32; ::_thesis: verum end; (id (TOP-REAL n)) . x = x by A19, A22, FUNCT_1:17; hence (id (TOP-REAL n)) . x = h . x by A23, A24, FINSEQ_1:14; ::_thesis: verum end; dom h = the carrier of (TOP-REAL n) by FUNCT_2:def_1; hence h = id (TOP-REAL n) by A19, A21, FUNCT_1:2; ::_thesis: verum end; percases ( (f . B) . i >= 0 or (f . B) . i < 0 ) ; supposeA27: (f . B) . i >= 0 ; ::_thesis: ( AutMt f = AxialSymmetry (i,n) or AutMt f = 1. (F_Real,n) ) A28: dom (f | S) = S by A5, RELAT_1:62; A29: (f . B) . i = 1 by A9, A17, A27, ABSVALUE:def_1; A30: for x being set st x in S holds (f | S) . x = (id S) . x proof let x be set ; ::_thesis: ( x in S implies (f | S) . x = (id S) . x ) assume A31: x in S ; ::_thesis: (f | S) . x = (id S) . x then consider j being Element of NAT such that A32: x = Base_FinSeq (n,j) and 1 <= j and j <= n ; A33: ( (f | S) . x = f . x & (id S) . x = x ) by A28, A31, FUNCT_1:17, FUNCT_1:47; percases ( j = i or j <> i ) ; suppose j = i ; ::_thesis: (f | S) . x = (id S) . x hence (f | S) . x = (id S) . x by A16, A29, A32, A33, MATRIXR2:def_4; ::_thesis: verum end; suppose j <> i ; ::_thesis: (f | S) . x = (id S) . x then not j in {i} by TARSKI:def_1; hence (f | S) . x = (id S) . x by A2, A32, A33, Lm9; ::_thesis: verum end; end; end; dom (id S) = S ; hence ( AutMt f = AxialSymmetry (i,n) or AutMt f = 1. (F_Real,n) ) by A2, A4, A18, A28, A30, FUNCT_1:2; ::_thesis: verum end; supposeA34: (f . B) . i < 0 ; ::_thesis: ( AutMt f = AxialSymmetry (i,n) or AutMt f = 1. (F_Real,n) ) set MA = Mx2Tran (AxialSymmetry (i,n)); Mx2Tran (AxialSymmetry (i,n)) is rotation by A8, Th27; then reconsider MAf = (Mx2Tran (AxialSymmetry (i,n))) * f as additive homogeneous rotation Function of (TOP-REAL n),(TOP-REAL n) by A2; A35: dom MAf = the carrier of (TOP-REAL n) by FUNCT_2:52; then A36: dom (MAf | S) = S by RELAT_1:62; A37: ( Mx2Tran (AxialSymmetry (i,n)) is {i} -support-yielding & {i} \/ {i} = {i} ) by A8, Th11; A38: for x being set st x in S holds (MAf | S) . x = (id S) . x proof let x be set ; ::_thesis: ( x in S implies (MAf | S) . x = (id S) . x ) assume A39: x in S ; ::_thesis: (MAf | S) . x = (id S) . x then consider j being Element of NAT such that A40: x = Base_FinSeq (n,j) and ( 1 <= j & j <= n ) ; A41: ( (MAf | S) . x = MAf . x & (id S) . x = x ) by A36, A39, FUNCT_1:17, FUNCT_1:47; percases ( j = i or j <> i ) ; supposeA42: j = i ; ::_thesis: (MAf | S) . x = (id S) . x A43: for k being Nat st 1 <= k & k <= n holds (MAf . B) . k = B . k proof let k be Nat; ::_thesis: ( 1 <= k & k <= n implies (MAf . B) . k = B . k ) assume A44: ( 1 <= k & k <= n ) ; ::_thesis: (MAf . B) . k = B . k then A45: k in Seg n by FINSEQ_1:1; percases ( k = i or k <> i ) ; supposeA46: k = i ; ::_thesis: (MAf . B) . k = B . k thus (MAf . B) . k = ((Mx2Tran (AxialSymmetry (i,n))) . (f . B)) . k by A35, FUNCT_1:12 .= - ((f . B) . k) by A45, A46, Th9 .= - (- 1) by A9, A17, A34, A46, ABSVALUE:def_1 .= B . k by A44, A46, MATRIXR2:75 ; ::_thesis: verum end; supposeA47: k <> i ; ::_thesis: (MAf . B) . k = B . k then A48: not k in {i} by TARSKI:def_1; thus (MAf . B) . k = ((Mx2Tran (AxialSymmetry (i,n))) . (f . B)) . k by A35, FUNCT_1:12 .= (f . B) . k by A8, A47, Th8 .= B . k by A2, A5, A48, Def1 ; ::_thesis: verum end; end; end; ( len (MAf . B) = n & len B = n ) by CARD_1:def_7; hence (MAf | S) . x = (id S) . x by A40, A41, A42, A43, FINSEQ_1:14; ::_thesis: verum end; suppose j <> i ; ::_thesis: (MAf | S) . x = (id S) . x then not j in {i} by TARSKI:def_1; hence (MAf | S) . x = (id S) . x by A2, A37, A40, A41, Lm9; ::_thesis: verum end; end; end; dom (id S) = S ; then A49: MAf = id (TOP-REAL n) by A18, A36, A38, FUNCT_1:2; A50: dom (Mx2Tran (AxialSymmetry (i,n))) = [#] (TOP-REAL n) by TOPS_2:def_5; set R = AutMt f; A51: rng (Mx2Tran (AxialSymmetry (i,n))) = [#] (TOP-REAL n) by TOPS_2:def_5; A52: Mx2Tran (AxialSymmetry (i,n)) is one-to-one by TOPS_2:def_5; A53: the carrier of (TOP-REAL n) c= rng f proof let x be set ; :: according to TARSKI:def_3 ::_thesis: ( not x in the carrier of (TOP-REAL n) or x in rng f ) assume A54: x in the carrier of (TOP-REAL n) ; ::_thesis: x in rng f then A55: (Mx2Tran (AxialSymmetry (i,n))) . x in rng (Mx2Tran (AxialSymmetry (i,n))) by A50, FUNCT_1:def_3; then A56: MAf . ((Mx2Tran (AxialSymmetry (i,n))) . x) = (Mx2Tran (AxialSymmetry (i,n))) . (f . ((Mx2Tran (AxialSymmetry (i,n))) . x)) by A35, A51, FUNCT_1:12; ( f . ((Mx2Tran (AxialSymmetry (i,n))) . x) in dom (Mx2Tran (AxialSymmetry (i,n))) & MAf . ((Mx2Tran (AxialSymmetry (i,n))) . x) = (Mx2Tran (AxialSymmetry (i,n))) . x ) by A35, A49, A51, A55, FUNCT_1:11, FUNCT_1:18; then x = f . ((Mx2Tran (AxialSymmetry (i,n))) . x) by A50, A52, A54, A56, FUNCT_1:def_4; hence x in rng f by A5, A51, A55, FUNCT_1:def_3; ::_thesis: verum end; rng f c= the carrier of (TOP-REAL n) by RELAT_1:def_19; then rng f = the carrier of (TOP-REAL n) by A53, XBOOLE_0:def_10; then A57: f = (Mx2Tran (AxialSymmetry (i,n))) " by A49, A51, A50, A52, FUNCT_1:42; ( f = Mx2Tran (AutMt f) & Det (AxialSymmetry (i,n)) <> 0. F_Real ) by A8, Def6, Th4; then AutMt f = (AxialSymmetry (i,n)) ~ by A57, MATRTOP1:43 .= AxialSymmetry (i,n) by A8, Th7 ; hence ( AutMt f = AxialSymmetry (i,n) or AutMt f = 1. (F_Real,n) ) ; ::_thesis: verum end; end; end; end; end; theorem Th36: :: MATRTOP3:36 for n being Nat for f1 being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st f1 is rotation holds ex f2 being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f2 is base_rotation & f2 * f1 is {n} -support-yielding ) proof let n be Nat; ::_thesis: for f1 being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st f1 is rotation holds ex f2 being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f2 is base_rotation & f2 * f1 is {n} -support-yielding ) let f1 be additive homogeneous Function of (TOP-REAL n),(TOP-REAL n); ::_thesis: ( f1 is rotation implies ex f2 being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f2 is base_rotation & f2 * f1 is {n} -support-yielding ) ) set TR = TOP-REAL n; assume A1: f1 is rotation ; ::_thesis: ex f2 being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f2 is base_rotation & f2 * f1 is {n} -support-yielding ) set cTR = the carrier of (TOP-REAL n); set f = f1; A2: dom f1 = the carrier of (TOP-REAL n) by FUNCT_2:52; A3: rng f1 c= the carrier of (TOP-REAL n) by RELAT_1:def_19; percases ( n = 0 or n > 0 ) ; supposeA4: n = 0 ; ::_thesis: ex f2 being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f2 is base_rotation & f2 * f1 is {n} -support-yielding ) take I = id (TOP-REAL n); ::_thesis: ( I is base_rotation & I * f1 is {n} -support-yielding ) A5: dom (id the carrier of (TOP-REAL n)) = the carrier of (TOP-REAL n) ; A6: for h being Function for x being set st h in dom I & (I . h) . x <> h . x holds x in {n} by A5, FUNCT_1:17; A7: the carrier of (TOP-REAL n) = {(0. (TOP-REAL n))} by A4, EUCLID:22, EUCLID:77; then ( rng (id the carrier of (TOP-REAL n)) = the carrier of (TOP-REAL n) & rng f1 = the carrier of (TOP-REAL n) ) by A3, ZFMISC_1:33; then f1 = id the carrier of (TOP-REAL n) by A2, A5, A7, FUNCT_1:7; then I * f1 = I by A2, RELAT_1:52; hence ( I is base_rotation & I * f1 is {n} -support-yielding ) by A6, Def1; ::_thesis: verum end; suppose n > 0 ; ::_thesis: ex f2 being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f2 is base_rotation & f2 * f1 is {n} -support-yielding ) then reconsider n1 = n - 1 as Nat by NAT_1:20; defpred S1[ Nat] means ( $1 <= n - 1 implies for S being Subset of (TOP-REAL n) st S = { (Base_FinSeq (n,i)) where i is Element of NAT : ( 1 <= i & i <= $1 ) } holds ex g being base_rotation Function of (TOP-REAL n),(TOP-REAL n) st (g * f1) | S = id S ); set S = { (Base_FinSeq (n,i)) where i is Element of NAT : ( 1 <= i & i <= n1 ) } ; { (Base_FinSeq (n,i)) where i is Element of NAT : ( 1 <= i & i <= n1 ) } c= the carrier of (TOP-REAL n) proof let x be set ; :: according to TARSKI:def_3 ::_thesis: ( not x in { (Base_FinSeq (n,i)) where i is Element of NAT : ( 1 <= i & i <= n1 ) } or x in the carrier of (TOP-REAL n) ) assume x in { (Base_FinSeq (n,i)) where i is Element of NAT : ( 1 <= i & i <= n1 ) } ; ::_thesis: x in the carrier of (TOP-REAL n) then consider j being Element of NAT such that A8: x = Base_FinSeq (n,j) and 1 <= j and j <= n1 ; len (Base_FinSeq (n,j)) = n by MATRIXR2:74; hence x in the carrier of (TOP-REAL n) by A8, TOPREAL3:46; ::_thesis: verum end; then reconsider S = { (Base_FinSeq (n,i)) where i is Element of NAT : ( 1 <= i & i <= n1 ) } as Subset of (TOP-REAL n) ; A9: for k being Nat st S1[k] holds S1[k + 1] proof let z be Nat; ::_thesis: ( S1[z] implies S1[z + 1] ) assume A10: S1[z] ; ::_thesis: S1[z + 1] set z1 = z + 1; set SS = { (Base_FinSeq (n,i)) where i is Element of NAT : ( 1 <= i & i <= z ) } ; assume A11: z + 1 <= n - 1 ; ::_thesis: for S being Subset of (TOP-REAL n) st S = { (Base_FinSeq (n,i)) where i is Element of NAT : ( 1 <= i & i <= z + 1 ) } holds ex g being base_rotation Function of (TOP-REAL n),(TOP-REAL n) st (g * f1) | S = id S then reconsider n1 = n - 1 as Element of NAT by INT_1:3; A12: n1 < n1 + 1 by NAT_1:13; then A13: z + 1 < n by A11, XXREAL_0:2; set B = Base_FinSeq (n,(z + 1)); set X = { i where i is Element of NAT : ( z + 1 <= i & i <= n ) } ; let S be Subset of (TOP-REAL n); ::_thesis: ( S = { (Base_FinSeq (n,i)) where i is Element of NAT : ( 1 <= i & i <= z + 1 ) } implies ex g being base_rotation Function of (TOP-REAL n),(TOP-REAL n) st (g * f1) | S = id S ) assume A14: S = { (Base_FinSeq (n,i)) where i is Element of NAT : ( 1 <= i & i <= z + 1 ) } ; ::_thesis: ex g being base_rotation Function of (TOP-REAL n),(TOP-REAL n) st (g * f1) | S = id S { (Base_FinSeq (n,i)) where i is Element of NAT : ( 1 <= i & i <= z ) } c= S proof let x be set ; :: according to TARSKI:def_3 ::_thesis: ( not x in { (Base_FinSeq (n,i)) where i is Element of NAT : ( 1 <= i & i <= z ) } or x in S ) assume x in { (Base_FinSeq (n,i)) where i is Element of NAT : ( 1 <= i & i <= z ) } ; ::_thesis: x in S then consider i being Element of NAT such that A15: ( x = Base_FinSeq (n,i) & 1 <= i ) and A16: i <= z ; i < z + 1 by A16, NAT_1:13; hence x in S by A14, A15; ::_thesis: verum end; then reconsider SS = { (Base_FinSeq (n,i)) where i is Element of NAT : ( 1 <= i & i <= z ) } as Subset of (TOP-REAL n) by XBOOLE_1:1; z < n1 by A11, NAT_1:13; then consider g being base_rotation Function of (TOP-REAL n),(TOP-REAL n) such that A17: (g * f1) | SS = id SS by A10; A18: n in NAT by ORDINAL1:def_12; then A19: n in { i where i is Element of NAT : ( z + 1 <= i & i <= n ) } by A13; n >= 1 by A12, NAT_1:14; then n in Seg n by A18; then A20: n in { i where i is Element of NAT : ( z + 1 <= i & i <= n ) } /\ (Seg n) by A19, XBOOLE_0:def_4; A21: card {(z + 1),n} = 2 by A11, A12, CARD_2:57; A22: 1 <= z + 1 by NAT_1:11; then A23: z + 1 in Seg n by A13; Base_FinSeq (n,(z + 1)) in S by A14, A22; then reconsider B = Base_FinSeq (n,(z + 1)) as Point of (TOP-REAL n) ; set gfB = (g * f1) . B; A24: z + 1 in { i where i is Element of NAT : ( z + 1 <= i & i <= n ) } by A13; then consider h being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) such that A25: ( h is { i where i is Element of NAT : ( z + 1 <= i & i <= n ) } -support-yielding & h is base_rotation ) and A26: ( card ( { i where i is Element of NAT : ( z + 1 <= i & i <= n ) } /\ (Seg n)) > 1 implies (h . ((g * f1) . B)) . (z + 1) >= 0 ) and A27: for i being Nat st i in { i where i is Element of NAT : ( z + 1 <= i & i <= n ) } /\ (Seg n) & i <> z + 1 holds (h . ((g * f1) . B)) . i = 0 by A23, Th30; reconsider hg = h * g as base_rotation Function of (TOP-REAL n),(TOP-REAL n) by A25; A28: dom (hg * f1) = the carrier of (TOP-REAL n) by FUNCT_2:52; A29: for x being set st x in SS holds ( ((h * g) * f1) . x = x & h . x = x ) proof let x be set ; ::_thesis: ( x in SS implies ( ((h * g) * f1) . x = x & h . x = x ) ) assume A30: x in SS ; ::_thesis: ( ((h * g) * f1) . x = x & h . x = x ) reconsider B = x as Point of (TOP-REAL n) by A30; ((g * f1) | SS) . x = (g * f1) . x by A30, FUNCT_1:49; then A31: (g * f1) . x = x by A17, A30, FUNCT_1:17; A32: ex i being Element of NAT st ( x = Base_FinSeq (n,i) & 1 <= i & i <= z ) by A30; A33: for j being Nat st j in { i where i is Element of NAT : ( z + 1 <= i & i <= n ) } /\ (Seg n) holds B . j = 0 proof let j be Nat; ::_thesis: ( j in { i where i is Element of NAT : ( z + 1 <= i & i <= n ) } /\ (Seg n) implies B . j = 0 ) assume A34: j in { i where i is Element of NAT : ( z + 1 <= i & i <= n ) } /\ (Seg n) ; ::_thesis: B . j = 0 then j in { i where i is Element of NAT : ( z + 1 <= i & i <= n ) } by XBOOLE_0:def_4; then ex I being Element of NAT st ( I = j & z + 1 <= I & I <= n ) ; then A35: z < j by NAT_1:13; j in Seg n by A34, XBOOLE_0:def_4; then ( 1 <= j & j <= n ) by FINSEQ_1:1; hence B . j = 0 by A32, A35, MATRIXR2:76; ::_thesis: verum end; A36: hg * f1 = h * (g * f1) by RELAT_1:36; then (h * (g * f1)) . x = h . ((g * f1) . x) by A28, A30, FUNCT_1:12; hence ( ((h * g) * f1) . x = x & h . x = x ) by A25, A31, A33, A36, Th33; ::_thesis: verum end; z + 1 in { i where i is Element of NAT : ( z + 1 <= i & i <= n ) } /\ (Seg n) by A23, A24, XBOOLE_0:def_4; then {(z + 1),n} c= { i where i is Element of NAT : ( z + 1 <= i & i <= n ) } /\ (Seg n) by A20, ZFMISC_1:32; then A37: 1 + 1 <= card ( { i where i is Element of NAT : ( z + 1 <= i & i <= n ) } /\ (Seg n)) by A21, NAT_1:43; A38: for x being set st x in S holds ((hg * f1) | S) . x = (id S) . x proof let x be set ; ::_thesis: ( x in S implies ((hg * f1) | S) . x = (id S) . x ) assume A39: x in S ; ::_thesis: ((hg * f1) | S) . x = (id S) . x A40: (id S) . x = x by A39, FUNCT_1:17; A41: hg * f1 = h * (g * f1) by RELAT_1:36; A42: ((hg * f1) | S) . x = (hg * f1) . x by A39, FUNCT_1:49; consider i being Element of NAT such that A43: x = Base_FinSeq (n,i) and A44: 1 <= i and A45: i <= z + 1 by A14, A39; percases ( i <= z or i = z + 1 ) by A45, NAT_1:8; suppose i <= z ; ::_thesis: ((hg * f1) | S) . x = (id S) . x then x in SS by A43, A44; hence ((hg * f1) | S) . x = (id S) . x by A29, A40, A42; ::_thesis: verum end; supposeA46: i = z + 1 ; ::_thesis: ((hg * f1) | S) . x = (id S) . x set H = h . ((g * f1) . B); A47: (h * (g * f1)) . x = h . ((g * f1) . B) by A28, A41, A43, A46, FUNCT_1:12; A48: len (h . ((g * f1) . B)) = n by CARD_1:def_7; A49: for j being Nat st j in Seg n & j < z + 1 holds (h . ((g * f1) . B)) . j = 0 proof let j be Nat; ::_thesis: ( j in Seg n & j < z + 1 implies (h . ((g * f1) . B)) . j = 0 ) assume that A50: j in Seg n and A51: j < z + 1 ; ::_thesis: (h . ((g * f1) . B)) . j = 0 A52: 1 <= j by A50, FINSEQ_1:1; j <= z by A51, NAT_1:13; then A53: Base_FinSeq (n,j) in SS by A50, A52; then reconsider Bnj = Base_FinSeq (n,j) as Point of (TOP-REAL n) ; ((h * g) * f1) . Bnj = Bnj by A29, A53; then A54: Bnj + (h . ((g * f1) . B)) = ((h * g) * f1) . (Bnj + B) by A41, A43, A46, A47, VECTSP_1:def_20; A55: len Bnj = n by CARD_1:def_7; |.(((h * g) * f1) . (Bnj + B)).| = |.(Bnj + B).| by A1, A25, Def4; then A56: |.(Bnj + B).| ^2 = ((|.Bnj.| ^2) + (2 * |((h . ((g * f1) . B)),Bnj)|)) + (|.(h . ((g * f1) . B)).| ^2) by A48, A54, A55, EUCLID_2:11; A57: Bnj = (0* n) +* (j,1) by MATRIXR2:def_4; len B = n by CARD_1:def_7; then A58: |.(Bnj + B).| ^2 = ((|.Bnj.| ^2) + (2 * |(B,Bnj)|)) + (|.B.| ^2) by A55, EUCLID_2:11; A59: j <= n by A50, FINSEQ_1:1; |.(h . ((g * f1) . B)).| = |.B.| by A1, A25, A43, A46, A47, Def4; then (B . j) * 1 = |((h . ((g * f1) . B)),Bnj)| by A56, A57, A58, TOPREALC:16 .= ((h . ((g * f1) . B)) . j) * 1 by A57, TOPREALC:16 ; hence (h . ((g * f1) . B)) . j = 0 by A51, A52, A59, MATRIXR2:76; ::_thesis: verum end; set H = h . ((g * f1) . B); set 0H = (0* n) +* ((z + 1),((h . ((g * f1) . B)) . (z + 1))); A60: len (0* n) = n by CARD_1:def_7; A61: for j being Nat st j in Seg n & j > z + 1 holds (h . ((g * f1) . B)) . j = 0 proof let j be Nat; ::_thesis: ( j in Seg n & j > z + 1 implies (h . ((g * f1) . B)) . j = 0 ) assume that A62: j in Seg n and A63: j > z + 1 ; ::_thesis: (h . ((g * f1) . B)) . j = 0 j <= n by A62, FINSEQ_1:1; then j in { i where i is Element of NAT : ( z + 1 <= i & i <= n ) } by A62, A63; then j in { i where i is Element of NAT : ( z + 1 <= i & i <= n ) } /\ (Seg n) by A62, XBOOLE_0:def_4; hence (h . ((g * f1) . B)) . j = 0 by A27, A63; ::_thesis: verum end; A64: for j being Nat st 1 <= j & j <= n holds (h . ((g * f1) . B)) . j = ((0* n) +* ((z + 1),((h . ((g * f1) . B)) . (z + 1)))) . j proof let j be Nat; ::_thesis: ( 1 <= j & j <= n implies (h . ((g * f1) . B)) . j = ((0* n) +* ((z + 1),((h . ((g * f1) . B)) . (z + 1)))) . j ) assume ( 1 <= j & j <= n ) ; ::_thesis: (h . ((g * f1) . B)) . j = ((0* n) +* ((z + 1),((h . ((g * f1) . B)) . (z + 1)))) . j then A65: j in Seg n by FINSEQ_1:1; then A66: j in dom (0* n) by A60, FINSEQ_1:def_3; percases ( j = z + 1 or j <> z + 1 ) ; suppose j = z + 1 ; ::_thesis: (h . ((g * f1) . B)) . j = ((0* n) +* ((z + 1),((h . ((g * f1) . B)) . (z + 1)))) . j hence (h . ((g * f1) . B)) . j = ((0* n) +* ((z + 1),((h . ((g * f1) . B)) . (z + 1)))) . j by A66, FUNCT_7:31; ::_thesis: verum end; supposeA67: j <> z + 1 ; ::_thesis: (h . ((g * f1) . B)) . j = ((0* n) +* ((z + 1),((h . ((g * f1) . B)) . (z + 1)))) . j then ( j > z + 1 or j < z + 1 ) by XXREAL_0:1; then A68: (h . ((g * f1) . B)) . j = 0 by A49, A61, A65; ((0* n) +* ((z + 1),((h . ((g * f1) . B)) . (z + 1)))) . j = (0* n) . j by A67, FUNCT_7:32; hence (h . ((g * f1) . B)) . j = ((0* n) +* ((z + 1),((h . ((g * f1) . B)) . (z + 1)))) . j by A68; ::_thesis: verum end; end; end; ( len (h . ((g * f1) . B)) = n & len ((0* n) +* ((z + 1),((h . ((g * f1) . B)) . (z + 1)))) = len (0* n) ) by CARD_1:def_7, FUNCT_7:97; then A69: (0* n) +* ((z + 1),((h . ((g * f1) . B)) . (z + 1))) = h . ((g * f1) . B) by A60, A64, FINSEQ_1:14; A70: |.((g * f1) . B).| = |.B.| by A1, Def4; A71: B = (0* n) +* ((z + 1),1) by MATRIXR2:def_4; then A72: |.B.| = abs 1 by A23, TOPREALC:13 .= 1 by ABSVALUE:def_1 ; ( |.(h . ((g * f1) . B)).| = |.((g * f1) . B).| & abs ((h . ((g * f1) . B)) . (z + 1)) = (h . ((g * f1) . B)) . (z + 1) ) by A25, A26, A37, Def4, ABSVALUE:def_1, NAT_1:13; then h . ((g * f1) . B) = B by A23, A69, A71, A72, A70, TOPREALC:13; hence ((hg * f1) | S) . x = (id S) . x by A39, A40, A41, A43, A46, A47, FUNCT_1:49; ::_thesis: verum end; end; end; take hg ; ::_thesis: (hg * f1) | S = id S ( dom (id S) = S & dom ((hg * f1) | S) = S ) by A28, RELAT_1:62; hence (hg * f1) | S = id S by A38, FUNCT_1:2; ::_thesis: verum end; A73: S1[ 0 ] proof assume 0 <= n - 1 ; ::_thesis: for S being Subset of (TOP-REAL n) st S = { (Base_FinSeq (n,i)) where i is Element of NAT : ( 1 <= i & i <= 0 ) } holds ex g being base_rotation Function of (TOP-REAL n),(TOP-REAL n) st (g * f1) | S = id S let S be Subset of (TOP-REAL n); ::_thesis: ( S = { (Base_FinSeq (n,i)) where i is Element of NAT : ( 1 <= i & i <= 0 ) } implies ex g being base_rotation Function of (TOP-REAL n),(TOP-REAL n) st (g * f1) | S = id S ) assume A74: S = { (Base_FinSeq (n,i)) where i is Element of NAT : ( 1 <= i & i <= 0 ) } ; ::_thesis: ex g being base_rotation Function of (TOP-REAL n),(TOP-REAL n) st (g * f1) | S = id S A75: S = {} proof assume S <> {} ; ::_thesis: contradiction then consider x being set such that A76: x in S by XBOOLE_0:def_1; ex i being Element of NAT st ( x = Base_FinSeq (n,i) & 1 <= i & i <= 0 ) by A74, A76; hence contradiction ; ::_thesis: verum end; take g = id (TOP-REAL n); ::_thesis: (g * f1) | S = id S (g * f1) | S = {} by A75; hence (g * f1) | S = id S by A75; ::_thesis: verum end; for k being Nat holds S1[k] from NAT_1:sch_2(A73, A9); then consider g being base_rotation Function of (TOP-REAL n),(TOP-REAL n) such that A77: (g * f1) | S = id S ; take g ; ::_thesis: ( g is base_rotation & g * f1 is {n} -support-yielding ) set gf = g * f1; thus g is base_rotation ; ::_thesis: g * f1 is {n} -support-yielding let p be Function; :: according to MATRTOP3:def_1 ::_thesis: for x being set st p in dom (g * f1) & ((g * f1) . p) . x <> p . x holds x in {n} let x be set ; ::_thesis: ( p in dom (g * f1) & ((g * f1) . p) . x <> p . x implies x in {n} ) assume that A78: p in dom (g * f1) and A79: ((g * f1) . p) . x <> p . x ; ::_thesis: x in {n} reconsider p = p as Point of (TOP-REAL n) by A78, FUNCT_2:52; len ((g * f1) . p) = n by CARD_1:def_7; then dom ((g * f1) . p) = Seg n by FINSEQ_1:def_3; then A80: ( not x in Seg n implies ((g * f1) . p) . x = {} ) by FUNCT_1:def_2; len p = n by CARD_1:def_7; then dom p = Seg n by FINSEQ_1:def_3; then A81: x in Seg n by A79, A80, FUNCT_1:def_2; assume A82: not x in {n} ; ::_thesis: contradiction reconsider x = x as Nat by A81; A83: x <= n by A81, FINSEQ_1:1; x <> n by A82, TARSKI:def_1; then x < n1 + 1 by A83, XXREAL_0:1; then A84: x <= n1 by NAT_1:13; ( x in NAT & 1 <= x ) by A81, FINSEQ_1:1; then Base_FinSeq (n,x) in S by A84; then Base_FinSeq (n,x) in Lin S by RLVECT_3:15; then ((g * f1) . p) . x = p . x by A1, A77, A81, Th32; hence contradiction by A79; ::_thesis: verum end; end; end; Lm10: for n being Nat for M being Matrix of n,F_Real st Mx2Tran M is base_rotation holds Det M = 1. F_Real proof let n be Nat; ::_thesis: for M being Matrix of n,F_Real st Mx2Tran M is base_rotation holds Det M = 1. F_Real let M be Matrix of n,F_Real; ::_thesis: ( Mx2Tran M is base_rotation implies Det M = 1. F_Real ) set TR = TOP-REAL n; set G = GFuncs the carrier of (TOP-REAL n); assume Mx2Tran M is base_rotation ; ::_thesis: Det M = 1. F_Real then consider F being FinSequence of (GFuncs the carrier of (TOP-REAL n)) such that A1: Mx2Tran M = Product F and A2: for k being Nat st k in dom F holds ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & F . k = Mx2Tran (Rotation (i,j,n,r)) ) by Def5; defpred S1[ Nat] means ( $1 <= len F implies ( Product (F | $1) is base_rotation Function of (TOP-REAL n),(TOP-REAL n) & ( for M being Matrix of n,F_Real st Mx2Tran M = Product (F | $1) holds Det M = 1. F_Real ) ) ); A3: for m being Nat st S1[m] holds S1[m + 1] proof let m be Nat; ::_thesis: ( S1[m] implies S1[m + 1] ) A4: ( n = 0 implies n = 0 ) ; set m1 = m + 1; assume A5: S1[m] ; ::_thesis: S1[m + 1] assume A6: m + 1 <= len F ; ::_thesis: ( Product (F | (m + 1)) is base_rotation Function of (TOP-REAL n),(TOP-REAL n) & ( for M being Matrix of n,F_Real st Mx2Tran M = Product (F | (m + 1)) holds Det M = 1. F_Real ) ) then reconsider P = Product (F | m) as base_rotation Function of (TOP-REAL n),(TOP-REAL n) by A5, NAT_1:13; set R = AutMt P; A7: ( width (AutMt P) = n & len (AutMt P) = n ) by MATRIX_1:24; 1 <= m + 1 by NAT_1:11; then A8: m + 1 in dom F by A6, FINSEQ_3:25; then consider i, j being Nat, r being real number such that A9: ( 1 <= i & i < j & j <= n ) and A10: F . (m + 1) = Mx2Tran (Rotation (i,j,n,r)) by A2; set O = Rotation (i,j,n,r); reconsider MO = Mx2Tran (Rotation (i,j,n,r)) as Element of (GFuncs the carrier of (TOP-REAL n)) by MONOID_0:73; F | (m + 1) = (F | m) ^ <*MO*> by A8, A10, FINSEQ_5:10; then A11: Product (F | (m + 1)) = (Product (F | m)) * MO by GROUP_4:6 .= (Mx2Tran (Rotation (i,j,n,r))) * P by MONOID_0:70 ; Mx2Tran (Rotation (i,j,n,r)) is base_rotation by A9, Th28; hence Product (F | (m + 1)) is base_rotation Function of (TOP-REAL n),(TOP-REAL n) by A11; ::_thesis: for M being Matrix of n,F_Real st Mx2Tran M = Product (F | (m + 1)) holds Det M = 1. F_Real A12: ( width (Rotation (i,j,n,r)) = n & len (Rotation (i,j,n,r)) = n ) by MATRIX_1:24; let M be Matrix of n,F_Real; ::_thesis: ( Mx2Tran M = Product (F | (m + 1)) implies Det M = 1. F_Real ) assume A13: Mx2Tran M = Product (F | (m + 1)) ; ::_thesis: Det M = 1. F_Real Mx2Tran M = (Mx2Tran (Rotation (i,j,n,r))) * (Mx2Tran (AutMt P)) by A11, A13, Def6 .= Mx2Tran ((AutMt P) * (Rotation (i,j,n,r))) by A4, A12, A7, MATRTOP1:32 ; then A14: M = (AutMt P) * (Rotation (i,j,n,r)) by MATRTOP1:34; P = Mx2Tran (AutMt P) by Def6; then A15: Det (AutMt P) = 1. F_Real by A5, A6, NAT_1:13; ( len (AutMt P) = n & Det (Rotation (i,j,n,r)) = 1. F_Real ) by A9, Th13, MATRIX_1:25; hence Det M = (1. F_Real) * (1. F_Real) by A14, A15, MATRIXR2:45 .= 1. F_Real ; ::_thesis: verum end; A16: F | (len F) = F by FINSEQ_1:58; A17: S1[ 0 ] proof assume 0 <= len F ; ::_thesis: ( Product (F | 0) is base_rotation Function of (TOP-REAL n),(TOP-REAL n) & ( for M being Matrix of n,F_Real st Mx2Tran M = Product (F | 0) holds Det M = 1. F_Real ) ) A18: Mx2Tran (1. (F_Real,n)) = id (TOP-REAL n) by MATRTOP1:33; F | 0 = <*> the carrier of (GFuncs the carrier of (TOP-REAL n)) ; then A19: Product (F | 0) = 1_ (GFuncs the carrier of (TOP-REAL n)) by GROUP_4:8 .= the_unity_wrt the multF of (GFuncs the carrier of (TOP-REAL n)) by GROUP_1:22 .= id (TOP-REAL n) by MONOID_0:75 ; hence Product (F | 0) is base_rotation Function of (TOP-REAL n),(TOP-REAL n) ; ::_thesis: for M being Matrix of n,F_Real st Mx2Tran M = Product (F | 0) holds Det M = 1. F_Real A20: n in NAT by ORDINAL1:def_12; ( n = 0 or n >= 1 ) by NAT_1:14; then A21: ( ( Det (1. (F_Real,n)) = 1_ F_Real & n >= 1 ) or ( Det (1. (F_Real,n)) = 1. F_Real & n = 0 ) ) by A20, MATRIXR2:41, MATRIX_7:16; let M be Matrix of n,F_Real; ::_thesis: ( Mx2Tran M = Product (F | 0) implies Det M = 1. F_Real ) thus ( Mx2Tran M = Product (F | 0) implies Det M = 1. F_Real ) by A18, A19, A21, MATRTOP1:34; ::_thesis: verum end; for m being Nat holds S1[m] from NAT_1:sch_2(A17, A3); hence Det M = 1. F_Real by A1, A16; ::_thesis: verum end; begin theorem Th37: :: MATRTOP3:37 for n being Nat for f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st f is rotation holds ( Det (AutMt f) = 1. F_Real iff f is base_rotation ) proof let n be Nat; ::_thesis: for f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st f is rotation holds ( Det (AutMt f) = 1. F_Real iff f is base_rotation ) let f be additive homogeneous Function of (TOP-REAL n),(TOP-REAL n); ::_thesis: ( f is rotation implies ( Det (AutMt f) = 1. F_Real iff f is base_rotation ) ) set TR = TOP-REAL n; set cTR = the carrier of (TOP-REAL n); set M = AutMt f; set MM = Mx2Tran (AutMt f); A1: ( len (AutMt f) = n & width (AutMt f) = n ) by MATRIX_1:24; A2: ( n = 0 implies n = 0 ) ; A3: Mx2Tran (AutMt f) = f by Def6; assume A4: f is rotation ; ::_thesis: ( Det (AutMt f) = 1. F_Real iff f is base_rotation ) then consider h being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) such that A5: h is base_rotation and A6: h * (Mx2Tran (AutMt f)) is {n} -support-yielding by A3, Th36; set R = AutMt h; A7: width (AutMt h) = n by MATRIX_1:24; A8: h = Mx2Tran (AutMt h) by Def6; A9: ( AutMt (h * (Mx2Tran (AutMt f))) = 1. (F_Real,n) or AutMt (h * (Mx2Tran (AutMt f))) = AxialSymmetry (n,n) ) by A4, A3, A5, A6, Th35; ( Det (AutMt f) = 1. F_Real implies Mx2Tran (AutMt f) is base_rotation ) proof assume A10: Det (AutMt f) = 1. F_Real ; ::_thesis: Mx2Tran (AutMt f) is base_rotation ( Det (AutMt h) = 1. F_Real & n in NAT ) by A5, A8, Lm10, ORDINAL1:def_12; then A11: Det ((AutMt f) * (AutMt h)) = (1. F_Real) * (1. F_Real) by A10, MATRIXR2:45 .= 1. F_Real ; A12: rng (Mx2Tran (AutMt f)) c= the carrier of (TOP-REAL n) by RELAT_1:def_19; A13: rng h = [#] (TOP-REAL n) by A5, TOPS_2:def_5; A14: ( dom h = [#] (TOP-REAL n) & h is one-to-one ) by A5, TOPS_2:def_5; A15: dom (h ") = [#] (TOP-REAL n) by A5, TOPS_2:def_5; A16: id (TOP-REAL n) = Mx2Tran (1. (F_Real,n)) by MATRTOP1:33; h * (Mx2Tran (AutMt f)) = id the carrier of (TOP-REAL n) proof assume A17: h * (Mx2Tran (AutMt f)) <> id the carrier of (TOP-REAL n) ; ::_thesis: contradiction n <> 0 proof A18: ( dom (h * (Mx2Tran (AutMt f))) = the carrier of (TOP-REAL n) & dom (id the carrier of (TOP-REAL n)) = the carrier of (TOP-REAL n) ) by FUNCT_2:52; assume n = 0 ; ::_thesis: contradiction then A19: the carrier of (TOP-REAL n) = {(0. (TOP-REAL n))} by EUCLID:22, EUCLID:77; rng (h * (Mx2Tran (AutMt f))) c= the carrier of (TOP-REAL n) by RELAT_1:def_19; then ( rng (id the carrier of (TOP-REAL n)) = the carrier of (TOP-REAL n) & rng (h * (Mx2Tran (AutMt f))) = the carrier of (TOP-REAL n) ) by A19, ZFMISC_1:33; hence contradiction by A17, A18, A19, FUNCT_1:7; ::_thesis: verum end; then A20: n in Seg n by FINSEQ_1:3; Mx2Tran (AutMt (h * (Mx2Tran (AutMt f)))) = h * (Mx2Tran (AutMt f)) by Def6; then Mx2Tran (AxialSymmetry (n,n)) = Mx2Tran ((AutMt f) * (AutMt h)) by A1, A2, A9, A8, A7, A16, A17, MATRTOP1:32; then AxialSymmetry (n,n) = (AutMt f) * (AutMt h) by MATRTOP1:34; hence contradiction by A11, A20, Th4; ::_thesis: verum end; then (h ") * (h * (Mx2Tran (AutMt f))) = h " by A15, RELAT_1:52; then h " = ((h ") * h) * (Mx2Tran (AutMt f)) by RELAT_1:36 .= (id the carrier of (TOP-REAL n)) * (Mx2Tran (AutMt f)) by A14, A13, TOPS_2:52 .= Mx2Tran (AutMt f) by A12, RELAT_1:53 ; hence Mx2Tran (AutMt f) is base_rotation by A5; ::_thesis: verum end; hence ( Det (AutMt f) = 1. F_Real iff f is base_rotation ) by A3, Lm10; ::_thesis: verum end; theorem Th38: :: MATRTOP3:38 for n being Nat for f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) holds ( not f is rotation or Det (AutMt f) = 1. F_Real or Det (AutMt f) = - (1. F_Real) ) proof let n be Nat; ::_thesis: for f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) holds ( not f is rotation or Det (AutMt f) = 1. F_Real or Det (AutMt f) = - (1. F_Real) ) let f be additive homogeneous Function of (TOP-REAL n),(TOP-REAL n); ::_thesis: ( not f is rotation or Det (AutMt f) = 1. F_Real or Det (AutMt f) = - (1. F_Real) ) set M = AutMt f; set MM = Mx2Tran (AutMt f); set TR = TOP-REAL n; A1: ( len (AutMt f) = n & width (AutMt f) = n ) by MATRIX_1:24; A2: ( n = 0 implies n = 0 ) ; A3: n in NAT by ORDINAL1:def_12; ( n = 0 or n >= 1 ) by NAT_1:14; then A4: ( ( Det (1. (F_Real,n)) = 1_ F_Real & n >= 1 ) or ( Det (1. (F_Real,n)) = 1. F_Real & n = 0 ) ) by A3, MATRIXR2:41, MATRIX_7:16; assume f is rotation ; ::_thesis: ( Det (AutMt f) = 1. F_Real or Det (AutMt f) = - (1. F_Real) ) then A5: Mx2Tran (AutMt f) is rotation by Def6; then consider h being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) such that A6: h is base_rotation and A7: h * (Mx2Tran (AutMt f)) is {n} -support-yielding by Th36; set R = AutMt h; A8: h = Mx2Tran (AutMt h) by Def6; then ( n in NAT & Det (AutMt h) = 1. F_Real ) by A6, Lm10, ORDINAL1:def_12; then A9: Det ((AutMt f) * (AutMt h)) = (Det (AutMt f)) * (1. F_Real) by MATRIXR2:45 .= Det (AutMt f) ; width (AutMt h) = n by MATRIX_1:24; then A10: h * (Mx2Tran (AutMt f)) = Mx2Tran ((AutMt f) * (AutMt h)) by A1, A2, A8, MATRTOP1:32; percases ( AutMt (h * (Mx2Tran (AutMt f))) = 1. (F_Real,n) or ( AutMt (h * (Mx2Tran (AutMt f))) <> 1. (F_Real,n) & AutMt (h * (Mx2Tran (AutMt f))) = AxialSymmetry (n,n) ) ) by A5, A6, A7, Th35; suppose AutMt (h * (Mx2Tran (AutMt f))) = 1. (F_Real,n) ; ::_thesis: ( Det (AutMt f) = 1. F_Real or Det (AutMt f) = - (1. F_Real) ) hence ( Det (AutMt f) = 1. F_Real or Det (AutMt f) = - (1. F_Real) ) by A4, A9, A10, Def6; ::_thesis: verum end; supposeA11: ( AutMt (h * (Mx2Tran (AutMt f))) <> 1. (F_Real,n) & AutMt (h * (Mx2Tran (AutMt f))) = AxialSymmetry (n,n) ) ; ::_thesis: ( Det (AutMt f) = 1. F_Real or Det (AutMt f) = - (1. F_Real) ) set cTR = the carrier of (TOP-REAL n); n <> 0 proof A12: ( dom (h * (Mx2Tran (AutMt f))) = the carrier of (TOP-REAL n) & dom (id the carrier of (TOP-REAL n)) = the carrier of (TOP-REAL n) ) by FUNCT_2:52; assume n = 0 ; ::_thesis: contradiction then A13: the carrier of (TOP-REAL n) = {(0. (TOP-REAL n))} by EUCLID:22, EUCLID:77; rng (h * (Mx2Tran (AutMt f))) c= the carrier of (TOP-REAL n) by RELAT_1:def_19; then ( rng (id the carrier of (TOP-REAL n)) = the carrier of (TOP-REAL n) & rng (h * (Mx2Tran (AutMt f))) = the carrier of (TOP-REAL n) ) by A13, ZFMISC_1:33; then h * (Mx2Tran (AutMt f)) = id (TOP-REAL n) by A12, A13, FUNCT_1:7 .= Mx2Tran (1. (F_Real,n)) by MATRTOP1:33 ; hence contradiction by A11, Def6; ::_thesis: verum end; then n in Seg n by FINSEQ_1:3; then Det (AxialSymmetry (n,n)) = - (1. F_Real) by Th4; hence ( Det (AutMt f) = 1. F_Real or Det (AutMt f) = - (1. F_Real) ) by A9, A10, A11, Def6; ::_thesis: verum end; end; end; theorem Th39: :: MATRTOP3:39 for i, n being Nat for f1, f2 being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st f1 is rotation & Det (AutMt f1) = - (1. F_Real) & i in Seg n & AutMt f2 = AxialSymmetry (i,n) holds f1 * f2 is base_rotation proof let i, n be Nat; ::_thesis: for f1, f2 being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st f1 is rotation & Det (AutMt f1) = - (1. F_Real) & i in Seg n & AutMt f2 = AxialSymmetry (i,n) holds f1 * f2 is base_rotation let f1, f2 be additive homogeneous Function of (TOP-REAL n),(TOP-REAL n); ::_thesis: ( f1 is rotation & Det (AutMt f1) = - (1. F_Real) & i in Seg n & AutMt f2 = AxialSymmetry (i,n) implies f1 * f2 is base_rotation ) set M = AutMt f1; set A = AutMt f2; assume that A1: f1 is rotation and A2: Det (AutMt f1) = - (1. F_Real) and A3: i in Seg n and A4: AutMt f2 = AxialSymmetry (i,n) ; ::_thesis: f1 * f2 is base_rotation A5: f2 = Mx2Tran (AxialSymmetry (i,n)) by A4, Def6; reconsider MF = (Mx2Tran (AutMt f1)) * f2 as additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) ; set A = AxialSymmetry (i,n); set R = AutMt MF; A6: ( n = 0 implies n = 0 ) ; A7: ( MF = Mx2Tran (AutMt MF) & width (AutMt f1) = n ) by Def6, MATRIX_1:24; ( len (AxialSymmetry (i,n)) = n & width (AxialSymmetry (i,n)) = n ) by MATRIX_1:24; then Mx2Tran (AutMt MF) = Mx2Tran ((AxialSymmetry (i,n)) * (AutMt f1)) by A5, A6, A7, MATRTOP1:32; then A8: AutMt MF = (AxialSymmetry (i,n)) * (AutMt f1) by MATRTOP1:34; ( n in NAT & Det (AxialSymmetry (i,n)) = - (1. F_Real) ) by A3, Th4, ORDINAL1:def_12; then A9: Det (AutMt MF) = (- (1. F_Real)) * (- (1. F_Real)) by A2, A8, MATRIXR2:45 .= 1. F_Real ; A10: Mx2Tran (AutMt f1) = f1 by Def6; f2 is rotation by A3, A5, Th27; hence f1 * f2 is base_rotation by A10, A1, A9, Th37; ::_thesis: verum end; Lm11: for n being Nat for f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st f is base_rotation holds AutMt f is Orthogonal proof let n be Nat; ::_thesis: for f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st f is base_rotation holds AutMt f is Orthogonal let f be additive homogeneous Function of (TOP-REAL n),(TOP-REAL n); ::_thesis: ( f is base_rotation implies AutMt f is Orthogonal ) set TR = TOP-REAL n; set G = GFuncs the carrier of (TOP-REAL n); assume f is base_rotation ; ::_thesis: AutMt f is Orthogonal then consider F being FinSequence of (GFuncs the carrier of (TOP-REAL n)) such that A1: f = Product F and A2: for k being Nat st k in dom F holds ex i, j being Nat ex r being real number st ( 1 <= i & i < j & j <= n & F . k = Mx2Tran (Rotation (i,j,n,r)) ) by Def5; A3: f = Mx2Tran (AutMt f) by Def6; defpred S1[ Nat] means ( $1 <= len F implies ( Product (F | $1) is base_rotation Function of (TOP-REAL n),(TOP-REAL n) & ( for M being Matrix of n,F_Real st Mx2Tran M = Product (F | $1) holds M is Orthogonal ) ) ); A4: for m being Nat st S1[m] holds S1[m + 1] proof let m be Nat; ::_thesis: ( S1[m] implies S1[m + 1] ) A5: ( n = 0 implies n = 0 ) ; set m1 = m + 1; assume A6: S1[m] ; ::_thesis: S1[m + 1] assume A7: m + 1 <= len F ; ::_thesis: ( Product (F | (m + 1)) is base_rotation Function of (TOP-REAL n),(TOP-REAL n) & ( for M being Matrix of n,F_Real st Mx2Tran M = Product (F | (m + 1)) holds M is Orthogonal ) ) then reconsider P = Product (F | m) as base_rotation Function of (TOP-REAL n),(TOP-REAL n) by A6, NAT_1:13; set R = AutMt P; A8: ( width (AutMt P) = n & len (AutMt P) = n ) by MATRIX_1:24; 1 <= m + 1 by NAT_1:11; then A9: m + 1 in dom F by A7, FINSEQ_3:25; then consider i, j being Nat, r being real number such that A10: ( 1 <= i & i < j & j <= n ) and A11: F . (m + 1) = Mx2Tran (Rotation (i,j,n,r)) by A2; set O = Rotation (i,j,n,r); reconsider MO = Mx2Tran (Rotation (i,j,n,r)) as Element of (GFuncs the carrier of (TOP-REAL n)) by MONOID_0:73; F | (m + 1) = (F | m) ^ <*MO*> by A9, A11, FINSEQ_5:10; then A12: Product (F | (m + 1)) = (Product (F | m)) * MO by GROUP_4:6 .= (Mx2Tran (Rotation (i,j,n,r))) * P by MONOID_0:70 ; Mx2Tran (Rotation (i,j,n,r)) is base_rotation by A10, Th28; hence Product (F | (m + 1)) is base_rotation Function of (TOP-REAL n),(TOP-REAL n) by A12; ::_thesis: for M being Matrix of n,F_Real st Mx2Tran M = Product (F | (m + 1)) holds M is Orthogonal A13: ( width (Rotation (i,j,n,r)) = n & len (Rotation (i,j,n,r)) = n ) by MATRIX_1:24; let M be Matrix of n,F_Real; ::_thesis: ( Mx2Tran M = Product (F | (m + 1)) implies M is Orthogonal ) assume A14: Mx2Tran M = Product (F | (m + 1)) ; ::_thesis: M is Orthogonal Mx2Tran M = (Mx2Tran (Rotation (i,j,n,r))) * (Mx2Tran (AutMt P)) by A12, A14, Def6 .= Mx2Tran ((AutMt P) * (Rotation (i,j,n,r))) by A5, A13, A8, MATRTOP1:32 ; then A15: M = (AutMt P) * (Rotation (i,j,n,r)) by MATRTOP1:34; P = Mx2Tran (AutMt P) by Def6; then A16: ( AutMt P is Orthogonal & n > i ) by A10, A6, A7, NAT_1:13, XXREAL_0:2; ( len (AutMt P) = n & Rotation (i,j,n,r) is Orthogonal & n > 0 ) by A10, Th19, MATRIX_1:25; hence M is Orthogonal by A15, A16, MATRIX_6:46; ::_thesis: verum end; A17: F | (len F) = F by FINSEQ_1:58; A18: S1[ 0 ] proof assume 0 <= len F ; ::_thesis: ( Product (F | 0) is base_rotation Function of (TOP-REAL n),(TOP-REAL n) & ( for M being Matrix of n,F_Real st Mx2Tran M = Product (F | 0) holds M is Orthogonal ) ) A19: Mx2Tran (1. (F_Real,n)) = id (TOP-REAL n) by MATRTOP1:33; F | 0 = <*> the carrier of (GFuncs the carrier of (TOP-REAL n)) ; then A20: Product (F | 0) = 1_ (GFuncs the carrier of (TOP-REAL n)) by GROUP_4:8 .= the_unity_wrt the multF of (GFuncs the carrier of (TOP-REAL n)) by GROUP_1:22 .= id (TOP-REAL n) by MONOID_0:75 ; hence Product (F | 0) is base_rotation Function of (TOP-REAL n),(TOP-REAL n) ; ::_thesis: for M being Matrix of n,F_Real st Mx2Tran M = Product (F | 0) holds M is Orthogonal A21: 1. (F_Real,n) is Orthogonal by MATRIX_6:58; let M be Matrix of n,F_Real; ::_thesis: ( Mx2Tran M = Product (F | 0) implies M is Orthogonal ) thus ( Mx2Tran M = Product (F | 0) implies M is Orthogonal ) by A19, A20, A21, MATRTOP1:34; ::_thesis: verum end; for m being Nat holds S1[m] from NAT_1:sch_2(A18, A4); hence AutMt f is Orthogonal by A1, A17, A3; ::_thesis: verum end; registration let n be Nat; let f be additive homogeneous rotation Function of (TOP-REAL n),(TOP-REAL n); cluster AutMt f -> Orthogonal ; coherence AutMt f is Orthogonal proof set TR = TOP-REAL n; set cTR = the carrier of (TOP-REAL n); set M = AutMt f; percases ( Det (AutMt f) = 1. F_Real or ( Det (AutMt f) = - (1. F_Real) & not f is base_rotation ) ) by Th37, Th38; suppose Det (AutMt f) = 1. F_Real ; ::_thesis: AutMt f is Orthogonal then f is base_rotation by Th37; hence AutMt f is Orthogonal by Lm11; ::_thesis: verum end; supposeA1: ( Det (AutMt f) = - (1. F_Real) & not f is base_rotation ) ; ::_thesis: AutMt f is Orthogonal A2: n > 0 proof A3: ( dom f = the carrier of (TOP-REAL n) & dom (id the carrier of (TOP-REAL n)) = the carrier of (TOP-REAL n) ) by FUNCT_2:52; assume n <= 0 ; ::_thesis: contradiction then n = 0 ; then A4: the carrier of (TOP-REAL n) = {(0. (TOP-REAL n))} by EUCLID:22, EUCLID:77; rng f c= the carrier of (TOP-REAL n) by RELAT_1:def_19; then ( rng (id the carrier of (TOP-REAL n)) = the carrier of (TOP-REAL n) & rng f = the carrier of (TOP-REAL n) ) by A4, ZFMISC_1:33; then f = id (TOP-REAL n) by A3, A4, FUNCT_1:7; hence contradiction by A1; ::_thesis: verum end; then A5: n in Seg n by FINSEQ_1:3; set A = AxialSymmetry (n,n); set MA = Mx2Tran (AxialSymmetry (n,n)); AutMt (Mx2Tran (AxialSymmetry (n,n))) = AxialSymmetry (n,n) by Def6; then A6: f * (Mx2Tran (AxialSymmetry (n,n))) is base_rotation by A1, A5, Th39; ( AxialSymmetry (n,n) is V244( F_Real ) & (AxialSymmetry (n,n)) ~ = AxialSymmetry (n,n) ) by A5, Th7; then (AxialSymmetry (n,n)) @ = (AxialSymmetry (n,n)) ~ by Th2; then A7: ( AxialSymmetry (n,n) is Orthogonal & AutMt (f * (Mx2Tran (AxialSymmetry (n,n)))) is Orthogonal ) by A6, Lm11, MATRIX_6:def_7; A8: AutMt (f * (Mx2Tran (AxialSymmetry (n,n)))) = (AutMt (Mx2Tran (AxialSymmetry (n,n)))) * (AutMt f) by Th29 .= (AxialSymmetry (n,n)) * (AutMt f) by Def6 ; (AxialSymmetry (n,n)) ~ is_reverse_of AxialSymmetry (n,n) by MATRIX_6:def_4; then A9: ((AxialSymmetry (n,n)) ~) * (AxialSymmetry (n,n)) = 1. (F_Real,n) by MATRIX_6:def_2; ( width ((AxialSymmetry (n,n)) ~) = n & len (AxialSymmetry (n,n)) = n & width (AxialSymmetry (n,n)) = n & len (AutMt f) = n ) by MATRIX_1:24; then ((AxialSymmetry (n,n)) ~) * ((AxialSymmetry (n,n)) * (AutMt f)) = (((AxialSymmetry (n,n)) ~) * (AxialSymmetry (n,n))) * (AutMt f) by MATRIX_3:33 .= AutMt f by A9, MATRIX_3:18 ; hence AutMt f is Orthogonal by A2, A7, A8, MATRIX_6:59; ::_thesis: verum end; end; end; end; registration let n be Nat; cluster Function-like quasi_total additive homogeneous rotation -> being_homeomorphism for Element of bool [: the carrier of (TOP-REAL n), the carrier of (TOP-REAL n):]; coherence for b1 being Function of (TOP-REAL n),(TOP-REAL n) st b1 is homogeneous & b1 is additive & b1 is rotation holds b1 is being_homeomorphism proof set TR = TOP-REAL n; set cTR = the carrier of (TOP-REAL n); let f be Function of (TOP-REAL n),(TOP-REAL n); ::_thesis: ( f is homogeneous & f is additive & f is rotation implies f is being_homeomorphism ) assume A1: ( f is homogeneous & f is additive & f is rotation ) ; ::_thesis: f is being_homeomorphism then reconsider F = f as additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) ; set M = AutMt F; percases ( Det (AutMt F) = 1. F_Real or ( Det (AutMt F) = - (1. F_Real) & not f is base_rotation ) ) by A1, Th37, Th38; suppose Det (AutMt F) = 1. F_Real ; ::_thesis: f is being_homeomorphism then f is base_rotation by A1, Th37; hence f is being_homeomorphism ; ::_thesis: verum end; supposeA2: ( Det (AutMt F) = - (1. F_Real) & not f is base_rotation ) ; ::_thesis: f is being_homeomorphism n <> 0 proof A3: ( dom f = the carrier of (TOP-REAL n) & dom (id the carrier of (TOP-REAL n)) = the carrier of (TOP-REAL n) ) by FUNCT_2:52; assume n = 0 ; ::_thesis: contradiction then A4: the carrier of (TOP-REAL n) = {(0. (TOP-REAL n))} by EUCLID:22, EUCLID:77; rng f c= the carrier of (TOP-REAL n) by RELAT_1:def_19; then ( rng (id the carrier of (TOP-REAL n)) = the carrier of (TOP-REAL n) & rng f = the carrier of (TOP-REAL n) ) by A4, ZFMISC_1:33; then f = id (TOP-REAL n) by A3, A4, FUNCT_1:7; hence contradiction by A2; ::_thesis: verum end; then A5: n in Seg n by FINSEQ_1:3; A6: dom f = the carrier of (TOP-REAL n) by FUNCT_2:52; set A = AxialSymmetry (n,n); set MA = Mx2Tran (AxialSymmetry (n,n)); A7: (Mx2Tran (AxialSymmetry (n,n))) " is being_homeomorphism by TOPS_2:56; A8: ( Mx2Tran (AxialSymmetry (n,n)) is one-to-one & rng (Mx2Tran (AxialSymmetry (n,n))) = [#] (TOP-REAL n) ) by TOPS_2:def_5; AutMt (Mx2Tran (AxialSymmetry (n,n))) = AxialSymmetry (n,n) by Def6; then A9: f * (Mx2Tran (AxialSymmetry (n,n))) is base_rotation by A1, A2, A5, Th39; (f * (Mx2Tran (AxialSymmetry (n,n)))) * ((Mx2Tran (AxialSymmetry (n,n))) ") = f * ((Mx2Tran (AxialSymmetry (n,n))) * ((Mx2Tran (AxialSymmetry (n,n))) ")) by RELAT_1:36 .= f * (id the carrier of (TOP-REAL n)) by A8, TOPS_2:52 .= f by A6, RELAT_1:51 ; hence f is being_homeomorphism by A9, A7, TOPS_2:57; ::_thesis: verum end; end; end; end; begin theorem :: MATRTOP3:40 for n being Nat for p, q being Point of (TOP-REAL n) st n = 1 & |.p.| = |.q.| holds ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is rotation & f . p = q & ( AutMt f = AxialSymmetry (n,n) or AutMt f = 1. (F_Real,n) ) ) proof let n be Nat; ::_thesis: for p, q being Point of (TOP-REAL n) st n = 1 & |.p.| = |.q.| holds ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is rotation & f . p = q & ( AutMt f = AxialSymmetry (n,n) or AutMt f = 1. (F_Real,n) ) ) let p, q be Point of (TOP-REAL n); ::_thesis: ( n = 1 & |.p.| = |.q.| implies ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is rotation & f . p = q & ( AutMt f = AxialSymmetry (n,n) or AutMt f = 1. (F_Real,n) ) ) ) set TR = TOP-REAL n; assume that A1: n = 1 and A2: |.p.| = |.q.| ; ::_thesis: ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is rotation & f . p = q & ( AutMt f = AxialSymmetry (n,n) or AutMt f = 1. (F_Real,n) ) ) percases ( p = q or p <> q ) ; supposeA3: p = q ; ::_thesis: ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is rotation & f . p = q & ( AutMt f = AxialSymmetry (n,n) or AutMt f = 1. (F_Real,n) ) ) take I = id (TOP-REAL n); ::_thesis: ( I is rotation & I . p = q & ( AutMt I = AxialSymmetry (n,n) or AutMt I = 1. (F_Real,n) ) ) id (TOP-REAL n) = Mx2Tran (1. (F_Real,1)) by A1, MATRTOP1:33; hence ( I is rotation & I . p = q & ( AutMt I = AxialSymmetry (n,n) or AutMt I = 1. (F_Real,n) ) ) by A1, A3, Def6, FUNCT_1:17; ::_thesis: verum end; supposeA4: p <> q ; ::_thesis: ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is rotation & f . p = q & ( AutMt f = AxialSymmetry (n,n) or AutMt f = 1. (F_Real,n) ) ) A5: len p = 1 by A1, CARD_1:def_7; then A6: p = <*(p . 1)*> by FINSEQ_1:40; A7: 1 in Seg 1 ; then reconsider f = Mx2Tran (AxialSymmetry (1,1)) as additive homogeneous rotation Function of (TOP-REAL n),(TOP-REAL n) by A1, Th27; take f ; ::_thesis: ( f is rotation & f . p = q & ( AutMt f = AxialSymmetry (n,n) or AutMt f = 1. (F_Real,n) ) ) A8: ( (q . 1) ^2 >= 0 & (p . 1) ^2 >= 0 ) by XREAL_1:63; A9: |.p.| = sqrt (Sum (sqr <*(p . 1)*>)) by A5, FINSEQ_1:40 .= sqrt (Sum <*((p . 1) ^2)*>) by RVSUM_1:55 .= sqrt ((p . 1) ^2) by RVSUM_1:73 ; A10: len q = 1 by A1, CARD_1:def_7; then A11: q = <*(q . 1)*> by FINSEQ_1:40; |.q.| = sqrt (Sum (sqr <*(q . 1)*>)) by A10, FINSEQ_1:40 .= sqrt (Sum <*((q . 1) ^2)*>) by RVSUM_1:55 .= sqrt ((q . 1) ^2) by RVSUM_1:73 ; then A12: (q . 1) ^2 = (p . 1) ^2 by A2, A8, A9, SQUARE_1:28; len (f . p) = 1 by A1, CARD_1:def_7; then f . p = <*((f . p) . 1)*> by FINSEQ_1:40 .= <*(- (p . 1))*> by A1, A7, Th9 .= q by A4, A6, A11, A12, SQUARE_1:40 ; hence ( f is rotation & f . p = q & ( AutMt f = AxialSymmetry (n,n) or AutMt f = 1. (F_Real,n) ) ) by A1, Def6; ::_thesis: verum end; end; end; theorem :: MATRTOP3:41 for n being Nat for p, q being Point of (TOP-REAL n) st n <> 1 & |.p.| = |.q.| holds ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is base_rotation & f . p = q ) proof let n be Nat; ::_thesis: for p, q being Point of (TOP-REAL n) st n <> 1 & |.p.| = |.q.| holds ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is base_rotation & f . p = q ) let p, q be Point of (TOP-REAL n); ::_thesis: ( n <> 1 & |.p.| = |.q.| implies ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is base_rotation & f . p = q ) ) set TR = TOP-REAL n; assume A1: n <> 1 ; ::_thesis: ( not |.p.| = |.q.| or ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is base_rotation & f . p = q ) ) assume A2: |.p.| = |.q.| ; ::_thesis: ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is base_rotation & f . p = q ) percases ( p = q or p <> q ) ; supposeA3: p = q ; ::_thesis: ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is base_rotation & f . p = q ) take I = id (TOP-REAL n); ::_thesis: ( I is base_rotation & I . p = q ) thus ( I is base_rotation & I . p = q ) by A3, FUNCT_1:17; ::_thesis: verum end; supposeA4: p <> q ; ::_thesis: ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is base_rotation & f . p = q ) A5: n <> 0 proof assume A6: n = 0 ; ::_thesis: contradiction then p = 0. (TOP-REAL n) by EUCLID:77; hence contradiction by A4, A6, EUCLID:77; ::_thesis: verum end; then A7: n >= 1 by NAT_1:14; defpred S1[ Nat] means ( $1 + 1 <= n implies ex f being base_rotation Function of (TOP-REAL n),(TOP-REAL n) ex X being set st ( card X = $1 & X c= Seg n & ( for k being Nat st k in X holds (f . p) . k = q . k ) ) ); A8: Sum (sqr q) >= 0 by RVSUM_1:86; A9: card (Seg n) = n by FINSEQ_1:57; A10: for m being Nat st S1[m] holds S1[m + 1] proof let m be Nat; ::_thesis: ( S1[m] implies S1[m + 1] ) assume A11: S1[m] ; ::_thesis: S1[m + 1] set m1 = m + 1; assume A12: (m + 1) + 1 <= n ; ::_thesis: ex f being base_rotation Function of (TOP-REAL n),(TOP-REAL n) ex X being set st ( card X = m + 1 & X c= Seg n & ( for k being Nat st k in X holds (f . p) . k = q . k ) ) then consider f being base_rotation Function of (TOP-REAL n),(TOP-REAL n), Xm being set such that A13: card Xm = m and A14: Xm c= Seg n and A15: for k being Nat st k in Xm holds (f . p) . k = q . k by A11, NAT_1:13; set fp = f . p; set sfp = sqr (f . p); set sq = sqr q; A16: m + 1 < n by A12, NAT_1:13; percases ( ex k being Nat st ( k in (Seg n) \ Xm & (sqr (f . p)) . k >= (sqr q) . k ) or for k being Nat st k in (Seg n) \ Xm holds (sqr (f . p)) . k < (sqr q) . k ) ; suppose ex k being Nat st ( k in (Seg n) \ Xm & (sqr (f . p)) . k >= (sqr q) . k ) ; ::_thesis: ex f being base_rotation Function of (TOP-REAL n),(TOP-REAL n) ex X being set st ( card X = m + 1 & X c= Seg n & ( for k being Nat st k in X holds (f . p) . k = q . k ) ) then consider k being Nat such that A17: k in (Seg n) \ Xm and A18: (sqr (f . p)) . k >= (sqr q) . k ; A19: k in Seg n by A17, XBOOLE_0:def_5; then A20: 1 <= k by FINSEQ_1:1; set Xmk = Xm \/ {k}; A21: ( (sqr (f . p)) . k = ((f . p) . k) ^2 & (sqr q) . k = (q . k) ^2 ) by VALUED_1:11; A22: {k} c= Seg n by A19, ZFMISC_1:31; then A23: Xm \/ {k} c= Seg n by A14, XBOOLE_1:8; A24: not k in Xm by A17, XBOOLE_0:def_5; then card (Xm \/ {k}) = m + 1 by A13, A14, CARD_2:41; then Xm \/ {k} c< Seg n by A9, A16, A23, XBOOLE_0:def_8; then consider z being set such that A25: z in Seg n and A26: not z in Xm \/ {k} by XBOOLE_0:6; reconsider z = z as Nat by A25; A27: 1 <= z by A25, FINSEQ_1:1; ((f . p) . z) ^2 >= 0 by XREAL_1:63; then A28: 0 + ((q . k) ^2) <= (((f . p) . z) ^2) + (((f . p) . k) ^2) by A18, A21, XREAL_1:7; A29: z <= n by A25, FINSEQ_1:1; A30: k <= n by A19, FINSEQ_1:1; not z in {k} by A26, XBOOLE_0:def_3; then A31: z <> k by TARSKI:def_1; now__::_thesis:_ex_g_being_base_rotation_Function_of_(TOP-REAL_n),(TOP-REAL_n)_st_ (_g_is_{k,z}_-support-yielding_&_(g_._(f_._p))_._k_=_q_._k_) percases ( z < k or z > k ) by A31, XXREAL_0:1; supposeA32: z < k ; ::_thesis: ex g being base_rotation Function of (TOP-REAL n),(TOP-REAL n) st ( g is {k,z} -support-yielding & (g . (f . p)) . k = q . k ) then consider r being real number such that A33: ((Mx2Tran (Rotation (z,k,n,r))) . (f . p)) . k = q . k by A27, A28, A30, Th25; ( Mx2Tran (Rotation (z,k,n,r)) is {k,z} -support-yielding & Mx2Tran (Rotation (z,k,n,r)) is base_rotation ) by A27, A30, A32, Th28, Th26; hence ex g being base_rotation Function of (TOP-REAL n),(TOP-REAL n) st ( g is {k,z} -support-yielding & (g . (f . p)) . k = q . k ) by A33; ::_thesis: verum end; supposeA34: z > k ; ::_thesis: ex g being base_rotation Function of (TOP-REAL n),(TOP-REAL n) st ( g is {k,z} -support-yielding & (g . (f . p)) . k = q . k ) then consider r being real number such that A35: ((Mx2Tran (Rotation (k,z,n,r))) . (f . p)) . k = q . k by A20, A28, A29, Th24; ( Mx2Tran (Rotation (k,z,n,r)) is {z,k} -support-yielding & Mx2Tran (Rotation (k,z,n,r)) is base_rotation ) by A20, A29, A34, Th28, Th26; hence ex g being base_rotation Function of (TOP-REAL n),(TOP-REAL n) st ( g is {k,z} -support-yielding & (g . (f . p)) . k = q . k ) by A35; ::_thesis: verum end; end; end; then consider g being base_rotation Function of (TOP-REAL n),(TOP-REAL n) such that A36: g is {k,z} -support-yielding and A37: (g . (f . p)) . k = q . k ; take gf = g * f; ::_thesis: ex X being set st ( card X = m + 1 & X c= Seg n & ( for k being Nat st k in X holds (gf . p) . k = q . k ) ) take Xm \/ {k} ; ::_thesis: ( card (Xm \/ {k}) = m + 1 & Xm \/ {k} c= Seg n & ( for k being Nat st k in Xm \/ {k} holds (gf . p) . k = q . k ) ) thus ( card (Xm \/ {k}) = m + 1 & Xm \/ {k} c= Seg n ) by A13, A14, A22, A24, CARD_2:41, XBOOLE_1:8; ::_thesis: for k being Nat st k in Xm \/ {k} holds (gf . p) . k = q . k let m be Nat; ::_thesis: ( m in Xm \/ {k} implies (gf . p) . m = q . m ) A38: dom gf = the carrier of (TOP-REAL n) by FUNCT_2:52; A39: dom g = the carrier of (TOP-REAL n) by FUNCT_2:52; assume A40: m in Xm \/ {k} ; ::_thesis: (gf . p) . m = q . m then A41: ( m in Xm or m in {k} ) by XBOOLE_0:def_3; percases ( m in Xm or m = k ) by A41, TARSKI:def_1; supposeA42: m in Xm ; ::_thesis: (gf . p) . m = q . m then m <> k by A17, XBOOLE_0:def_5; then not m in {k,z} by A26, A40, TARSKI:def_2; then (g . (f . p)) . m = (f . p) . m by A36, A39, Def1; hence (gf . p) . m = (f . p) . m by A38, FUNCT_1:12 .= q . m by A15, A42 ; ::_thesis: verum end; suppose m = k ; ::_thesis: (gf . p) . m = q . m hence (gf . p) . m = q . m by A37, A38, FUNCT_1:12; ::_thesis: verum end; end; end; supposeA43: for k being Nat st k in (Seg n) \ Xm holds (sqr (f . p)) . k < (sqr q) . k ; ::_thesis: ex f being base_rotation Function of (TOP-REAL n),(TOP-REAL n) ex X being set st ( card X = m + 1 & X c= Seg n & ( for k being Nat st k in X holds (f . p) . k = q . k ) ) A44: Sum (sqr (f . p)) >= 0 by RVSUM_1:86; ( Sum (sqr q) >= 0 & |.p.| = |.(f . p).| ) by Def4, RVSUM_1:86; then A45: Sum (sqr (f . p)) = Sum (sqr q) by A2, A44, SQUARE_1:28; f . p = @ (@ (f . p)) ; then A46: len (sqr (f . p)) = len (f . p) by RVSUM_1:143; q = @ (@ q) ; then A47: len (sqr q) = len q by RVSUM_1:143; ( len (f . p) = n & len q = n ) by CARD_1:def_7; then reconsider sfp = sqr (f . p), sq = sqr q as Element of n -tuples_on REAL by A46, A47, FINSEQ_2:92; m < n by A16, NAT_1:13; then Xm <> Seg n by A13, FINSEQ_1:57; then Xm c< Seg n by A14, XBOOLE_0:def_8; then consider z being set such that A48: z in Seg n and A49: not z in Xm by XBOOLE_0:6; reconsider z = z as Nat by A48; A50: z in (Seg n) \ Xm by A48, A49, XBOOLE_0:def_5; for k being Nat st k in Seg n holds sfp . k <= sq . k proof let k be Nat; ::_thesis: ( k in Seg n implies sfp . k <= sq . k ) assume A51: k in Seg n ; ::_thesis: sfp . k <= sq . k percases ( k in (Seg n) \ Xm or k in Xm ) by A51, XBOOLE_0:def_5; suppose k in (Seg n) \ Xm ; ::_thesis: sfp . k <= sq . k hence sfp . k <= sq . k by A43; ::_thesis: verum end; suppose k in Xm ; ::_thesis: sfp . k <= sq . k then (f . p) . k = q . k by A15; then sfp . k = (q . k) ^2 by VALUED_1:11 .= sq . k by VALUED_1:11 ; hence sfp . k <= sq . k ; ::_thesis: verum end; end; end; then sfp . z >= sq . z by A45, A48, RVSUM_1:83; hence ex f being base_rotation Function of (TOP-REAL n),(TOP-REAL n) ex X being set st ( card X = m + 1 & X c= Seg n & ( for k being Nat st k in X holds (f . p) . k = q . k ) ) by A43, A50; ::_thesis: verum end; end; end; reconsider n1 = n - 1 as Nat by A5, NAT_1:14, NAT_1:21; A52: n1 + 1 = n ; A53: S1[ 0 ] proof assume 0 + 1 <= n ; ::_thesis: ex f being base_rotation Function of (TOP-REAL n),(TOP-REAL n) ex X being set st ( card X = 0 & X c= Seg n & ( for k being Nat st k in X holds (f . p) . k = q . k ) ) take f = id (TOP-REAL n); ::_thesis: ex X being set st ( card X = 0 & X c= Seg n & ( for k being Nat st k in X holds (f . p) . k = q . k ) ) take X = {} ; ::_thesis: ( card X = 0 & X c= Seg n & ( for k being Nat st k in X holds (f . p) . k = q . k ) ) thus ( card X = 0 & X c= Seg n ) by XBOOLE_1:2; ::_thesis: for k being Nat st k in X holds (f . p) . k = q . k let k be Nat; ::_thesis: ( k in X implies (f . p) . k = q . k ) assume k in X ; ::_thesis: (f . p) . k = q . k hence (f . p) . k = q . k ; ::_thesis: verum end; for m being Nat holds S1[m] from NAT_1:sch_2(A53, A10); then consider f being base_rotation Function of (TOP-REAL n),(TOP-REAL n), X being set such that A54: ( card X = n1 & X c= Seg n ) and A55: for k being Nat st k in X holds (f . p) . k = q . k by A52; card ((Seg n) \ X) = n - n1 by A9, A54, CARD_2:44; then consider z being set such that A56: {z} = (Seg n) \ X by CARD_2:42; set fp = f . p; ( Sum (sqr (f . p)) >= 0 & |.p.| = |.(f . p).| ) by Def4, RVSUM_1:86; then A57: Sum (sqr q) = Sum (sqr (f . p)) by A2, A8, SQUARE_1:28; A58: z in {z} by TARSKI:def_1; then A59: z in Seg n by A56, XBOOLE_0:def_5; reconsider z = z as Nat by A56, A58; set fpz = (f . p) +* (z,(q . z)); A60: len (f . p) = n by CARD_1:def_7; then A61: dom (f . p) = Seg n by FINSEQ_1:def_3; A62: for k being Nat st 1 <= k & k <= n holds ((f . p) +* (z,(q . z))) . k = q . k proof let k be Nat; ::_thesis: ( 1 <= k & k <= n implies ((f . p) +* (z,(q . z))) . k = q . k ) assume ( 1 <= k & k <= n ) ; ::_thesis: ((f . p) +* (z,(q . z))) . k = q . k then A63: k in Seg n by FINSEQ_1:1; percases ( k = z or k <> z ) ; suppose k = z ; ::_thesis: ((f . p) +* (z,(q . z))) . k = q . k hence ((f . p) +* (z,(q . z))) . k = q . k by A61, A63, FUNCT_7:31; ::_thesis: verum end; supposeA64: k <> z ; ::_thesis: ((f . p) +* (z,(q . z))) . k = q . k then not k in (Seg n) \ X by A56, TARSKI:def_1; then A65: k in X by A63, XBOOLE_0:def_5; ((f . p) +* (z,(q . z))) . k = (f . p) . k by A64, FUNCT_7:32; hence ((f . p) +* (z,(q . z))) . k = q . k by A55, A65; ::_thesis: verum end; end; end; A66: ( len ((f . p) +* (z,(q . z))) = len (f . p) & len q = n ) by CARD_1:def_7, FUNCT_7:97; then A67: (f . p) +* (z,(q . z)) = q by A60, A62, FINSEQ_1:14; then A68: Sum (sqr q) = ((Sum (sqr (f . p))) - (((f . p) . z) ^2)) + ((q . z) ^2) by A59, A61, Th3; percases ( (f . p) . z = q . z or (f . p) . z = - (q . z) ) by A68, A57, SQUARE_1:40; suppose (f . p) . z = q . z ; ::_thesis: ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is base_rotation & f . p = q ) then f . p = q by A67, FUNCT_7:35; hence ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is base_rotation & f . p = q ) ; ::_thesis: verum end; supposeA69: (f . p) . z = - (q . z) ; ::_thesis: ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is base_rotation & f . p = q ) 1 < n by A1, A7, XXREAL_0:1; then 1 + 1 <= n by NAT_1:13; then consider h being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) such that A70: h is base_rotation and A71: h . (f . p) = (f . p) +* (z,(- ((f . p) . z))) by A59, Th34; dom (h * f) = the carrier of (TOP-REAL n) by FUNCT_2:52; then (h * f) . p = (f . p) +* (z,(- ((f . p) . z))) by A71, FUNCT_1:12 .= q by A60, A62, A66, A69, FINSEQ_1:14 ; hence ex f being additive homogeneous Function of (TOP-REAL n),(TOP-REAL n) st ( f is base_rotation & f . p = q ) by A70; ::_thesis: verum end; end; end; end; end;