The Mizar article:

Operations on Subspaces in Real Linear Space

by
Wojciech A. Trybulec

Received September 20, 1989

Copyright (c) 1989 Association of Mizar Users

MML identifier: RLSUB_2
[ MML identifier index ]


environ

 vocabulary RLVECT_1, RLSUB_1, BOOLE, ARYTM_1, FUNCT_1, RELAT_1, TARSKI,
      MCART_1, BINOP_1, LATTICES, RLSUB_2;
 notation TARSKI, XBOOLE_0, ZFMISC_1, SUBSET_1, BINOP_1, RELAT_1, FUNCT_1,
      LATTICES, REAL_1, RELSET_1, STRUCT_0, RLVECT_1, RLSUB_1, DOMAIN_1;
 constructors BINOP_1, LATTICES, REAL_1, RLSUB_1, DOMAIN_1, MEMBERED, XBOOLE_0;
 clusters LATTICES, RLVECT_1, RLSUB_1, STRUCT_0, RELSET_1, SUBSET_1, MEMBERED,
      ZFMISC_1, XBOOLE_0;
 requirements NUMERALS, SUBSET, BOOLE;
 definitions LATTICES, RLSUB_1, TARSKI, STRUCT_0, XBOOLE_0;
 theorems BINOP_1, FUNCT_1, FUNCT_2, LATTICES, MCART_1, ORDERS_1, RLSUB_1,
      RLVECT_1, TARSKI, ZFMISC_1, RELAT_1, XBOOLE_0, XBOOLE_1, VECTSP_1,
      XCMPLX_0;
 schemes BINOP_1, FUNCT_1, ORDERS_2, RELSET_1, XBOOLE_0;

begin

 reserve V for RealLinearSpace;
 reserve W,W1,W2,W3 for Subspace of V;
 reserve u,u1,u2,v,v1,v2 for VECTOR of V;
 reserve a,a1,a2 for Real;
 reserve X,Y for set;
 reserve x,y,y1,y2 for set;

::
::  Definitions of sum and intersection of subspaces.
::

 definition let V; let W1,W2;
  func W1 + W2 -> strict Subspace of V means
   :Def1: the carrier of it = {v + u : v in W1 & u in W2};
  existence
   proof set VS = {v + u : v in W1 & u in W2};
         VS c= the carrier of V
        proof let x be set;
          assume x in VS;
           then ex v1,v2 st x = v1 + v2 & v1 in W1 & v2 in W2;
         hence thesis;
        end;
      then reconsider VS as Subset of V;
         0.V in W1 & 0.V in W2 & 0.V = 0.V + 0.V by RLSUB_1:25,RLVECT_1:10;
       then A1: 0.V in VS;
      reconsider V1 = the carrier of W1, V2 = the carrier of W2 as
           Subset of V by RLSUB_1:def 2;
       A2: VS = {v + u : v in V1 & u in V2}
        proof
         thus VS c= {v + u : v in V1 & u in V2}
          proof let x be set;
            assume x in VS;
             then consider v,u such that A3: x = v + u and A4: v in W1 & u in
 W2;
                v in V1 & u in V2 by A4,RLVECT_1:def 1;
           hence thesis by A3;
          end;
         let x be set;
          assume x in {v + u : v in V1 & u in V2};
           then consider v,u such that A5: x = v + u and A6: v in V1 & u in V2;
              v in W1 & u in W2 by A6,RLVECT_1:def 1;
         hence thesis by A5;
        end;
         V1 is lineary-closed & V2 is lineary-closed by RLSUB_1:42;
       then VS is lineary-closed by A2,RLSUB_1:9;
    hence thesis by A1,RLSUB_1:43;
   end;
  uniqueness by RLSUB_1:38;
 end;

 definition let V; let W1,W2;
  func W1 /\ W2 -> strict Subspace of V means
   :Def2: the carrier of it = (the carrier of W1) /\ (the carrier of W2);
  existence
   proof set VV = the carrier of V;
         set VW1 = the carrier of W1; set VW2 = the carrier of W2;
        VW1 c= VV & VW2 c= VV by RLSUB_1:def 2;
      then VW1 /\ VW2 c= VV /\ VV by XBOOLE_1:27;
     then reconsider V1 = VW1, V2 = VW2, V3 = VW1 /\ VW2 as Subset of
                                         the carrier of V by RLSUB_1:def 2;
        0.V in W1 & 0.V in W2 by RLSUB_1:25;
      then 0.V in VW1 & 0.V in VW2 by RLVECT_1:def 1;
      then A1: VW1 /\ VW2 <> {} by XBOOLE_0:def 3;
        V1 is lineary-closed & V2 is lineary-closed by RLSUB_1:42;
      then V3 is lineary-closed by RLSUB_1:10;
    hence thesis by A1,RLSUB_1:43;
   end;
  uniqueness by RLSUB_1:38;
 end;

::
::  Definitional theorems of sum and intersection of subspaces.
::

canceled 4;

theorem Th5:
 x in W1 + W2 iff
  (ex v1,v2 st v1 in W1 & v2 in W2 & x = v1 + v2)
   proof
    thus x in W1 + W2 implies (ex v1,v2 st v1 in W1 & v2 in W2 & x = v1 + v2)
     proof assume x in W1 + W2;
        then x in the carrier of W1 + W2 by RLVECT_1:def 1;
        then x in {v + u : v in W1 & u in W2} by Def1;
       then consider v1,v2 such that A1: x = v1 + v2 & v1 in W1 & v2 in W2;
      take v1,v2;
      thus thesis by A1;
     end;
     given v1,v2 such that A2: v1 in W1 & v2 in W2 & x = v1 + v2;
        x in {v + u : v in W1 & u in W2} by A2;
      then x in the carrier of W1 + W2 by Def1;
    hence thesis by RLVECT_1:def 1;
   end;

theorem Th6:
 v in W1 or v in W2 implies v in W1 + W2
  proof assume A1: v in W1 or v in W2;
      now per cases by A1;
     suppose A2: v in W1;
         v = v + 0.V & 0.V in W2 by RLSUB_1:25,RLVECT_1:10;
      hence thesis by A2,Th5;
     suppose A3: v in W2;
         v = 0.V + v & 0.V in W1 by RLSUB_1:25,RLVECT_1:10;
      hence thesis by A3,Th5;
    end;
   hence thesis;
  end;

theorem Th7:
 x in W1 /\ W2 iff x in W1 & x in W2
  proof
      x in W1 /\ W2 iff x in the carrier of W1 /\ W2 by RLVECT_1:def 1;
    then x in W1 /\ W2 iff x in
 (the carrier of W1) /\ (the carrier of W2) by Def2;
    then x in W1 /\ W2 iff x in the carrier of W1 & x in the carrier of W2
                                                                   by XBOOLE_0:
def 3;
   hence thesis by RLVECT_1:def 1;
  end;

 Lm1: W1 + W2 = W2 + W1
  proof set A = {v + u : v in W1 & u in W2};
        set B = {v + u : v in W2 & u in W1};
    A1: the carrier of W1 + W2 = A & the carrier of W2 + W1 = B by Def1;
    A2: A c= B
     proof let x be set;
       assume x in A;
        then consider v,u such that A3: x = v + u and A4: v in W1 & u in W2;
      thus thesis by A3,A4;
     end;
      B c= A
     proof let x be set;
       assume x in B;
        then consider v,u such that A5: x = v + u and A6: v in W2 & u in W1;
      thus thesis by A5,A6;
     end;
   then A = B by A2,XBOOLE_0:def 10;
   hence thesis by A1,RLSUB_1:38;
  end;

 Lm2: the carrier of W1 c= the carrier of W1 + W2
  proof let x be set;
  set A = {v + u : v in W1 & u in W2};
    assume x in the carrier of W1;
     then reconsider v = x as Element of W1;
     reconsider v as VECTOR of V by RLSUB_1:18;
        v in W1 & 0.V in
 W2 & v = v + 0.V by RLSUB_1:25,RLVECT_1:10,def 1;
      then x in A;
   hence thesis by Def1;
  end;

 Lm3:for W2 being strict Subspace of V holds
  the carrier of W1 c= the carrier of W2 implies W1 + W2 = W2
  proof let W2 be strict Subspace of V;
   assume A1: the carrier of W1 c= the carrier of W2;
      the carrier of W1 + W2 = the carrier of W2
     proof
      thus the carrier of W1 + W2 c= the carrier of W2
       proof
        let x be set;
         assume x in the carrier of W1 + W2;
           then x in {v + u : v in W1 & u in W2} by Def1;
          then consider v,u such that A2: x = v + u and A3: v in W1 and A4: u
in W2;
             W1 is Subspace of W2 by A1,RLSUB_1:36;
           then v in W2 by A3,RLSUB_1:16;
           then v + u in W2 by A4,RLSUB_1:28;
        hence thesis by A2,RLVECT_1:def 1;
       end;
        W1 + W2 = W2 + W1 by Lm1;
      hence thesis by Lm2;
     end;
   hence thesis by RLSUB_1:38;
  end;

theorem
  for W being strict Subspace of V holds
 W + W = W by Lm3;

theorem
   W1 + W2 = W2 + W1 by Lm1;

theorem Th10:
 W1 + (W2 + W3) = (W1 + W2) + W3
  proof set A = {v + u : v in W1 & u in W2};
        set B = {v + u : v in W2 & u in W3};
        set C = {v + u : v in W1 + W2 & u in W3};
        set D = {v + u : v in W1 & u in W2 + W3};
    A1: the carrier of W1 + (W2 + W3) = D & the carrier of (W1 + W2) + W3 = C
                                                                        by Def1
;
    A2: D c= C
     proof let x be set;
       assume x in D;
        then consider v,u such that A3: x = v + u and
                                    A4: v in W1 and A5: u in W2 + W3;
           u in the carrier of W2 + W3 by A5,RLVECT_1:def 1;
         then u in B by Def1;
        then consider u1,u2 such that A6: u = u1 + u2 and
                                      A7: u1 in W2 and A8: u2 in W3;
         A9: v + u = (v + u1) + u2 by A6,RLVECT_1:def 6;
           v + u1 in A by A4,A7;
         then v + u1 in the carrier of W1 + W2 by Def1;
         then v + u1 in W1 + W2 by RLVECT_1:def 1;
      hence thesis by A3,A8,A9;
     end;
      C c= D
     proof let x be set;
       assume x in C;
        then consider v,u such that A10: x = v + u and
                                    A11: v in W1 + W2 and A12: u in W3;
           v in the carrier of W1 + W2 by A11,RLVECT_1:def 1;
         then v in A by Def1;
        then consider u1,u2 such that A13: v = u1 + u2 and
                                      A14: u1 in W1 and A15: u2 in W2;
         A16: v + u =u1 + (u2 + u) by A13,RLVECT_1:def 6;
           u2 + u in B by A12,A15;
         then u2 + u in the carrier of W2 + W3 by Def1;
         then u2 + u in W2 + W3 by RLVECT_1:def 1;
      hence thesis by A10,A14,A16;
     end;
    then D = C by A2,XBOOLE_0:def 10;
   hence thesis by A1,RLSUB_1:38;
  end;

theorem Th11:
 W1 is Subspace of W1 + W2 & W2 is Subspace of W1 + W2
  proof the carrier of W1 c= the carrier of W1 + W2 by Lm2;
   hence W1 is Subspace of W1 + W2 by RLSUB_1:36;
      the carrier of W2 c= the carrier of W2 + W1 by Lm2;
    then the carrier of W2 c= the carrier of W1 + W2 by Lm1;
   hence thesis by RLSUB_1:36;
  end;

theorem Th12:
 for W2 being strict Subspace of V holds
 W1 is Subspace of W2 iff W1 + W2 = W2
  proof let W2 be strict Subspace of V;
   thus W1 is Subspace of W2 implies W1 + W2 = W2
    proof assume W1 is Subspace of W2;
      then the carrier of W1 c= the carrier of W2 by RLSUB_1:def 2;
     hence thesis by Lm3;
    end;
   thus thesis by Th11;
  end;

theorem Th13:
 for W being strict Subspace of V holds
 (0).V + W = W & W + (0).V = W
  proof let W be strict Subspace of V;
      (0).V is Subspace of W by RLSUB_1:51;
    then the carrier of (0).V c= the carrier of W by RLSUB_1:def 2;
   hence (0).V + W = W by Lm3;
   hence thesis by Lm1;
  end;

theorem Th14:
 (0).V + (Omega).V = the RLSStruct of V & (Omega).
V + (0).V = the RLSStruct of V
  proof
   thus (0).V + (Omega).V = (Omega).V by Th13
                     .= the RLSStruct of V by RLSUB_1:def 4;
   hence thesis by Lm1;
  end;

theorem Th15:
 (Omega).V + W = the RLSStruct of V & W + (Omega).V = the RLSStruct of V
  proof
A1: the carrier of V = the carrier of the RLSStruct of V;
      the carrier of W c= the carrier of V by RLSUB_1:def 2;
    then the carrier of W c= the carrier of (Omega).V by A1,RLSUB_1:def 4;
    then W + (Omega).V = (Omega).V by Lm3
                  .= the RLSStruct of V by RLSUB_1:def 4;
   hence thesis by Lm1;
  end;

theorem
   for V being strict RealLinearSpace holds
 (Omega).V + (Omega).V = V by Th15;

theorem
   for W being strict Subspace of V holds
 W /\ W = W
  proof let W be strict Subspace of V;
     the carrier of W = (the carrier of W) /\ (the carrier of W);
   hence thesis by Def2;
  end;

theorem Th18:
 W1 /\ W2 = W2 /\ W1
  proof the carrier of W1 /\ W2 = (the carrier of W2) /\
 (the carrier of W1) by Def2;
   hence thesis by Def2;
  end;

theorem Th19:
 W1 /\ (W2 /\ W3) = (W1 /\ W2) /\ W3
  proof
        set V1 = the carrier of W1; set V2 = the carrier of W2;
        set V3 = the carrier of W3;
      the carrier of W1 /\ (W2 /\ W3) = V1 /\ (the carrier of W2 /\ W3) by Def2
        .= V1 /\ (V2 /\ V3) by Def2
        .= (V1 /\ V2) /\ V3 by XBOOLE_1:16
        .= (the carrier of W1 /\ W2) /\ V3 by Def2;
   hence thesis by Def2;
  end;

Lm4: the carrier of W1 /\ W2 c= the carrier of W1
 proof
     the carrier of W1 /\ W2 = (the carrier of W1) /\
 (the carrier of W2) by Def2;
  hence thesis by XBOOLE_1:17;
 end;

theorem Th20:
 W1 /\ W2 is Subspace of W1 & W1 /\ W2 is Subspace of W2
  proof the carrier of W1 /\ W2 c= the carrier of W1 by Lm4;
   hence W1 /\ W2 is Subspace of W1 by RLSUB_1:36;
      the carrier of W2 /\ W1 c= the carrier of W2 by Lm4;
    then the carrier of W1 /\ W2 c= the carrier of W2 by Th18;
   hence thesis by RLSUB_1:36;
  end;

theorem Th21:
 for W1 being strict Subspace of V holds
 W1 is Subspace of W2 iff W1 /\ W2 = W1
  proof let W1 be strict Subspace of V;
   thus W1 is Subspace of W2 implies W1 /\ W2 = W1
    proof
      assume W1 is Subspace of W2;
       then A1: the carrier of W1 c= the carrier of W2 by RLSUB_1:def 2;
         the carrier of W1 /\ W2 = (the carrier of W1) /\ (the carrier of W2)
                                                                    by Def2;
       then the carrier of W1 /\ W2 = the carrier of W1 by A1,XBOOLE_1:28;
     hence thesis by RLSUB_1:38;
    end;
   thus thesis by Th20;
  end;

theorem Th22:
 (0).V /\ W = (0).V & W /\ (0).V = (0).V
  proof
    A1: the carrier of (0).V /\ W
           = (the carrier of (0).V) /\ (the carrier of W) by Def2
          .= {0.V} /\ (the carrier of W) by RLSUB_1:def 3;
      0.V in W by RLSUB_1:25;
    then 0.V in the carrier of W by RLVECT_1:def 1;
    then {0.V} c= the carrier of W by ZFMISC_1:37;
    then {0.V} /\ (the carrier of W) = {0.V} & the carrier of (0).V = {0.V}
                                                        by RLSUB_1:def 3,
XBOOLE_1:28;
   hence (0).V /\ W = (0).V by A1,RLSUB_1:38;
   hence thesis by Th18;
  end;

theorem
   (0).V /\ (Omega).V = (0).V & (Omega).V /\ (0).V = (0).V by Th22;

theorem Th24:
 for W being strict Subspace of V holds
 (Omega).V /\ W = W & W /\ (Omega).V = W
  proof let W be strict Subspace of V;
   (Omega).V = the RLSStruct of V by RLSUB_1:def 4;
    then A1: the carrier of (Omega).
V /\ W = (the carrier of V) /\ (the carrier of W)
                                                              by Def2;
      the carrier of W c= the carrier of V by RLSUB_1:def 2;
    then the carrier of (Omega).V /\ W = the carrier of W by A1,XBOOLE_1:28;
   hence (Omega).V /\ W = W by RLSUB_1:38;
   hence thesis by Th18;
  end;

theorem
   for V being strict RealLinearSpace holds
 (Omega).V /\ (Omega).V = V
  proof let V be strict RealLinearSpace;
   thus (Omega).V /\ (Omega).V = (Omega).V by Th24
                     .= V by RLSUB_1:def 4;
  end;

Lm5: the carrier of W1 /\ W2 c= the carrier of W1 + W2
 proof the carrier of W1 /\ W2 c= the carrier of W1 &
       the carrier of W1 c= the carrier of W1 + W2 by Lm2,Lm4;
  hence thesis by XBOOLE_1:1;
 end;

theorem
   W1 /\ W2 is Subspace of W1 + W2
  proof the carrier of W1 /\ W2 c= the carrier of W1 + W2 by Lm5;
   hence thesis by RLSUB_1:36;
  end;

Lm6: the carrier of (W1 /\ W2) + W2 = the carrier of W2
 proof
  thus the carrier of (W1 /\ W2) + W2 c= the carrier of W2
   proof let x be set;
     assume x in the carrier of (W1 /\ W2) + W2;
       then x in {u + v : u in W1 /\ W2 & v in W2} by Def1;
      then consider u,v such that A1: x = u + v and
                                  A2: u in W1 /\ W2 and A3: v in W2;
         u in W2 by A2,Th7;
       then u + v in W2 by A3,RLSUB_1:28;
    hence thesis by A1,RLVECT_1:def 1;
   end;
  let x be set;
   assume A4: x in the carrier of W2;
      the carrier of W2 c= the carrier of W2 + (W1 /\ W2) by Lm2;
    then the carrier of W2 c= the carrier of (W1 /\ W2) + W2 by Lm1;
  hence thesis by A4;
 end;

theorem Th27:
 for W2 being strict Subspace of V holds
 (W1 /\ W2) + W2 = W2
  proof let W2 be strict Subspace of V;
     the carrier of (W1 /\ W2) + W2 = the carrier of W2 by Lm6;
   hence thesis by RLSUB_1:38;
  end;

Lm7: the carrier of W1 /\ (W1 + W2) = the carrier of W1
 proof
  thus the carrier of W1 /\ (W1 + W2) c= the carrier of W1
   proof let x be set;
     assume A1: x in the carrier of W1 /\ (W1 + W2);
        the carrier of W1 /\ (W1 + W2) =
         (the carrier of W1) /\ (the carrier of W1 + W2) by Def2;
    hence thesis by A1,XBOOLE_0:def 3;
   end;
  let x be set;
   assume A2: x in the carrier of W1;
       the carrier of W1 c= the carrier of V by RLSUB_1:def 2;
    then reconsider x1 = x as Element of V by A2;

       x1 + 0.V = x1 & 0.V in W2 & x in W1 by A2,RLSUB_1:25,RLVECT_1:10,def 1;
     then x in {u + v : u in W1 & v in W2};
     then x in the carrier of W1 + W2 by Def1;
     then x in
 (the carrier of W1) /\ (the carrier of W1 + W2) by A2,XBOOLE_0:def 3;
  hence thesis by Def2;
 end;

theorem Th28:
 for W1 being strict Subspace of V holds
 W1 /\ (W1 + W2) = W1
  proof let W1 be strict Subspace of V;
     the carrier of W1 /\ (W1 + W2) = the carrier of W1 by Lm7;
   hence thesis by RLSUB_1:38;
  end;

Lm8: the carrier of (W1 /\ W2) + (W2 /\ W3) c= the carrier of W2 /\ (W1 + W3)
 proof
  let x be set;
   assume x in the carrier of (W1 /\ W2) + (W2 /\ W3);
     then x in {u + v : u in W1 /\ W2 & v in W2 /\ W3} by Def1;
    then consider u,v such that A1: x = u + v and
                                A2: u in W1 /\ W2 & v in W2 /\ W3;
       u in W1 & u in W2 & v in W2 & v in W3 by A2,Th7;
     then x in W1 + W3 & x in W2 by A1,Th5,RLSUB_1:28;
     then x in W2 /\ (W1 + W3) by Th7;
  hence thesis by RLVECT_1:def 1;
 end;

theorem
   (W1 /\ W2) + (W2 /\ W3) is Subspace of W2 /\ (W1 + W3)
  proof
      the carrier of (W1 /\ W2) + (W2 /\ W3) c= the carrier of W2 /\ (W1 + W3)
                                                                     by Lm8;
   hence thesis by RLSUB_1:36;
  end;

Lm9: W1 is Subspace of W2 implies
      the carrier of W2 /\ (W1 + W3) = the carrier of (W1 /\ W2) + (W2 /\ W3)
 proof assume A1: W1 is Subspace of W2;
  thus the carrier of W2 /\ (W1 + W3) c= the carrier of (W1 /\ W2) + (W2 /\ W3
)
   proof let x be set;
     assume x in the carrier of W2 /\ (W1 + W3);
      then A2: x in (the carrier of W2) /\ (the carrier of W1 + W3) by Def2;
      then x in the carrier of W1 + W3 by XBOOLE_0:def 3;
      then x in {u + v : u in W1 & v in W3} by Def1;
     then consider u1,v1 such that A3: x = u1 + v1 and
                                   A4: u1 in W1 and A5: v1 in W3;
      A6: u1 in W2 by A1,A4,RLSUB_1:16;
      then A7: u1 in W1 /\ W2 by A4,Th7;

        x in the carrier of W2 by A2,XBOOLE_0:def 3;
      then u1 + v1 in W2 by A3,RLVECT_1:def 1;
      then (v1 + u1) - u1 in W2 by A6,RLSUB_1:31;
      then v1 + (u1 - u1) in W2 by RLVECT_1:42;
      then v1 + 0.V in W2 by RLVECT_1:28;
      then v1 in W2 by RLVECT_1:10;
      then v1 in W2 /\ W3 by A5,Th7;
      then x in (W1 /\ W2) + (W2 /\ W3) by A3,A7,Th5;
    hence thesis by RLVECT_1:def 1;
   end;
 thus thesis by Lm8;
end;

theorem Th30:
 W1 is Subspace of W2 implies W2 /\ (W1 + W3) = (W1 /\ W2) + (W2 /\ W3)
  proof assume W1 is Subspace of W2;
    then the carrier of W2 /\ (W1 + W3) = the carrier of (W1 /\ W2) + (W2 /\
 W3)
                                                                        by Lm9;
   hence thesis by RLSUB_1:38;
  end;

Lm10: the carrier of W2 + (W1 /\ W3) c= the carrier of (W1 + W2) /\ (W2 + W3)
   proof let x be set;
     assume x in the carrier of W2 + (W1 /\ W3);
       then x in {u + v : u in W2 & v in W1 /\ W3} by Def1;
      then consider u,v such that A1: x = u + v and
                                  A2: u in W2 and A3: v in W1 /\ W3;
         v in W1 & v in W3 & x = v + u by A1,A3,Th7;
       then x in {v1 + v2 : v1 in W1 & v2 in W2} &
            x in {u1 + u2 : u1 in W2 & u2 in W3} by A2;
       then x in the carrier of W1 + W2 & x in the carrier of W2 + W3 by Def1;
       then x in (the carrier of W1 + W2) /\ (the carrier of W2 + W3) by
XBOOLE_0:def 3;
    hence thesis by Def2;
   end;

theorem
   W2 + (W1 /\ W3) is Subspace of (W1 + W2) /\ (W2 + W3)
  proof the carrier of W2 + (W1 /\ W3) c= the carrier of (W1 + W2) /\ (W2 + W3
)
                                                                     by Lm10;
   hence thesis by RLSUB_1:36;
  end;

Lm11: W1 is Subspace of W2 implies
        the carrier of W2 + (W1 /\ W3) = the carrier of (W1 + W2) /\ (W2 + W3)
 proof assume A1: W1 is Subspace of W2;
   reconsider V2 = the carrier of W2 as Subset of V by RLSUB_1:
def 2;
  thus the carrier of W2 + (W1 /\ W3) c= the carrier of (W1 + W2) /\ (W2 + W3)
                                                                     by Lm10;
  let x be set;
   assume x in the carrier of (W1 + W2) /\ (W2 + W3);
     then x in (the carrier of W1 + W2) /\ (the carrier of W2 + W3) by Def2;
     then x in the carrier of W1 + W2 by XBOOLE_0:def 3;
     then x in {u1 + u2 : u1 in W1 & u2 in W2} by Def1;
    then consider u1,u2 such that A2: x = u1 + u2 and
                                  A3: u1 in W1 and A4: u2 in W2;
     A5: u1 in the carrier of W1 by A3,RLVECT_1:def 1;
     A6: u2 in the carrier of W2 by A4,RLVECT_1:def 1;
       the carrier of W1 c= the carrier of W2 by A1,RLSUB_1:def 2;
     then u1 in the carrier of W2 & V2 is lineary-closed
                                                  by A5,RLSUB_1:42;
     then u1 + u2 in V2 by A6,RLSUB_1:def 1;
     then A7: u1 + u2 in W2 by RLVECT_1:def 1;
     A8: 0.V in W1 /\ W3 by RLSUB_1:25;
       (u1 + u2) + 0.V = u1 + u2 by RLVECT_1:10;
     then x in {u + v : u in W2 & v in W1 /\ W3} by A2,A7,A8;
  hence thesis by Def1;
 end;

theorem
   W1 is Subspace of W2 implies W2 + (W1 /\ W3) = (W1 + W2) /\ (W2 + W3)
  proof assume W1 is Subspace of W2;
    then the carrier of W2 + (W1 /\ W3) = the carrier of (W1 + W2) /\ (W2 + W3
)
                                                                      by Lm11;
   hence thesis by RLSUB_1:38;
  end;

theorem Th33:
 W1 is strict Subspace of W3 implies W1 + (W2 /\ W3) = (W1 + W2) /\ W3
  proof assume A1: W1 is strict Subspace of W3;
   thus (W1 + W2) /\ W3 = W3 /\ (W1 + W2) by Th18
                      .= (W1 /\ W3) + (W3 /\ W2) by A1,Th30
                      .= W1 + (W3 /\ W2) by A1,Th21
                      .= W1 + (W2 /\ W3) by Th18;
  end;

theorem
   for W1,W2 being strict Subspace of V holds
 W1 + W2 = W2 iff W1 /\ W2 = W1
  proof let W1,W2 be strict Subspace of V;
      (W1 + W2 = W2 iff W1 is Subspace of W2) &
    (W1 /\ W2 = W1 iff W1 is Subspace of W2) by Th12,Th21;
   hence thesis;
  end;

theorem
   for W2,W3 being strict Subspace of V holds
 W1 is Subspace of W2 implies W1 + W3 is Subspace of W2 + W3
  proof let W2,W3 be strict Subspace of V;
   assume A1: W1 is Subspace of W2;
      (W1 + W3) + (W2 + W3) = (W1 + W3) + (W3 + W2) by Lm1
                         .= ((W1 + W3) + W3) + W2 by Th10
                         .= (W1 + (W3 + W3)) + W2 by Th10
                         .= (W1 + W3) + W2 by Lm3
                         .= W1 + (W3 + W2) by Th10
                         .= W1 + (W2 + W3) by Lm1
                         .= (W1 + W2) + W3 by Th10
                         .= W2 + W3 by A1,Th12;
   hence thesis by Th12;
  end;

theorem
   (ex W st the carrier of W = (the carrier of W1) \/ (the carrier of W2)) iff
  W1 is Subspace of W2 or W2 is Subspace of W1
   proof
     set VW1 = the carrier of W1; set VW2 = the carrier of W2;
    thus (ex W st the carrier of W = (the carrier of W1) \/
 (the carrier of W2))
          implies W1 is Subspace of W2 or W2 is Subspace of W1
     proof given W such that
            A1: the carrier of W = (the carrier of W1) \/ (the carrier of W2);
        set VW = the carrier of W;
       assume not W1 is Subspace of W2 & not W2 is Subspace of W1;
         then A2: not VW1 c= VW2 & not VW2 c= VW1 by RLSUB_1:36;
        then consider x being set such that A3: x in VW1 and A4: not x in VW2
                                                                   by TARSKI:
def 3;
        consider y being set such that A5: y in VW2 and A6: not y in VW1
                                                                 by A2,TARSKI:
def 3;
        reconsider x as Element of VW1 by A3;

        reconsider x as VECTOR of V by RLSUB_1:18;
        reconsider y as Element of VW2 by A5;

        reconsider y as VECTOR of V by RLSUB_1:18;
        reconsider A1 = VW as Subset of V by RLSUB_1:def 2;
           x in VW & y in VW & A1 is lineary-closed by A1,RLSUB_1:42,XBOOLE_0:
def 2;
         then A7: x + y in VW by RLSUB_1:def 1;

         A8: now assume A9: x + y in VW1;
           reconsider A2 = VW1 as Subset of V by RLSUB_1:def 2;
              A2 is lineary-closed by RLSUB_1:42;
            then (y + x) - x in VW1 by A9,RLSUB_1:6;
            then y + (x - x) in VW1 by RLVECT_1:42;
            then y + 0.V in VW1 by RLVECT_1:28;
          hence contradiction by A6,RLVECT_1:10;
         end;
           now assume A10: x + y in VW2;
           reconsider A2 = VW2 as Subset of V by RLSUB_1:def 2;
              A2 is lineary-closed by RLSUB_1:42;
            then (x + y) - y in VW2 by A10,RLSUB_1:6;
            then x + (y - y) in VW2 by RLVECT_1:42;
            then x + 0.V in VW2 by RLVECT_1:28;
          hence contradiction by A4,RLVECT_1:10;
         end;
      hence thesis by A1,A7,A8,XBOOLE_0:def 2;
     end;
     assume A11: W1 is Subspace of W2 or W2 is Subspace of W1;
      A12: now assume W1 is Subspace of W2;
        then VW1 c= VW2 by RLSUB_1:def 2;
        then VW1 \/ VW2 = VW2 by XBOOLE_1:12;
       hence thesis;
      end;
        now assume W2 is Subspace of W1;
        then VW2 c= VW1 by RLSUB_1:def 2;
        then VW1 \/ VW2 = VW1 by XBOOLE_1:12;
       hence thesis;
      end;
    hence thesis by A11,A12;
   end;

::
::  Introduction of a set of subspaces of real linear space.
::

 definition let V;
  func Subspaces(V) -> set means
   :Def3: for x holds x in it iff x is strict Subspace of V;
  existence
   proof
     defpred P[set] means
        (ex W being strict Subspace of V st $1 = the carrier of W);
     consider B being set such that
      A1: for x holds x in B iff x in bool(the carrier of V) & P[x]
           from Separation;
     defpred Q[set,set] means
       (ex W being strict Subspace of V st $2 = W & $1 = the carrier of W);
      A2: for x,y1,y2 st Q[x,y1] & Q[x,y2]
         holds y1 = y2 by RLSUB_1:38;
     consider f being Function such that
      A3: for x,y holds [x,y] in f iff x in B & Q[x,y]
           from GraphFunc(A2);
        for x holds x in B iff ex y st [x,y] in f
       proof let x;
        thus x in B implies ex y st [x,y] in f
         proof assume A4: x in B;
           then consider W being strict Subspace of V such that
           A5: x = the carrier of W by A1;
           reconsider y = W as set;
          take y;
          thus thesis by A3,A4,A5;
         end;
         given y such that A6: [x,y] in f;
        thus thesis by A3,A6;
       end;
      then A7: B = dom f by RELAT_1:def 4;
        for y holds y in rng f iff y is strict Subspace of V
       proof let y;
        thus y in rng f implies y is strict Subspace of V
         proof assume y in rng f;
           then consider x such that A8: x in dom f & y = f.x by FUNCT_1:def 5;
              [x,y] in f by A8,FUNCT_1:def 4;
            then ex W being strict Subspace of V st y = W &
 x = the carrier of W by A3;
          hence thesis;
         end;
         assume y is strict Subspace of V;
          then reconsider W = y as strict Subspace of V;
          reconsider x = the carrier of W as set;
             the carrier of W c= the carrier of V by RLSUB_1:def 2;
           then A9: x in dom f by A1,A7;
           then [x,y] in f by A3,A7;
           then y = f.x by A9,FUNCT_1:def 4;
        hence thesis by A9,FUNCT_1:def 5;
       end;
    hence thesis;
   end;
  uniqueness
   proof let D1,D2 be set;
     assume A10: for x holds x in D1 iff x is strict Subspace of V;
     assume A11: for x holds x in D2 iff x is strict Subspace of V;
        now let x;
       thus x in D1 implies x in D2
        proof assume x in D1;
          then x is strict Subspace of V by A10;
         hence thesis by A11;
        end;
        assume x in D2;
         then x is strict Subspace of V by A11;
       hence x in D1 by A10;
      end;
    hence thesis by TARSKI:2;
   end;
 end;

 definition let V;
  cluster Subspaces(V) -> non empty;
  coherence
  proof
   consider x being strict Subspace of V;
     x in Subspaces(V) by Def3;
   hence thesis;
 end;
end;



canceled 2;

theorem
   for V being strict RealLinearSpace holds
 V in Subspaces(V)
  proof let V be strict RealLinearSpace;
     (Omega).V in Subspaces(V) by Def3;
   hence thesis by RLSUB_1:def 4;
  end;

::
::  Introduction of the direct sum of subspaces and
::  linear complement of subspace.
::

definition let V; let W1,W2;
 pred V is_the_direct_sum_of W1,W2 means
  :Def4: the RLSStruct of V = W1 + W2 & W1 /\ W2 = (0).V;
end;

Lm12:
 for V being RealLinearSpace, W being strict Subspace of V holds
 (for v being VECTOR of V holds v in W) implies W = the RLSStruct of V
 proof let V be RealLinearSpace, W be strict Subspace of V;
  assume A1: for v being VECTOR of V holds v in W;
     now let v be VECTOR of V;
    thus v in W implies v in (Omega).V
     proof assume v in W;
         v in the RLSStruct of V by RLVECT_1:3;
      hence v in (Omega).V by RLSUB_1:def 4;
     end;
     assume v in (Omega).V;
    thus v in W by A1;
   end;
   then W = (Omega).V by RLSUB_1:39;
  hence thesis by RLSUB_1:def 4;
 end;

Lm13:
 for V being RealLinearSpace, W1,W2 being Subspace of V holds
 W1 + W2 = the RLSStruct of V iff
 for v being VECTOR of V
  ex v1,v2 being VECTOR of V st v1 in W1 & v2 in W2 & v = v1 + v2
  proof let V be RealLinearSpace, W1,W2 be Subspace of V;
   thus W1 + W2 = the RLSStruct of V implies
    for v being VECTOR of V
     ex v1,v2 being VECTOR of V st v1 in W1 & v2 in W2 & v = v1 + v2
    proof assume A1: W1 + W2 = the RLSStruct of V;
     let v be VECTOR of V;
        v in the RLSStruct of V by RLVECT_1:3;
     hence thesis by A1,Th5;
    end;
    assume
    A2: for v being VECTOR of V
        ex v1,v2 being VECTOR of V st v1 in W1 & v2 in W2 & v = v1 + v2;
       now let u be VECTOR of V;
          ex v1,v2 being VECTOR of V st v1 in W1 & v2 in W2 & u = v1 + v2 by A2
;
      hence u in W1 + W2 by Th5;
     end;
   hence thesis by Lm12;
  end;

Lm14:
for V being add-associative right_zeroed right_complementable
  (non empty LoopStr),
    v,v1,v2 being Element of V holds
  v = v1 + v2 iff v1 = v - v2
  proof
   let V be add-associative right_zeroed right_complementable
        (non empty LoopStr),
       v,v1,v2 be Element of V;
    thus v = v1 + v2 implies v1 = v - v2
      proof
        assume v = v1 + v2;
        hence v - v2 = v1 + (v2 - v2) by RLVECT_1:42
                    .= v1 + 0.V by VECTSP_1:66
                    .= v1 by RLVECT_1:10;
      end;
    thus v1 = v - v2 implies v = v1 + v2
      proof
        assume v1 = v - v2;
        hence v1 + v2 = v + (-v2) + v2 by RLVECT_1:def 11
                      .= v + (-v2 + v2) by RLVECT_1:def 6
                      .= v + 0.V by RLVECT_1:16
                      .= v by RLVECT_1:10;
      end;
  end;

Lm15:
for V being RealLinearSpace, W being Subspace of V
 ex C being strict Subspace of V st V is_the_direct_sum_of C,W
proof let V be RealLinearSpace; let W be Subspace of V;
  defpred P[set] means
      ex W1 being strict Subspace of V st $1 = W1 & W /\ W1 = (0).V;
  consider X such that A1: x in X iff (x in Subspaces(V) & P[x])
     from Separation;
     W /\ (0).V = (0).V & (0).V in Subspaces(V) by Def3,Th22;
  then reconsider X as non empty set by A1;
   A2: now let x be Element of X;
       x in Subspaces(V) by A1;
    hence x is strict Subspace of V by Def3;
   end;
  defpred P[set,set] means  ex W1,W2 being strict Subspace of V st
        $1 = W1 & $2 = W2 & W1 is Subspace of W2;
  consider R being Relation of X such that
   A3: for x,y being Element of X holds
       [x,y] in R iff P[x,y]  from Rel_On_Dom_Ex;
   defpred P[set,set] means [$1,$2] in R;
   A4: for x being Element of X holds  P[x,x]
   proof
   now let x be Element of X;
     reconsider W = x as strict Subspace of V by A2;
        W is Subspace of W by RLSUB_1:33;
    hence [x,x] in R by A3;
   end;
   hence thesis;
   end;
 A5: for x,y being Element of X st P[x,y] & P[y,x] holds x=y
  proof
   now let x,y be Element of X;
    assume that A6: [x,y] in R and A7: [y,x] in R;
 A8:     ex W1,W2 being strict Subspace of V st
 x = W1 & y = W2 & W1 is Subspace of W2 by A3,A6;
        ex W3,W4 being strict Subspace of V st
 y = W3 & x = W4 & W3 is Subspace of W4 by A3,A7;
    hence x = y by A8,RLSUB_1:34;
   end;
   hence thesis;
   end;
 A9: for x,y,z being Element of X st P[x,y] & P[y,z] holds P[x,z]
  proof
   now let x,y,z be Element of X;
    assume that A10: [x,y] in R and A11: [y,z] in R;
     consider W1,W2 being strict Subspace of V such that
     A12: x = W1 & y = W2 & W1 is Subspace of W2 by A3,A10;
     consider W3,W4 being strict Subspace of V such that
     A13: y = W3 & z = W4 & W3 is Subspace of W4 by A3,A11;
        W1 is strict Subspace of W4 by A12,A13,RLSUB_1:35;
    hence [x,z] in R by A3,A12,A13;
   end;
  hence thesis;
  end;
 A14:   for Y st Y c= X &
     (for x,y being Element of X st x in Y & y in Y
       holds P[x,y] or P[y,x]) holds
     ex y being Element of X st
      for x being Element of X st x in Y holds P[x,y]
  proof
   for Y st Y c= X &
        (for x,y being Element of X st x in Y & y in Y holds
          [x,y] in R or [y,x] in R)
        ex y being Element of X st
         for x being Element of X st x in Y holds [x,y] in R
    proof let Y;
      assume that A15: Y c= X and
        A16: for x,y being Element of X st x in Y & y in Y holds
             [x,y] in R or [y,x] in R;
          now per cases;
         suppose A17: Y = {};
           consider y being Element of X;
          take y' = y;
          let x be Element of X;
           assume x in Y;
          hence [x,y'] in R by A17;
         suppose A18: Y <> {};
         defpred P[set,set] means ex W1 being strict Subspace of V st
                 $1 = W1 & $2 = the carrier of W1;
           A19: for x,y1,y2 st x in Y & P[x,y1] & P[x,y2] holds y1 = y2;
           A20: for x st x in Y ex y st P[x,y]
            proof let x;
              assume x in Y;
                then consider W1 being strict Subspace of V such that
               A21: x = W1 and W /\ W1 = (0).V by A1,A15;
               reconsider y = the carrier of W1 as set;
             take y;
             take W1;
             thus thesis by A21;
            end;
          consider f being Function such that
           A22: dom f = Y and
           A23: for x st x in Y holds P[x,f.x] from FuncEx(A19,A20);
          set Z = union(rng f);
             now let x;
             assume x in Z;
              then consider Y' being set such that A24: x in Y' and
               A25: Y' in rng f by TARSKI:def 4;
              consider y such that A26: y in dom f and A27: f.y = Y'
                                                     by A25,FUNCT_1:def 5;
              consider W1 being strict Subspace of V such that y = W1 and
              A28: f.y = the carrier of W1 by A22,A23,A26;
                the carrier of W1 c= the carrier of V & x in the carrier of W1
                                                      by A24,A27,A28,RLSUB_1:
def 2
;
            hence x in the carrier of V;
           end;
          then reconsider Z as Subset of V by TARSKI:def 3;
          A29: rng f <> {} by A18,A22,RELAT_1:65;
          consider z being Element of rng f;
          consider z1 being set such that A30: z1 in dom f and
                                          A31: f.z1 = z by A29,FUNCT_1:def 5;
          consider W3 being strict Subspace of V such that
A32:         z1 = W3 & f.z1 = the carrier of W3 by A22,A23,A30;


           A33: Z <> {} by A29,A31,A32,ORDERS_1:91;
             Z is lineary-closed
            proof
             thus for v1,v2 being VECTOR of V st
              v1 in Z & v2 in Z holds v1 + v2 in Z
              proof let v1,v2 be VECTOR of V;
                assume that A34: v1 in Z and A35: v2 in Z;
                 consider Y1 being set such that A36: v1 in Y1 and
                  A37: Y1 in rng f by A34,TARSKI:def 4;
                 consider Y2 being set such that A38: v2 in Y2 and
                  A39: Y2 in rng f by A35,TARSKI:def 4;
                 consider y1 such that A40: y1 in dom f and
                                       A41: f.y1 = Y1 by A37,FUNCT_1:def 5;
                 consider y2 such that A42: y2 in dom f and
                                       A43: f.y2 = Y2 by A39,FUNCT_1:def 5;
                 consider W1 being strict Subspace of V such that
                  A44: y1 = W1 and
                  A45: f.y1 = the carrier of W1 by A22,A23,A40;
                 consider W2 being strict Subspace of V such that
                  A46: y2 = W2 and
                  A47: f.y2 = the carrier of W2 by A22,A23,A42;
                 reconsider y1,y2 as Element of X by A15,A22,A40,A42;
                    now per cases by A16,A22,A40,A42;
                   suppose [y1,y2] in R;
                     then ex W3,W4 being strict Subspace of V
                          st y1 = W3 & y2 = W4 & W3 is Subspace of W4 by A3;
                       then the carrier of W1 c= the carrier of W2
                                                            by A44,A46,RLSUB_1:
def 2;
                      then v1 in W2 & v2 in W2 by A36,A38,A41,A43,A45,A47,
RLVECT_1:def 1;
                      then v1 + v2 in W2 by RLSUB_1:28;
                      then A48: v1 + v2 in the carrier of W2 by RLVECT_1:def 1;
                        f.y2 in rng f by A42,FUNCT_1:def 5;
                    hence v1 + v2 in Z by A47,A48,TARSKI:def 4;
                   suppose [y2,y1] in R;
                     then ex W3,W4 being strict Subspace of V
                     st y2 = W3 & y1 = W4 & W3 is Subspace of W4 by A3;
                       then the carrier of W2 c= the carrier of W1
                                                            by A44,A46,RLSUB_1:
def 2;
                      then v1 in W1 & v2 in W1 by A36,A38,A41,A43,A45,A47,
RLVECT_1:def 1;
                      then v1 + v2 in W1 by RLSUB_1:28;
                      then A49: v1 + v2 in the carrier of W1 by RLVECT_1:def 1;
                        f.y1 in rng f by A40,FUNCT_1:def 5;
                    hence v1 + v2 in Z by A45,A49,TARSKI:def 4;
                  end;
               hence v1 + v2 in Z;
              end;
             let a; let v1 be VECTOR of V;
              assume v1 in Z;
               then consider Y1 being set such that A50: v1 in Y1 and
                A51: Y1 in rng f by TARSKI:def 4;
               consider y1 such that A52: y1 in dom f and
                                     A53: f.y1 = Y1 by A51,FUNCT_1:def 5;
               consider W1 being strict Subspace of V such that
                       y1 = W1 and
                A54: f.y1 = the carrier of W1 by A22,A23,A52;
                  v1 in W1 by A50,A53,A54,RLVECT_1:def 1;
                then a * v1 in W1 by RLSUB_1:29;
                then A55: a * v1 in the carrier of W1 by RLVECT_1:def 1;
                  f.y1 in rng f by A52,FUNCT_1:def 5;
             hence a * v1 in Z by A54,A55,TARSKI:def 4;
            end;
          then consider E being strict Subspace of V such that
           A56: Z = the carrier of E by A33,RLSUB_1:43;
             now let u be VECTOR of V;
            thus u in W /\ E implies u in (0).V
             proof assume A57: u in W /\ E;
                then u in E by Th7;
                then u in Z by A56,RLVECT_1:def 1;
               then consider Y1 being set such that A58: u in Y1 and
                A59: Y1 in rng f by TARSKI:def 4;
               consider y1 such that A60: y1 in dom f and
                                     A61: f.y1 = Y1 by A59,FUNCT_1:def 5;
               consider W1 being strict Subspace of V such that
                A62: y1 = W1 and
                A63: f.y1 = the carrier of W1 by A22,A23,A60;
A64:                u in W1 & u in W by A57,A58,A61,A63,Th7,RLVECT_1:def 1;

                  ex W2 being strict Subspace of V st
 y1 = W2 & W /\ W2 = (0).V by A1,A15,A22,A60;
              hence u in (0).V by A62,A64,Th7;
             end;
             assume u in (0).V;
              then u in the carrier of (0).V by RLVECT_1:def 1;
              then u in {0.V} by RLSUB_1:def 3;
              then u = 0.V by TARSKI:def 1;
            hence u in W /\ E by RLSUB_1:25;
           end;
           then A65: W /\ E = (0).V by RLSUB_1:39;
             E in Subspaces(V) by Def3;
          then reconsider y' = E as Element of X by A1,A65;
        take y = y';
        let x be Element of X;
         assume A66: x in Y;
          then consider W1 being strict Subspace of V such that
           A67: x = W1 and
           A68: f.x = the carrier of W1 by A23;
             now let u be VECTOR of V;
             assume u in W1;
              then A69: u in the carrier of W1 by RLVECT_1:def 1;
                the carrier of W1 in rng f by A22,A66,A68,FUNCT_1:def 5;
              then u in Z by A69,TARSKI:def 4;
            hence u in E by A56,RLVECT_1:def 1;
           end;
           then W1 is strict Subspace of E by RLSUB_1:37;
        hence [x,y] in R by A3,A67;
       end;
      hence thesis;
     end;
    hence thesis;
   end;
  consider x being Element of X such that
   A70: for y being Element of X st x <> y holds not P[x,y]
                                          from Zorn_Max(A4,A5,A9,A14);
  consider L being strict Subspace of V such that A71: x = L and
   A72: W /\ L = (0).V by A1;
 take L;
 thus the RLSStruct of V = L + W
  proof assume not thesis;
    then consider v being VECTOR of V such that
     A73: for v1,v2 being VECTOR of V holds
      not v1 in L or not v2 in W or v <> v1 + v2 by Lm13;
     A74: not v in L + W by A73,Th5;
    set A = {a * v : not contradiction};
    A75: 1 * v in A;
       now let x;
       assume x in A;
        then ex a st x = a * v;
      hence x in the carrier of V;
     end;
    then reconsider A as Subset of V by TARSKI:def 3;
       A is lineary-closed
      proof
       thus for v1,v2 being VECTOR of V st
        v1 in A & v2 in A holds v1 + v2 in A
        proof let v1,v2 be VECTOR of V;
          assume v1 in A;
           then consider a1 such that A76: v1 = a1 * v;
          assume v2 in A;
           then consider a2 such that A77: v2 = a2 * v;
              v1 + v2 = (a1 + a2) * v by A76,A77,RLVECT_1:def 9;
         hence v1 + v2 in A;
        end;
       let a; let v1 be VECTOR of V;
        assume v1 in A;
         then consider a1 such that A78: v1 = a1 * v;
            a * v1 = (a * a1) * v by A78,RLVECT_1:def 9;
       hence thesis;
      end;
    then consider Z being strict Subspace of V such that
    A79: the carrier of Z = A by A75,RLSUB_1:43;
       now let u be VECTOR of V;
      thus u in Z /\ (W + L) implies u in (0).V
       proof assume A80: u in Z /\ (W + L);
          then u in Z by Th7;
          then u in A by A79,RLVECT_1:def 1;
         then consider a such that A81: u = a * v;
            now assume A82: a <> 0;
              u in W + L by A80,Th7;
            then a" * (a * v) in W + L by A81,RLSUB_1:29;
            then (a" * a) * v in W + L by RLVECT_1:def 9;
            then 1 * v in W + L by A82,XCMPLX_0:def 7;
            then 1 * v in L + W by Lm1;
           hence contradiction by A74,RLVECT_1:def 9;
          end;
         then u = 0.V by A81,RLVECT_1:23;
        hence thesis by RLSUB_1:25;
       end;
       assume u in (0).V;
        then u in the carrier of (0).V by RLVECT_1:def 1;
        then u in {0.V} by RLSUB_1:def 3;
        then u = 0.V by TARSKI:def 1;
      hence u in Z /\ (W + L) by RLSUB_1:25;
     end;
     then A83: Z /\ (W + L) = (0).V by RLSUB_1:39;
       now let u be VECTOR of V;
      thus u in (Z + L) /\ W implies u in (0).V
       proof assume A84: u in (Z + L) /\ W;
          then u in Z + L by Th7;
         then consider v1,v2 being VECTOR of V such that
         A85: v1 in Z and A86: v2 in L and
                                       A87: u = v1 + v2 by Th5;
          A88: v1 = u - v2 by A87,Lm14;
          A89: u in W by A84,Th7;
          then u in W + L & v2 in W + L by A86,Th6;
          then v1 in W + L by A88,RLSUB_1:31;
          then v1 in Z /\ (W + L) by A85,Th7;
          then v1 in the carrier of (0).V by A83,RLVECT_1:def 1;
          then v1 in {0.V} by RLSUB_1:def 3;
          then v1 = 0.V by TARSKI:def 1;
          then v2 = u by A87,RLVECT_1:10;
        hence u in (0).V by A72,A86,A89,Th7;
       end;
       assume u in (0).V;
        then u in the carrier of (0).V by RLVECT_1:def 1;
        then u in {0.V} by RLSUB_1:def 3;
        then u = 0.V by TARSKI:def 1;
      hence u in (Z + L) /\ W by RLSUB_1:25;
     end;
     then (Z + L) /\ W = (0).V by RLSUB_1:39;
     then W /\ (Z + L) = (0).V & (Z + L) in Subspaces(V) by Def3,Th18;
    then reconsider x1 = Z + L as Element of X by A1;
       v = 0.V + v & v = v + 0.V & 0.V in W & 0.V in L by RLSUB_1:25,RLVECT_1:
10
;
     then A90: not v in L by A73;
       v in A by A75,RLVECT_1:def 9;
     then v in Z by A79,RLVECT_1:def 1;
     then A91: Z + L <> L by A90,Th6;
       L is Subspace of Z + L by Th11;
     then [x,x1] in R by A3,A71;
   hence contradiction by A70,A71,A91;
  end;
 thus L /\ W = (0).V by A72,Th18;
end;

definition let V be RealLinearSpace; let W be Subspace of V;
 mode Linear_Compl of W -> Subspace of V means
  :Def5: V is_the_direct_sum_of it,W;
 existence
  proof
      ex C being strict Subspace of V st V is_the_direct_sum_of C,W by Lm15;
   hence thesis;
  end;
end;

definition let V be RealLinearSpace; let W be Subspace of V;
 cluster strict Linear_Compl of W;
  existence
   proof
     consider C being strict Subspace of V such that
A1:      V is_the_direct_sum_of C,W by Lm15;
       C is Linear_Compl of W by A1,Def5;
    hence thesis;
   end;
end;

Lm16: V is_the_direct_sum_of W1,W2 implies V is_the_direct_sum_of W2,W1
 proof assume V is_the_direct_sum_of W1,W2;
   then the RLSStruct of V = W1 + W2 & W1 /\ W2 = (0).V by Def4;
   then the RLSStruct of V = W2 + W1 & W2 /\ W1 = (0).V by Lm1,Th18;
  hence thesis by Def4;
 end;

canceled 2;

theorem
   for V being RealLinearSpace, W1,W2 being Subspace of V holds
 V is_the_direct_sum_of W1,W2 implies W2 is Linear_Compl of W1
  proof let V be RealLinearSpace, W1,W2 be Subspace of V;
   assume V is_the_direct_sum_of W1,W2;
    then V is_the_direct_sum_of W2,W1 &
         the RLSStruct of V = W1 + W2 by Def4,Lm16;
   hence thesis by Def5;
  end;

theorem Th43:
 for V being RealLinearSpace, W being Subspace of V,
     L being Linear_Compl of W holds
 V is_the_direct_sum_of L,W & V is_the_direct_sum_of W,L
  proof let V be RealLinearSpace, W be Subspace of V,
            L be Linear_Compl of W;
   thus V is_the_direct_sum_of L,W by Def5;
   hence thesis by Lm16;
  end;

::
::  Theorems concerning the direct sum of a subspaces,
::  linear complement of a subspace and coset of a subspace.
::

theorem Th44:
 for V being RealLinearSpace, W being Subspace of V,
  L being Linear_Compl of W holds
 W + L = the RLSStruct of V & L + W = the RLSStruct of V
  proof let V be RealLinearSpace, W be Subspace of V,
            L be Linear_Compl of W;
     V is_the_direct_sum_of W,L by Th43;
   hence W + L = the RLSStruct of V by Def4;
   hence thesis by Lm1;
  end;

theorem Th45:
 for V being RealLinearSpace, W being Subspace of V,
   L being Linear_Compl of W holds
 W /\ L = (0).V & L /\ W = (0).V
  proof let V be RealLinearSpace, W be Subspace of V,
            L be Linear_Compl of W;
     V is_the_direct_sum_of W,L by Th43;
   hence W /\ L = (0).V by Def4;
   hence thesis by Th18;
  end;

theorem
   V is_the_direct_sum_of W1,W2 implies V is_the_direct_sum_of W2,W1 by Lm16;

theorem Th47:
 for V being RealLinearSpace holds
 V is_the_direct_sum_of (0).V,(Omega).V &
  V is_the_direct_sum_of (Omega).V,(0).V
   proof let V be RealLinearSpace;
      (0).V + (Omega).V = the RLSStruct of V & (0).V = (0).V /\ (Omega).
V by Th14,Th22;
    hence V is_the_direct_sum_of (0).V,(Omega).V by Def4;
    hence thesis by Lm16;
   end;

theorem
   for V being RealLinearSpace, W being Subspace of V,
   L being Linear_Compl of W holds
 W is Linear_Compl of L
  proof let V be RealLinearSpace, W be Subspace of V,
            L be Linear_Compl of W;
     V is_the_direct_sum_of L,W by Def5;
    then V is_the_direct_sum_of W,L by Lm16;
   hence thesis by Def5;
  end;

theorem
   for V being RealLinearSpace holds
 (0).V is Linear_Compl of (Omega).V &
  (Omega).V is Linear_Compl of (0).V
   proof let V be RealLinearSpace;
       V is_the_direct_sum_of (0).V,(Omega).V &
      V is_the_direct_sum_of (Omega).V,(0).V by Th47;
    hence thesis by Def5;
   end;

reserve C for Coset of W;
reserve C1 for Coset of W1;
reserve C2 for Coset of W2;

theorem Th50:
 C1 meets C2 implies C1 /\ C2 is Coset of W1 /\ W2
   proof assume A1: C1 /\ C2 <> {};
     consider v being Element of C1 /\ C2;
     reconsider v as Element of V by A1,TARSKI:def 3;
     set C = C1 /\ C2;
        v in C1 by A1,XBOOLE_0:def 3;
      then A2: C1 = v + W1 by RLSUB_1:94;
        v in C2 by A1,XBOOLE_0:def 3;
      then A3: C2 = v + W2 by RLSUB_1:94;
        C is Coset of W1 /\ W2
       proof
        take v;
        thus C c= v + W1 /\ W2
         proof let x;
          assume A4: x in C;
            then x in C1 by XBOOLE_0:def 3;
           then consider u1 such that A5: u1 in W1 and A6: x = v + u1
                                                             by A2,RLSUB_1:79;
              x in C2 by A4,XBOOLE_0:def 3;
           then consider u2 such that A7: u2 in W2 and A8: x = v + u2
                                                            by A3,RLSUB_1:79;
              u1 = u2 by A6,A8,RLVECT_1:21;
            then u1 in W1 /\ W2 by A5,A7,Th7;
            then x in {v + u : u in W1 /\ W2} by A6;
          hence thesis by RLSUB_1:def 5;
         end;
        let x;
         assume x in v + (W1 /\ W2);
          then consider u such that A9: u in
 W1 /\ W2 and A10: x = v + u by RLSUB_1:79;
             u in W1 & u in W2 by A9,Th7;
           then x in {v + u1 : u1 in W1} & x in {v + u2 : u2 in W2} by A10;
           then x in C1 & x in C2 by A2,A3,RLSUB_1:def 5;
        hence thesis by XBOOLE_0:def 3;
       end;
    hence thesis;
   end;

Lm17: ex C st v in C
 proof
   reconsider C = v + W as Coset of W by RLSUB_1:def 6;
  take C;
  thus thesis by RLSUB_1:59;
 end;

theorem Th51:
 for V being RealLinearSpace, W1,W2 being Subspace of V holds
 V is_the_direct_sum_of W1,W2 iff
  for C1 being Coset of W1, C2 being Coset of W2
   ex v being VECTOR of V st C1 /\ C2 = {v}
   proof let V be RealLinearSpace, W1,W2 be Subspace of V;
      set VW1 = the carrier of W1; set VW2 = the carrier of W2;
      A1: VW1 is Coset of W1 & VW2 is Coset of W2 by RLSUB_1:90;
    thus V is_the_direct_sum_of W1,W2 implies
     for C1 being Coset of W1, C2 being Coset of W2
      ex v being VECTOR of V st C1 /\ C2 = {v}
     proof assume A2: V is_the_direct_sum_of W1,W2;
      let C1 be Coset of W1, C2 be Coset of W2;
       consider v1 being VECTOR of V such that A3: C1 = v1 + W1 by RLSUB_1:def
6;
       consider v2 being VECTOR of V such that A4: C2 = v2 + W2 by RLSUB_1:def
6;
        A5: the RLSStruct of V = W1 + W2 by A2,Def4;
          v1 in the RLSStruct of V by RLVECT_1:3;
       then consider v11,v12 being VECTOR of V such that A6: v11 in W1 and
        A7: v12 in W2 and A8: v1 = v11 + v12 by A5,Th5;
          v2 in the RLSStruct of V by RLVECT_1:3;
       then consider v21,v22 being VECTOR of V such that A9: v21 in W1 and
        A10: v22 in W2 and A11: v2 = v21 + v22 by A5,Th5;
      take v = v12 + v21;
         {v} = C1 /\ C2
        proof
         thus A12: {v} c= C1 /\ C2
          proof let x;
            assume x in {v};
             then A13: x = v by TARSKI:def 1;
               v12 = v1 - v11 by A8,Lm14;
             then v12 in C1 by A3,A6,RLSUB_1:80;
             then C1 = v12 + W1 by RLSUB_1:94;
             then A14: x in C1 by A9,A13,RLSUB_1:79;
               v21 = v2 - v22 by A11,Lm14;
             then v21 in C2 by A4,A10,RLSUB_1:80;
             then C2 = v21 + W2 & v = v21 + v12 by RLSUB_1:94;
             then x in C2 by A7,A13,RLSUB_1:79;
           hence thesis by A14,XBOOLE_0:def 3;
          end;
         let x;
          assume A15: x in C1 /\ C2;
           then C1 meets C2 by XBOOLE_0:4;
           then reconsider C = C1 /\ C2 as Coset of W1 /\ W2 by Th50;
              W1 /\ W2 = (0).V by A2,Def4;
           then A16: ex u being VECTOR of V st C = {u} by RLSUB_1:89;
              v in {v} by TARSKI:def 1;
            hence thesis by A12,A15,A16,TARSKI:def 1;
         thus thesis;
        end;
       hence C1 /\ C2 = {v};
      end;
     assume
     A17: for C1 being Coset of W1, C2 being Coset of W2
         ex v being VECTOR of V st C1 /\ C2 = {v};
            now let u be VECTOR of V;
             consider C1 being Coset of W1 such that A18: u in C1 by Lm17;
             consider v being VECTOR of V such that A19: C1 /\ VW2 = {v} by A1,
A17;
              A20: v in {v} by TARSKI:def 1;
              then v in VW2 by A19,XBOOLE_0:def 3;
              then A21: v in W2 by RLVECT_1:def 1;
                v in C1 by A19,A20,XBOOLE_0:def 3;
             then consider v1 being VECTOR of V such that
             A22: v1 in W1 and A23: u - v1 = v by A18,RLSUB_1:96;
                u = v1 + v by A23,Lm14;
           hence u in W1 + W2 by A21,A22,Th5;
          end;
    hence the RLSStruct of V = W1 + W2 by Lm12;
      consider v being VECTOR of V such that A24: VW1 /\ VW2 = {v} by A1,A17;
         0.V in W1 & 0.V in W2 by RLSUB_1:25;
       then 0.V in VW1 & 0.V in VW2 by RLVECT_1:def 1;
       then A25: 0.V in {v} by A24,XBOOLE_0:def 3;
         the carrier of (0).V = {0.V} by RLSUB_1:def 3
                           .= VW1 /\ VW2 by A24,A25,TARSKI:def 1
                           .= the carrier of W1 /\ W2 by Def2;
    hence thesis by RLSUB_1:38;
   end;

::
::  Decomposition of a vector.
::

theorem
   for V being RealLinearSpace, W1,W2 being Subspace of V holds
 W1 + W2 = the RLSStruct of V iff
  for v being VECTOR of V
   ex v1,v2 being VECTOR of V st v1 in W1 & v2 in W2 & v = v1 + v2 by Lm13;

theorem Th53:
 V is_the_direct_sum_of W1,W2 &
  v = v1 + v2 & v = u1 + u2 & v1 in W1 & u1 in W1 & v2 in W2 & u2 in W2 implies
   v1 = u1 & v2 = u2
     proof assume A1: V is_the_direct_sum_of W1,W2;
       assume that A2: v = v1 + v2 & v = u1 + u2 and A3: v1 in W1 & u1 in
 W1 and
                   A4: v2 in W2 & u2 in W2;
        reconsider C2 = v1 + W2 as Coset of W2 by RLSUB_1:def 6;
        reconsider C1 = the carrier of W1 as Coset of W1 by RLSUB_1:90;
         A5: u1 = (v1 + v2) - u2 by A2,Lm14
              .= v1 + (v2 - u2) by RLVECT_1:42;
           v2 - u2 in W2 by A4,RLSUB_1:31;
         then v1 in C1 & v1 in C2 & u1 in C1 & u1 in C2
                                      by A3,A5,RLSUB_1:59,79,RLVECT_1:def 1;
         then A6: v1 in C1 /\ C2 & u1 in C1 /\ C2 by XBOOLE_0:def 3;
        consider u being VECTOR of V such that A7: C1 /\ C2 = {u} by A1,Th51;
A8:         v1 = u & u1 = u by A6,A7,TARSKI:def 1;
      hence v1 = u1;
      thus v2 = u2 by A2,A8,RLVECT_1:21;
     end;

Lm18: X c< Y implies ex x st x in Y & not x in X
 proof assume X c< Y;
    then not Y c= X by XBOOLE_1:60;
    then A1: Y \ X <> {} by XBOOLE_1:37;
    consider x being Element of Y \ X;
  take x;
  thus thesis by A1,XBOOLE_0:def 4;
 end;

theorem
   V = W1 + W2 &
  (ex v st for v1,v2,u1,u2 st
    v = v1 + v2 & v = u1 + u2 & v1 in W1 & u1 in W1 & v2 in W2 & u2 in W2 holds
     v1 = u1 & v2 = u2) implies V is_the_direct_sum_of W1,W2
  proof assume A1: V = W1 + W2;
    given v such that A2: for v1,v2,u1,u2 st v = v1 + v2 & v = u1 + u2 &
     v1 in W1 & u1 in W1 & v2 in W2 & u2 in W2 holds v1 = u1 & v2 = u2;
      A3: the carrier of (0).V = {0.V} by RLSUB_1:def 3;
    assume not thesis;
      then W1 /\ W2 <> (0).V by A1,Def4;
      then the carrier of W1 /\ W2 <> the carrier of (0).V &
           (0).V is Subspace of W1 /\ W2 by RLSUB_1:38,51;
      then the carrier of W1 /\ W2 <> {0.V} & {0.V} c= the carrier of W1 /\ W2
       by A3,RLSUB_1:def 2;
      then {0.V} c< the carrier of W1 /\ W2 by XBOOLE_0:def 8;
      then consider x such that A4: x in the carrier of W1 /\ W2 and
                               A5: not x in {0.V} by Lm18;
A6:      x in W1 /\ W2 by A4,RLVECT_1:def 1;
      then A7: x in W1 & x in W2 by Th7;
      A8: x <> 0.V by A5,TARSKI:def 1;
        x in V by A6,RLSUB_1:17;
     then reconsider u = x as VECTOR of V by RLVECT_1:def 1;
     consider v1,v2 such that A9: v1 in W1 & v2 in W2 and A10: v = v1 + v2
                                                                by A1,Lm13;
      A11: v = v1 + v2 + 0.V by A10,RLVECT_1:10
          .= (v1 + v2) + (u - u) by RLVECT_1:28
          .= ((v1 + v2) + u) - u by RLVECT_1:42
          .= ((v1 + u) + v2) - u by RLVECT_1:def 6
          .= (v1 + u) + (v2 - u) by RLVECT_1:42;
        v1 + u in W1 & v2 - u in W2 by A7,A9,RLSUB_1:28,31;
      then v2 - u = v2 by A2,A9,A10,A11
                 .= v2 - 0.V by RLVECT_1:26;
   hence thesis by A8,RLVECT_1:37;
  end;

reserve t1,t2 for Element of [:the carrier of V, the carrier of V:];



definition let V; let v; let W1,W2;
 assume A1: V is_the_direct_sum_of W1,W2;
 func v |-- (W1,W2) -> Element of [:the carrier of V, the carrier of V:] means
  :Def6: v = it`1 + it`2 & it`1 in W1 & it`2 in W2;
 existence
  proof W1 + W2 = the RLSStruct of V by A1,Def4;
    then consider v1,v2 such that A2: v1 in W1 & v2 in
 W2 & v = v1 + v2 by Lm13;
   take [v1,v2];
      [v1,v2]`1 = v1 & [v1,v2]`2 = v2 by MCART_1:7;
   hence thesis by A2;
  end;
 uniqueness
  proof let t1,t2;
    assume v = t1`1 + t1`2 & t1`1 in W1 & t1`2 in W2 &
           v = t2`1 + t2`2 & t2`1 in W1 & t2`2 in W2;
     then t1`1 = t2`1 & t1`2 = t2`2 & t1 = [t1`1,t1`2] & t2 = [t2`1,t2`2]
                                                    by A1,Th53,MCART_1:23;
   hence thesis;
  end;
end;

canceled 4;

theorem Th59:
 V is_the_direct_sum_of W1,W2 implies
  (v |-- (W1,W2))`1 = (v |-- (W2,W1))`2
   proof assume A1: V is_the_direct_sum_of W1,W2;
     then A2: v = (v |-- (W1,W2))`1 + (v |-- (W1,W2))`2 & (v |-- (W1,W2))`1 in
 W1 &
             (v |-- (W1,W2))`2 in W2 by Def6;
     A3: V is_the_direct_sum_of W2,W1 by A1,Lm16;
     then A4: v = (v |-- (W2,W1))`2 + (v |-- (W2,W1))`1 by Def6;
       (v |-- (W2,W1))`1 in W2 & (v |-- (W2,W1))`2 in W1 by A3,Def6;
    hence thesis by A1,A2,A4,Th53;
   end;

theorem Th60:
 V is_the_direct_sum_of W1,W2 implies
  (v |-- (W1,W2))`2 = (v |-- (W2,W1))`1
   proof assume A1: V is_the_direct_sum_of W1,W2;
     then A2: v = (v |-- (W1,W2))`1 + (v |-- (W1,W2))`2 & (v |-- (W1,W2))`1 in
 W1 &
             (v |-- (W1,W2))`2 in W2 by Def6;
     A3: V is_the_direct_sum_of W2,W1 by A1,Lm16;
     then A4: v = (v |-- (W2,W1))`2 + (v |-- (W2,W1))`1 by Def6;
       (v |-- (W2,W1))`1 in W2 & (v |-- (W2,W1))`2 in W1 by A3,Def6;
    hence thesis by A1,A2,A4,Th53;
   end;

theorem
   for V being RealLinearSpace, W being Subspace of V,
   L being Linear_Compl of W, v being VECTOR of V,
   t being Element of [:the carrier of V, the carrier of V:] holds
 t`1 + t`2 = v & t`1 in W & t`2 in L implies t = v |-- (W,L)
  proof let V be RealLinearSpace, W be Subspace of V,
            L be Linear_Compl of W;
     V is_the_direct_sum_of W,L by Th43;
   hence thesis by Def6;
  end;

theorem
   for V being RealLinearSpace, W being Subspace of V,
   L being Linear_Compl of W, v being VECTOR of V holds
 (v |-- (W,L))`1 + (v |-- (W,L))`2 = v
  proof let V be RealLinearSpace, W be Subspace of V,
            L be Linear_Compl of W;
     V is_the_direct_sum_of W,L by Th43;
   hence thesis by Def6;
  end;

theorem
   for V being RealLinearSpace, W being Subspace of V,
   L being Linear_Compl of W, v being VECTOR of V holds
 (v |-- (W,L))`1 in W & (v |-- (W,L))`2 in L
  proof let V be RealLinearSpace, W be Subspace of V,
            L be Linear_Compl of W;
     V is_the_direct_sum_of W,L by Th43;
   hence thesis by Def6;
  end;

theorem
   for V being RealLinearSpace, W being Subspace of V,
   L being Linear_Compl of W, v being VECTOR of V holds
 (v |-- (W,L))`1 = (v |-- (L,W))`2
  proof let V be RealLinearSpace, W be Subspace of V,
            L be Linear_Compl of W;
     V is_the_direct_sum_of W,L by Th43;
   hence thesis by Th59;
  end;

theorem
   for V being RealLinearSpace, W being Subspace of V,
   L being Linear_Compl of W, v being VECTOR of V holds
 (v |-- (W,L))`2 = (v |-- (L,W))`1
  proof let V be RealLinearSpace, W be Subspace of V,
            L be Linear_Compl of W;
     V is_the_direct_sum_of W,L by Th43;
   hence thesis by Th60;
  end;

::
::  Introduction of operations on set of subspaces as binary operations.
::

reserve A1,A2,B for Element of Subspaces(V);

 definition let V;
  func SubJoin(V) -> BinOp of Subspaces(V) means
   :Def7: for A1,A2,W1,W2 st A1 = W1 & A2 = W2 holds it.(A1,A2) = W1 + W2;
  existence
   proof
     defpred P[Element of Subspaces(V),Element of Subspaces(V),
     Element of Subspaces(V)] means
      for W1,W2 st $1 = W1 & $2 = W2 holds $3 = W1 + W2;
     A1: for A1,A2 ex B st P[A1,A2,B]
      proof let A1,A2;
        reconsider W1 = A1, W2 = A2 as Subspace of V by Def3;
        reconsider C = W1 + W2 as Element of Subspaces(V) by Def3;
       take C;
       thus thesis;
      end;
    ex o being BinOp of Subspaces(V) st
    for a,b being Element of Subspaces(V) holds P[a,b,o.(a,b)]
                  from BinOpEx(A1);
    hence thesis;
   end;
  uniqueness
   proof let o1,o2 be BinOp of Subspaces(V);
     assume A2: for A1,A2,W1,W2 st A1 = W1 & A2 = W2 holds o1.(A1,A2) = W1 + W2
;
     assume A3: for A1,A2,W1,W2 st A1 = W1 & A2 = W2 holds o2.(A1,A2) = W1 + W2
;
        now let x,y be set;
        assume A4: x in Subspaces(V) & y in Subspaces(V);
         then reconsider A = x, B = y as Element of Subspaces(V);
         reconsider W1 = x, W2 = y as Subspace of V by A4,Def3;
            o1.(A,B) = W1 + W2 & o2.(A,B) = W1 + W2 by A2,A3;
          then o1.(x,y) = o2.(x,y) & o1.[x,y] = o1.(x,y) & o2.[x,y] = o2.(x,y)
                                                                  by BINOP_1:
def 1;
       hence o1.[x,y] = o2.[x,y];
      end;
    hence thesis by FUNCT_2:118;
   end;
 end;

 definition let V;
  func SubMeet(V) -> BinOp of Subspaces(V) means
   :Def8: for A1,A2,W1,W2 st A1 = W1 & A2 = W2 holds it.(A1,A2) = W1 /\ W2;
  existence
   proof
     defpred P[Element of Subspaces(V),Element of Subspaces(V),
     Element of Subspaces(V)] means
      for W1,W2 st $1 = W1 & $2 = W2 holds $3 = W1 /\ W2;
     A1: for A1,A2 ex B st P[A1,A2,B]
      proof let A1,A2;
        reconsider W1 = A1, W2 = A2 as Subspace of V by Def3;
        reconsider C = W1 /\ W2 as Element of Subspaces(V) by Def3;
       take C;
       thus thesis;
      end;
    ex o being BinOp of Subspaces(V) st
    for a,b being Element of Subspaces(V) holds P[a,b,o.(a,b)]
                  from BinOpEx(A1);
    hence thesis;
   end;
  uniqueness
   proof let o1,o2 be BinOp of Subspaces(V);
     assume A2: for A1,A2,W1,W2 st A1 = W1 & A2 = W2 holds o1.(A1,A2) = W1 /\
W2;
     assume A3: for A1,A2,W1,W2 st A1 = W1 & A2 = W2 holds o2.(A1,A2) = W1 /\
W2;
        now let x,y be set;
        assume A4: x in Subspaces(V) & y in Subspaces(V);
         then reconsider A = x, B = y as Element of Subspaces(V);
         reconsider W1 = x, W2 = y as Subspace of V by A4,Def3;
            o1.(A,B) = W1 /\ W2 & o2.(A,B) = W1 /\ W2 by A2,A3;
          then o1.(x,y) = o2.(x,y) & o1.[x,y] = o1.(x,y) & o2.[x,y] = o2.(x,y)
                                                                  by BINOP_1:
def 1;
       hence o1.[x,y] = o2.[x,y];
      end;
    hence thesis by FUNCT_2:118;
   end;
 end;

::
::  Definitional theorems of functions SubJoin, SubMeet.
::



definition let X be non empty set, m,u be BinOp of X;
 cluster LattStr(#X,m,u#) -> non empty;
 coherence
  proof
   thus the carrier of LattStr(#X,m,u#) is non empty;
  end;
end;

canceled 4;

theorem Th70:
 LattStr (# Subspaces(V), SubJoin(V), SubMeet(V) #) is Lattice
  proof set S = LattStr (# Subspaces(V), SubJoin(V), SubMeet(V) #);
    A1: for A,B being Element of S
         holds A "\/" B = B "\/" A
     proof let A,B be Element of S;
       reconsider W1 = A, W2 = B as Subspace of V by Def3;
      thus A "\/" B = SubJoin(V).(A,B) by LATTICES:def 1
                 .= W1 + W2 by Def7
                 .= W2 + W1 by Lm1
                 .= SubJoin(V).(B,A) by Def7
                 .= B "\/" A by LATTICES:def 1;
     end;
    A2: for A,B,C being Element of S
         holds A "\/" (B "\/" C) = (A "\/" B) "\/" C
     proof let A,B,C be Element of S;
       reconsider W1 = A, W2 = B, W3 = C as Subspace of V by Def3;
       reconsider AB = W1 + W2, BC = W2 + W3 as
              Element of S by Def3;
      thus A "\/" (B "\/" C) = SubJoin(V).(A,B "\/" C) by LATTICES:def 1
                        .= SubJoin(V).(A,SubJoin(V).(B,C)) by LATTICES:def 1
                        .= SubJoin(V).(A,BC) by Def7
                        .= W1 + (W2 + W3) by Def7
                        .= (W1 + W2) + W3 by Th10
                        .= SubJoin(V).(AB,C) by Def7
                        .= SubJoin(V).(SubJoin(V).(A,B),C) by Def7
                        .= SubJoin(V).(A "\/" B,C) by LATTICES:def 1
                        .= (A "\/" B) "\/" C by LATTICES:def 1;
     end;
    A3: for A,B being Element of S
         holds (A "/\" B) "\/" B = B
     proof let A,B be Element of S;
       reconsider W1 = A, W2 = B as strict Subspace of V by Def3;
       reconsider AB = W1 /\ W2 as Element of S by Def3;
      thus (A "/\" B) "\/" B = SubJoin(V).(A "/\" B,B) by LATTICES:def 1
                        .= SubJoin(V).(SubMeet(V).(A,B),B) by LATTICES:def 2
                        .= SubJoin(V).(AB,B) by Def8
                        .= (W1 /\ W2) + W2 by Def7
                        .= B by Th27;
     end;
    A4: for A,B being Element of S
         holds A "/\" B = B "/\" A
     proof let A,B be Element of S;
       reconsider W1 = A, W2 = B as Subspace of V by Def3;
      thus A "/\" B = SubMeet(V).(A,B) by LATTICES:def 2
                 .= W1 /\ W2 by Def8
                 .= W2 /\ W1 by Th18
                 .= SubMeet(V).(B,A) by Def8
                 .= B "/\" A by LATTICES:def 2;
     end;
    A5: for A,B,C being Element of S
         holds A "/\" (B "/\" C) = (A "/\" B) "/\" C
     proof let A,B,C be Element of S;
       reconsider W1 = A, W2 = B, W3 = C as Subspace of V by Def3;
       reconsider AB = W1 /\ W2, BC = W2 /\ W3 as
              Element of S by Def3;
      thus A "/\" (B "/\" C) = SubMeet(V).(A,B "/\" C) by LATTICES:def 2
                        .= SubMeet(V).(A,SubMeet(V).(B,C)) by LATTICES:def 2
                        .= SubMeet(V).(A,BC) by Def8
                        .= W1 /\ (W2 /\ W3) by Def8
                        .= (W1 /\ W2) /\ W3 by Th19
                        .= SubMeet(V).(AB,C) by Def8
                        .= SubMeet(V).(SubMeet(V).(A,B),C) by Def8
                        .= SubMeet(V).(A "/\" B,C) by LATTICES:def 2
                        .= (A "/\" B) "/\" C by LATTICES:def 2;
     end;
      for A,B being Element of S
         holds A "/\" (A "\/" B) = A
     proof let A,B be Element of S;
       reconsider W1 = A, W2 = B as strict Subspace of V by Def3;
       reconsider AB = W1 + W2 as Element of S by Def3;
      thus A "/\" (A "\/" B) = SubMeet(V).(A,A "\/" B) by LATTICES:def 2
                        .= SubMeet(V).(A,SubJoin(V).(A,B)) by LATTICES:def 1
                        .= SubMeet(V).(A,AB) by Def7
                        .= W1 /\ (W1 + W2) by Def8
                        .= A by Th28;
     end;
  then S is join-commutative join-associative meet-absorbing
        meet-commutative meet-associative join-absorbing
   by A1,A2,A3,A4,A5,LATTICES:def 4,def 5,def 6,def 7,def 8,def 9;
   hence thesis by LATTICES:def 10;
  end;

definition let V;
  cluster LattStr (# Subspaces(V), SubJoin(V), SubMeet(V) #) -> Lattice-like;
  coherence by Th70;
end;

theorem Th71:
 for V being RealLinearSpace holds
 LattStr (# Subspaces(V), SubJoin(V), SubMeet(V) #) is lower-bounded
  proof let V be RealLinearSpace;
   set S = LattStr (# Subspaces(V), SubJoin(V), SubMeet(V) #);
      ex C being Element of S st
     for A being Element of S holds C "/\" A = C & A "/\" C = C
      proof
        reconsider C = (0).V as Element of S by Def3;
       take C;
       let A be Element of S;
        reconsider W = A as Subspace of V by Def3;
       thus C "/\" A = SubMeet(V).(C,A) by LATTICES:def 2
                  .= (0).V /\ W by Def8
                  .= C by Th22;
       hence A "/\" C = C;
      end;
   hence thesis by LATTICES:def 13;
  end;

theorem Th72:
 for V being RealLinearSpace holds
 LattStr (# Subspaces(V), SubJoin(V), SubMeet(V) #) is upper-bounded
  proof let V be RealLinearSpace;
   set S = LattStr (# Subspaces(V), SubJoin(V), SubMeet(V) #);
      ex C being Element of S st
     for A being Element of S holds C "\/" A = C & A "\/" C = C
      proof
        reconsider C = (Omega).V as Element of S by Def3;
       take C;
       let A be Element of S;
        reconsider W = A as Subspace of V by Def3;
       thus C "\/" A = SubJoin(V).(C,A) by LATTICES:def 1
                  .= (Omega).V + W by Def7
                  .= the RLSStruct of V by Th15
                  .= C by RLSUB_1:def 4;
       hence A "\/" C = C;
      end;
   hence thesis by LATTICES:def 14;
  end;

theorem Th73:
 for V being RealLinearSpace holds
 LattStr (# Subspaces(V), SubJoin(V), SubMeet(V) #) is 01_Lattice
  proof let V be RealLinearSpace;
      LattStr (# Subspaces(V), SubJoin(V), SubMeet(V) #)
    is lower-bounded upper-bounded Lattice by Th71,Th72;
   hence thesis;
  end;

theorem Th74:
 for V being RealLinearSpace holds
 LattStr (# Subspaces(V), SubJoin(V), SubMeet(V) #) is modular
  proof let V be RealLinearSpace;
    set S = LattStr (# Subspaces(V), SubJoin(V), SubMeet(V) #);
      for A,B,C being Element of S st A [= C
     holds A "\/" (B "/\" C) = (A "\/" B) "/\" C
      proof let A,B,C be Element of S;
        assume A1: A [= C;
         reconsider W1 = A, W2 = B, W3 = C as strict Subspace of V by Def3;
         reconsider BC = W2 /\ W3 as Element of S by Def3;
         reconsider AB = W1 + W2 as Element of S by Def3;
            W1 + W3 = SubJoin(V).(A,C) by Def7
                      .= A "\/" C by LATTICES:def 1
                      .= W3 by A1,LATTICES:def 3;
          then A2: W1 is Subspace of W3 by Th12;
       thus A "\/" (B "/\" C) = SubJoin(V).(A,B "/\" C) by LATTICES:def 1
                         .= SubJoin(V).(A,SubMeet(V).(B,C)) by LATTICES:def 2
                         .= SubJoin(V).(A,BC) by Def8
                         .= W1 + (W2 /\ W3) by Def7
                         .= (W1 + W2) /\ W3 by A2,Th33
                         .= SubMeet(V).(AB,C) by Def8
                         .= SubMeet(V).(SubJoin(V).(A,B),C) by Def7
                         .= SubMeet(V).(A "\/" B,C) by LATTICES:def 1
                         .= (A "\/" B) "/\" C by LATTICES:def 2;
      end;
   hence thesis by LATTICES:def 12;
  end;

reserve l for Lattice;
reserve a,b for Element of l;

Lm19: a is_a_complement_of b iff a "\/" b = Top l & a "/\" b = Bottom l
 proof
    not (a "\/" b = Top l & a "/\" b = Bottom
 l) implies not a is_a_complement_of b
   proof assume not (a "\/" b = Top l & a "/\" b = Bottom l);
    hence not (a "\/" b = Top l & b "\/" a = Top l & a "/\" b = Bottom
 l & b "/\"
 a = Bottom l);
   end;
  hence a is_a_complement_of b implies a "\/" b = Top l & a "/\" b = Bottom
 l;
   assume a "\/" b = Top l & a "/\" b = Bottom l;
  hence a "\/" b = Top l & b "\/" a = Top
 l & a "/\" b = Bottom l & b "/\" a = Bottom l;
 end;

Lm20: (for a holds a "/\" b = b) implies b = Bottom l
 proof assume A1: for a holds a "/\" b = b;
then for a holds a "/\" b = b & b "/\" a = b;
then l is lower-bounded by LATTICES:def 13;
  hence Bottom l = Bottom l "/\" b by LATTICES:def 16
           .= b by A1;
 end;

Lm21: (for a holds a "\/" b = b) implies b = Top l
 proof assume A1: for a holds a "\/" b = b;
then for a holds b "\/" a = b & a "\/" b = b;
then l is upper-bounded by LATTICES:def 14;
  hence Top l = Top l "\/" b by LATTICES:44
            .= b by A1;
 end;

theorem Th75:
 for V being RealLinearSpace holds
 LattStr (# Subspaces(V), SubJoin(V), SubMeet(V) #) is complemented
  proof let V be RealLinearSpace;
    reconsider S = LattStr (# Subspaces(V), SubJoin(V), SubMeet(V) #)
 as 01_Lattice by Th73;
    reconsider S0 = S as 0_Lattice;
    reconsider S1 = S as 1_Lattice;
    reconsider Z = (0).V, I = (Omega).V as Element of S by Def3;
    reconsider Z0 = Z as Element of S0;
    reconsider I1 = I as Element of S1;
       now let A be Element of S1;
       reconsider W = A as Subspace of V by Def3;
      thus A "\/" I1 = SubJoin(V).(A,I1) by LATTICES:def 1
                  .= W + (Omega).V by Def7
                  .= the RLSStruct of V by Th15
                  .= (Omega).V by RLSUB_1:def 4;
     end;
     then A1: Top S = I by Lm21;
       now let A be Element of S0;
       reconsider W = A as Subspace of V by Def3;
      thus A "/\" Z0 = SubMeet(V).(A,Z0) by LATTICES:def 2
                  .= W /\ (0).V by Def8
                  .= Z0 by Th22;
     end;
     then A2: Bottom S = Z by Lm20;
       now let A be Element of S;
       reconsider W = A as Subspace of V by Def3;
       consider L being strict Linear_Compl of W;
       reconsider B' = L as Element of S by Def3;
      take B = B';
        A3: B "\/" A = SubJoin(V).(B,A) by LATTICES:def 1
                 .= L + W by Def7
                 .= the RLSStruct of V by Th44
                 .= Top S by A1,RLSUB_1:def 4;
          B "/\" A = SubMeet(V).(B,A) by LATTICES:def 2
              .= L /\ W by Def8
              .= Bottom S by A2,Th45;
      hence B is_a_complement_of A by A3,Lm19;
     end;
   hence thesis by LATTICES:def 19;
  end;

definition let V;
  cluster LattStr (# Subspaces(V), SubJoin(V), SubMeet(V) #) ->
    lower-bounded upper-bounded modular complemented;
  coherence by Th71,Th72,Th74,Th75;
end;

::
::  Theorems concerning operations on subspaces (continuation). Proven
::  on the basis that set of subspaces with operations is a lattice.
::

theorem
   for V being RealLinearSpace,
     W1,W2,W3 being strict Subspace of V holds
 W1 is Subspace of W2 implies W1 /\ W3 is Subspace of W2 /\ W3
  proof let V be RealLinearSpace,
            W1,W2,W3 be strict Subspace of V;
   set S = LattStr (# Subspaces(V), SubJoin(V), SubMeet(V) #);
     reconsider A = W1, B = W2, C = W3, AC = W1 /\ W3, BC = W2 /\ W3
        as Element of S by Def3;
    assume A1: W1 is Subspace of W2;
         A "\/" B = SubJoin(V).(A,B) by LATTICES:def 1
             .= W1 + W2 by Def7
             .= B by A1,Th12;
       then A [= B by LATTICES:def 3;
       then A "/\" C [= B "/\" C by LATTICES:27;
       then A2: (A "/\" C) "\/" (B "/\" C) = (B "/\" C) by LATTICES:def 3;
       A3: B "/\" C = SubMeet(V).(B,C) by LATTICES:def 2
             .= W2 /\ W3 by Def8;
         (A "/\" C) "\/" (B "/\" C) = SubJoin(V).(A "/\" C,B "/\"
 C) by LATTICES:def 1
           .= SubJoin(V).(SubMeet(V).(A,C),B "/\" C) by LATTICES:def 2
           .= SubJoin(V).(SubMeet(V).(A,C),SubMeet(V).(B,C)) by LATTICES:def 2
           .= SubJoin(V).(SubMeet(V).(A,C),BC) by Def8
           .= SubJoin(V).(AC,BC) by Def8
           .= (W1 /\ W3) + (W2 /\ W3) by Def7;
   hence thesis by A2,A3,Th12;
  end;

::
::  Auxiliary theorems.
::

theorem
   X c< Y implies ex x st x in Y & not x in X by Lm18;

theorem
  for V being add-associative right_zeroed right_complementable
 (non empty LoopStr),
    v,v1,v2 being Element of V holds
 v = v1 + v2 iff v1 = v - v2 by Lm14;

theorem
   for V being RealLinearSpace, W being strict Subspace of V holds
 (for v being VECTOR of V holds v in W) implies W = the RLSStruct of V by Lm12;

theorem
   ex C st v in C by Lm17;

canceled 3;

theorem
   (for a holds a "/\" b = b) implies b = Bottom l by Lm20;

theorem
   (for a holds a "\/" b = b) implies b = Top l by Lm21;

Back to top