:: The Fundamental Logic Structure in Quantum Mechanics
:: by Pawe{\l} Sadowski, Andrzej Trybulec and Konrad Raczkowski
::
:: Received December 18, 1989
:: Copyright (c) 1990-2021 Association of Mizar Users


definition
let X be non empty set ;
let S be SigmaField of X;
func Probabilities S -> set means :Def1: :: QMAX_1:def 1
for x being object holds
( x in it iff x is Probability of S );
existence
ex b1 being set st
for x being object holds
( x in b1 iff x is Probability of S )
proof end;
uniqueness
for b1, b2 being set st ( for x being object holds
( x in b1 iff x is Probability of S ) ) & ( for x being object holds
( x in b2 iff x is Probability of S ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def1 defines Probabilities QMAX_1:def 1 :
for X being non empty set
for S being SigmaField of X
for b3 being set holds
( b3 = Probabilities S iff for x being object holds
( x in b3 iff x is Probability of S ) );

registration
let X be non empty set ;
let S be SigmaField of X;
cluster Probabilities S -> non empty ;
coherence
not Probabilities S is empty
proof end;
end;

definition
attr c1 is strict ;
struct QM_Str -> ;
aggr QM_Str(# Observables, FStates, Quantum_Probability #) -> QM_Str ;
sel Observables c1 -> non empty set ;
sel FStates c1 -> non empty set ;
sel Quantum_Probability c1 -> Function of [: the Observables of c1, the FStates of c1:],(Probabilities Borel_Sets);
end;

definition
let Q be QM_Str ;
func Obs Q -> set equals :: QMAX_1:def 2
the Observables of Q;
coherence
the Observables of Q is set
;
func Sts Q -> set equals :: QMAX_1:def 3
the FStates of Q;
coherence
the FStates of Q is set
;
end;

:: deftheorem defines Obs QMAX_1:def 2 :
for Q being QM_Str holds Obs Q = the Observables of Q;

:: deftheorem defines Sts QMAX_1:def 3 :
for Q being QM_Str holds Sts Q = the FStates of Q;

registration
let Q be QM_Str ;
cluster Obs Q -> non empty ;
coherence
not Obs Q is empty
;
cluster Sts Q -> non empty ;
coherence
not Sts Q is empty
;
end;

definition
let Q be QM_Str ;
let A1 be Element of Obs Q;
let s be Element of Sts Q;
func Meas (A1,s) -> Probability of Borel_Sets equals :: QMAX_1:def 4
the Quantum_Probability of Q . [A1,s];
coherence
the Quantum_Probability of Q . [A1,s] is Probability of Borel_Sets
proof end;
end;

:: deftheorem defines Meas QMAX_1:def 4 :
for Q being QM_Str
for A1 being Element of Obs Q
for s being Element of Sts Q holds Meas (A1,s) = the Quantum_Probability of Q . [A1,s];

set X = {0};

consider P being Function of Borel_Sets,REAL such that
Lm1: for D being Subset of REAL st D in Borel_Sets holds
( ( 0 in D implies P . D = 1 ) & ( not 0 in D implies P . D = 0 ) )
by PROB_1:28;

Lm2: for A being Event of Borel_Sets holds 0 <= P . A
proof end;

Lm3: P . REAL = 1
proof end;

Lm4: for A, B being Event of Borel_Sets st A misses B holds
P . (A \/ B) = (P . A) + (P . B)

proof end;

for ASeq being SetSequence of Borel_Sets st ASeq is non-ascending holds
( P * ASeq is convergent & lim (P * ASeq) = P . (Intersection ASeq) )

proof end;

then reconsider P = P as Probability of Borel_Sets by Lm2, Lm3, Lm4, PROB_1:def 8;

set f = {[[0,0],P]};

dom {[[0,0],P]} = {[0,0]}
by RELAT_1:9;

then Lm5: dom {[[0,0],P]} = [:{0},{0}:]
by ZFMISC_1:29;

( rng {[[0,0],P]} = {P} & P in Probabilities Borel_Sets )
by Def1, RELAT_1:9;

then rng {[[0,0],P]} c= Probabilities Borel_Sets
by ZFMISC_1:31;

then reconsider Y = {[[0,0],P]} as Function of [:{0},{0}:],(Probabilities Borel_Sets) by Lm5, FUNCT_2:def 1, RELSET_1:4;

Lm6: now :: thesis: ( ( for A1, A2 being Element of Obs QM_Str(# {0},{0},Y #) st ( for s being Element of Sts QM_Str(# {0},{0},Y #) holds Meas (A1,s) = Meas (A2,s) ) holds
A1 = A2 ) & ( for s1, s2 being Element of Sts QM_Str(# {0},{0},Y #) st ( for A being Element of Obs QM_Str(# {0},{0},Y #) holds Meas (A,s1) = Meas (A,s2) ) holds
s1 = s2 ) & ( for s1, s2 being Element of Sts QM_Str(# {0},{0},Y #)
for t being Real st 0 <= t & t <= 1 holds
ex s being Element of Sts QM_Str(# {0},{0},Y #) st
for A being Element of Obs QM_Str(# {0},{0},Y #)
for E being Event of Borel_Sets holds (Meas (A,s)) . E = (t * ((Meas (A,s1)) . E)) + ((1 - t) * ((Meas (A,s2)) . E)) ) )
thus for A1, A2 being Element of Obs QM_Str(# {0},{0},Y #) st ( for s being Element of Sts QM_Str(# {0},{0},Y #) holds Meas (A1,s) = Meas (A2,s) ) holds
A1 = A2 :: thesis: ( ( for s1, s2 being Element of Sts QM_Str(# {0},{0},Y #) st ( for A being Element of Obs QM_Str(# {0},{0},Y #) holds Meas (A,s1) = Meas (A,s2) ) holds
s1 = s2 ) & ( for s1, s2 being Element of Sts QM_Str(# {0},{0},Y #)
for t being Real st 0 <= t & t <= 1 holds
ex s being Element of Sts QM_Str(# {0},{0},Y #) st
for A being Element of Obs QM_Str(# {0},{0},Y #)
for E being Event of Borel_Sets holds (Meas (A,s)) . E = (t * ((Meas (A,s1)) . E)) + ((1 - t) * ((Meas (A,s2)) . E)) ) )
proof
let A1, A2 be Element of Obs QM_Str(# {0},{0},Y #); :: thesis: ( ( for s being Element of Sts QM_Str(# {0},{0},Y #) holds Meas (A1,s) = Meas (A2,s) ) implies A1 = A2 )
A1 = 0 by TARSKI:def 1;
hence ( ( for s being Element of Sts QM_Str(# {0},{0},Y #) holds Meas (A1,s) = Meas (A2,s) ) implies A1 = A2 ) by TARSKI:def 1; :: thesis: verum
end;
thus for s1, s2 being Element of Sts QM_Str(# {0},{0},Y #) st ( for A being Element of Obs QM_Str(# {0},{0},Y #) holds Meas (A,s1) = Meas (A,s2) ) holds
s1 = s2 :: thesis: for s1, s2 being Element of Sts QM_Str(# {0},{0},Y #)
for t being Real st 0 <= t & t <= 1 holds
ex s being Element of Sts QM_Str(# {0},{0},Y #) st
for A being Element of Obs QM_Str(# {0},{0},Y #)
for E being Event of Borel_Sets holds (Meas (A,s)) . E = (t * ((Meas (A,s1)) . E)) + ((1 - t) * ((Meas (A,s2)) . E))
proof
let s1, s2 be Element of Sts QM_Str(# {0},{0},Y #); :: thesis: ( ( for A being Element of Obs QM_Str(# {0},{0},Y #) holds Meas (A,s1) = Meas (A,s2) ) implies s1 = s2 )
s1 = 0 by TARSKI:def 1;
hence ( ( for A being Element of Obs QM_Str(# {0},{0},Y #) holds Meas (A,s1) = Meas (A,s2) ) implies s1 = s2 ) by TARSKI:def 1; :: thesis: verum
end;
thus for s1, s2 being Element of Sts QM_Str(# {0},{0},Y #)
for t being Real st 0 <= t & t <= 1 holds
ex s being Element of Sts QM_Str(# {0},{0},Y #) st
for A being Element of Obs QM_Str(# {0},{0},Y #)
for E being Event of Borel_Sets holds (Meas (A,s)) . E = (t * ((Meas (A,s1)) . E)) + ((1 - t) * ((Meas (A,s2)) . E)) :: thesis: verum
proof
let s1, s2 be Element of Sts QM_Str(# {0},{0},Y #); :: thesis: for t being Real st 0 <= t & t <= 1 holds
ex s being Element of Sts QM_Str(# {0},{0},Y #) st
for A being Element of Obs QM_Str(# {0},{0},Y #)
for E being Event of Borel_Sets holds (Meas (A,s)) . E = (t * ((Meas (A,s1)) . E)) + ((1 - t) * ((Meas (A,s2)) . E))

let t be Real; :: thesis: ( 0 <= t & t <= 1 implies ex s being Element of Sts QM_Str(# {0},{0},Y #) st
for A being Element of Obs QM_Str(# {0},{0},Y #)
for E being Event of Borel_Sets holds (Meas (A,s)) . E = (t * ((Meas (A,s1)) . E)) + ((1 - t) * ((Meas (A,s2)) . E)) )

assume ( 0 <= t & t <= 1 ) ; :: thesis: ex s being Element of Sts QM_Str(# {0},{0},Y #) st
for A being Element of Obs QM_Str(# {0},{0},Y #)
for E being Event of Borel_Sets holds (Meas (A,s)) . E = (t * ((Meas (A,s1)) . E)) + ((1 - t) * ((Meas (A,s2)) . E))

take s2 ; :: thesis: for A being Element of Obs QM_Str(# {0},{0},Y #)
for E being Event of Borel_Sets holds (Meas (A,s2)) . E = (t * ((Meas (A,s1)) . E)) + ((1 - t) * ((Meas (A,s2)) . E))

let A be Element of Obs QM_Str(# {0},{0},Y #); :: thesis: for E being Event of Borel_Sets holds (Meas (A,s2)) . E = (t * ((Meas (A,s1)) . E)) + ((1 - t) * ((Meas (A,s2)) . E))
let E be Event of Borel_Sets ; :: thesis: (Meas (A,s2)) . E = (t * ((Meas (A,s1)) . E)) + ((1 - t) * ((Meas (A,s2)) . E))
( s1 = 0 & s2 = 0 ) by TARSKI:def 1;
hence (Meas (A,s2)) . E = (t * ((Meas (A,s1)) . E)) + ((1 - t) * ((Meas (A,s2)) . E)) ; :: thesis: verum
end;
end;

definition
let IT be QM_Str ;
attr IT is Quantum_Mechanics-like means :Def5: :: QMAX_1:def 5
( ( for A1, A2 being Element of Obs IT st ( for s being Element of Sts IT holds Meas (A1,s) = Meas (A2,s) ) holds
A1 = A2 ) & ( for s1, s2 being Element of Sts IT st ( for A being Element of Obs IT holds Meas (A,s1) = Meas (A,s2) ) holds
s1 = s2 ) & ( for s1, s2 being Element of Sts IT
for t being Real st 0 <= t & t <= 1 holds
ex s being Element of Sts IT st
for A being Element of Obs IT
for E being Event of Borel_Sets holds (Meas (A,s)) . E = (t * ((Meas (A,s1)) . E)) + ((1 - t) * ((Meas (A,s2)) . E)) ) );
end;

:: deftheorem Def5 defines Quantum_Mechanics-like QMAX_1:def 5 :
for IT being QM_Str holds
( IT is Quantum_Mechanics-like iff ( ( for A1, A2 being Element of Obs IT st ( for s being Element of Sts IT holds Meas (A1,s) = Meas (A2,s) ) holds
A1 = A2 ) & ( for s1, s2 being Element of Sts IT st ( for A being Element of Obs IT holds Meas (A,s1) = Meas (A,s2) ) holds
s1 = s2 ) & ( for s1, s2 being Element of Sts IT
for t being Real st 0 <= t & t <= 1 holds
ex s being Element of Sts IT st
for A being Element of Obs IT
for E being Event of Borel_Sets holds (Meas (A,s)) . E = (t * ((Meas (A,s1)) . E)) + ((1 - t) * ((Meas (A,s2)) . E)) ) ) );

registration
cluster strict Quantum_Mechanics-like for QM_Str ;
existence
ex b1 being QM_Str st
( b1 is strict & b1 is Quantum_Mechanics-like )
by Def5, Lm6;
end;

definition
mode Quantum_Mechanics is Quantum_Mechanics-like QM_Str ;
end;

definition
attr c1 is strict ;
struct OrthoRelStr -> RelStr , ComplStr ;
aggr OrthoRelStr(# carrier, InternalRel, Compl #) -> OrthoRelStr ;
end;

definition
let X1 be set ;
let Inv be Function of X1,X1;
pred Inv is_an_involution means :: QMAX_1:def 6
for x1 being Element of X1 holds Inv . (Inv . x1) = x1;
end;

:: deftheorem defines is_an_involution QMAX_1:def 6 :
for X1 being set
for Inv being Function of X1,X1 holds
( Inv is_an_involution iff for x1 being Element of X1 holds Inv . (Inv . x1) = x1 );

definition
let W be OrthoRelStr ;
pred W is_a_Quantum_Logic means :: QMAX_1:def 7
( the InternalRel of W partially_orders the carrier of W & the Compl of W is_an_involution & ( for x, y being Element of W st [x,y] in the InternalRel of W holds
[( the Compl of W . y),( the Compl of W . x)] in the InternalRel of W ) );
end;

:: deftheorem defines is_a_Quantum_Logic QMAX_1:def 7 :
for W being OrthoRelStr holds
( W is_a_Quantum_Logic iff ( the InternalRel of W partially_orders the carrier of W & the Compl of W is_an_involution & ( for x, y being Element of W st [x,y] in the InternalRel of W holds
[( the Compl of W . y),( the Compl of W . x)] in the InternalRel of W ) ) );

definition
let Q be Quantum_Mechanics;
func Prop Q -> set equals :: QMAX_1:def 8
[:(Obs Q),Borel_Sets:];
coherence
[:(Obs Q),Borel_Sets:] is set
;
end;

:: deftheorem defines Prop QMAX_1:def 8 :
for Q being Quantum_Mechanics holds Prop Q = [:(Obs Q),Borel_Sets:];

registration
let Q be Quantum_Mechanics;
cluster Prop Q -> non empty ;
coherence
not Prop Q is empty
;
end;

definition
let Q be Quantum_Mechanics;
let p be Element of Prop Q;
:: original: `1
redefine func p `1 -> Element of Obs Q;
coherence
p `1 is Element of Obs Q
by MCART_1:10;
:: original: `2
redefine func p `2 -> Event of Borel_Sets ;
coherence
p `2 is Event of Borel_Sets
by MCART_1:10;
end;

theorem Th1: :: QMAX_1:1
for Q being Quantum_Mechanics
for s being Element of Sts Q
for p being Element of Prop Q
for E being Event of Borel_Sets st E = (p `2) ` holds
(Meas ((p `1),s)) . (p `2) = 1 - ((Meas ((p `1),s)) . E)
proof end;

definition
let Q be Quantum_Mechanics;
let p be Element of Prop Q;
func 'not' p -> Element of Prop Q equals :: QMAX_1:def 9
[(p `1),((p `2) `)];
coherence
[(p `1),((p `2) `)] is Element of Prop Q
proof end;
involutiveness
for b1, b2 being Element of Prop Q st b1 = [(b2 `1),((b2 `2) `)] holds
b2 = [(b1 `1),((b1 `2) `)]
by MCART_1:21;
end;

:: deftheorem defines 'not' QMAX_1:def 9 :
for Q being Quantum_Mechanics
for p being Element of Prop Q holds 'not' p = [(p `1),((p `2) `)];

definition
let Q be Quantum_Mechanics;
let p, q be Element of Prop Q;
pred p |- q means :: QMAX_1:def 10
for s being Element of Sts Q holds (Meas ((p `1),s)) . (p `2) <= (Meas ((q `1),s)) . (q `2);
reflexivity
for p being Element of Prop Q
for s being Element of Sts Q holds (Meas ((p `1),s)) . (p `2) <= (Meas ((p `1),s)) . (p `2)
;
end;

:: deftheorem defines |- QMAX_1:def 10 :
for Q being Quantum_Mechanics
for p, q being Element of Prop Q holds
( p |- q iff for s being Element of Sts Q holds (Meas ((p `1),s)) . (p `2) <= (Meas ((q `1),s)) . (q `2) );

definition
let Q be Quantum_Mechanics;
let p, q be Element of Prop Q;
pred p <==> q means :: QMAX_1:def 11
( p |- q & q |- p );
reflexivity
for p being Element of Prop Q holds
( p |- p & p |- p )
;
symmetry
for p, q being Element of Prop Q st p |- q & q |- p holds
( q |- p & p |- q )
;
end;

:: deftheorem defines <==> QMAX_1:def 11 :
for Q being Quantum_Mechanics
for p, q being Element of Prop Q holds
( p <==> q iff ( p |- q & q |- p ) );

theorem Th2: :: QMAX_1:2
for Q being Quantum_Mechanics
for p, q being Element of Prop Q holds
( p <==> q iff for s being Element of Sts Q holds (Meas ((p `1),s)) . (p `2) = (Meas ((q `1),s)) . (q `2) )
proof end;

theorem :: QMAX_1:3
for Q being Quantum_Mechanics
for p being Element of Prop Q holds p |- p ;

theorem Th4: :: QMAX_1:4
for Q being Quantum_Mechanics
for p, q, r being Element of Prop Q st p |- q & q |- r holds
p |- r
proof end;

theorem :: QMAX_1:5
for Q being Quantum_Mechanics
for p being Element of Prop Q holds p <==> p ;

theorem :: QMAX_1:6
for Q being Quantum_Mechanics
for p, q being Element of Prop Q st p <==> q holds
q <==> p ;

theorem Th7: :: QMAX_1:7
for Q being Quantum_Mechanics
for p, q, r being Element of Prop Q st p <==> q & q <==> r holds
p <==> r by Th4;

theorem :: QMAX_1:8
canceled;

::$CT
theorem Th8: :: QMAX_1:9
for Q being Quantum_Mechanics
for p, q being Element of Prop Q st p |- q holds
'not' q |- 'not' p
proof end;

definition
let Q be Quantum_Mechanics;
func PropRel Q -> Equivalence_Relation of (Prop Q) means :Def12: :: QMAX_1:def 12
for p, q being Element of Prop Q holds
( [p,q] in it iff p <==> q );
existence
ex b1 being Equivalence_Relation of (Prop Q) st
for p, q being Element of Prop Q holds
( [p,q] in b1 iff p <==> q )
proof end;
uniqueness
for b1, b2 being Equivalence_Relation of (Prop Q) st ( for p, q being Element of Prop Q holds
( [p,q] in b1 iff p <==> q ) ) & ( for p, q being Element of Prop Q holds
( [p,q] in b2 iff p <==> q ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def12 defines PropRel QMAX_1:def 12 :
for Q being Quantum_Mechanics
for b2 being Equivalence_Relation of (Prop Q) holds
( b2 = PropRel Q iff for p, q being Element of Prop Q holds
( [p,q] in b2 iff p <==> q ) );

theorem Th9: :: QMAX_1:10
for Q being Quantum_Mechanics
for B, C being Subset of (Prop Q) st B in Class (PropRel Q) & C in Class (PropRel Q) holds
for a, b, c, d being Element of Prop Q st a in B & b in B & c in C & d in C & a |- c holds
b |- d
proof end;

definition
let Q be Quantum_Mechanics;
func OrdRel Q -> Relation of (Class (PropRel Q)) means :Def13: :: QMAX_1:def 13
for B, C being Subset of (Prop Q) holds
( [B,C] in it iff ( B in Class (PropRel Q) & C in Class (PropRel Q) & ( for p, q being Element of Prop Q st p in B & q in C holds
p |- q ) ) );
existence
ex b1 being Relation of (Class (PropRel Q)) st
for B, C being Subset of (Prop Q) holds
( [B,C] in b1 iff ( B in Class (PropRel Q) & C in Class (PropRel Q) & ( for p, q being Element of Prop Q st p in B & q in C holds
p |- q ) ) )
proof end;
uniqueness
for b1, b2 being Relation of (Class (PropRel Q)) st ( for B, C being Subset of (Prop Q) holds
( [B,C] in b1 iff ( B in Class (PropRel Q) & C in Class (PropRel Q) & ( for p, q being Element of Prop Q st p in B & q in C holds
p |- q ) ) ) ) & ( for B, C being Subset of (Prop Q) holds
( [B,C] in b2 iff ( B in Class (PropRel Q) & C in Class (PropRel Q) & ( for p, q being Element of Prop Q st p in B & q in C holds
p |- q ) ) ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def13 defines OrdRel QMAX_1:def 13 :
for Q being Quantum_Mechanics
for b2 being Relation of (Class (PropRel Q)) holds
( b2 = OrdRel Q iff for B, C being Subset of (Prop Q) holds
( [B,C] in b2 iff ( B in Class (PropRel Q) & C in Class (PropRel Q) & ( for p, q being Element of Prop Q st p in B & q in C holds
p |- q ) ) ) );

theorem Th10: :: QMAX_1:11
for Q being Quantum_Mechanics
for p, q being Element of Prop Q holds
( p |- q iff [(Class ((PropRel Q),p)),(Class ((PropRel Q),q))] in OrdRel Q )
proof end;

theorem Th11: :: QMAX_1:12
for Q being Quantum_Mechanics
for B, C being Subset of (Prop Q) st B in Class (PropRel Q) & C in Class (PropRel Q) holds
for p1, q1 being Element of Prop Q st p1 in B & q1 in B & 'not' p1 in C holds
'not' q1 in C
proof end;

theorem :: QMAX_1:13
for Q being Quantum_Mechanics
for B, C being Subset of (Prop Q) st B in Class (PropRel Q) & C in Class (PropRel Q) holds
for p, q being Element of Prop Q st 'not' p in C & 'not' q in C & p in B holds
q in B
proof end;

definition
let Q be Quantum_Mechanics;
func InvRel Q -> Function of (Class (PropRel Q)),(Class (PropRel Q)) means :Def14: :: QMAX_1:def 14
for p being Element of Prop Q holds it . (Class ((PropRel Q),p)) = Class ((PropRel Q),('not' p));
existence
ex b1 being Function of (Class (PropRel Q)),(Class (PropRel Q)) st
for p being Element of Prop Q holds b1 . (Class ((PropRel Q),p)) = Class ((PropRel Q),('not' p))
proof end;
uniqueness
for b1, b2 being Function of (Class (PropRel Q)),(Class (PropRel Q)) st ( for p being Element of Prop Q holds b1 . (Class ((PropRel Q),p)) = Class ((PropRel Q),('not' p)) ) & ( for p being Element of Prop Q holds b2 . (Class ((PropRel Q),p)) = Class ((PropRel Q),('not' p)) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def14 defines InvRel QMAX_1:def 14 :
for Q being Quantum_Mechanics
for b2 being Function of (Class (PropRel Q)),(Class (PropRel Q)) holds
( b2 = InvRel Q iff for p being Element of Prop Q holds b2 . (Class ((PropRel Q),p)) = Class ((PropRel Q),('not' p)) );

theorem :: QMAX_1:14
for Q being Quantum_Mechanics holds OrthoRelStr(# (Class (PropRel Q)),(OrdRel Q),(InvRel Q) #) is_a_Quantum_Logic
proof end;