begin
deffunc H1( non empty LattStr ) -> set = the carrier of $1;
deffunc H2( LattStr ) -> Element of bool [:[: the carrier of $1, the carrier of $1:], the carrier of $1:] = the L_join of $1;
deffunc H3( LattStr ) -> Element of bool [:[: the carrier of $1, the carrier of $1:], the carrier of $1:] = the L_meet of $1;
deffunc H4( multMagma ) -> Element of bool [:[: the carrier of $1, the carrier of $1:], the carrier of $1:] = the multF of $1;
theorem Th1:
theorem Th2:
:: deftheorem defines directed QUANTAL1:def 1 :
for L being non empty LattStr
for X being Subset of L holds
( X is directed iff for Y being finite Subset of X ex x being Element of L st
( "\/" (Y,L) [= x & x in X ) );
theorem
:: deftheorem defines with_left-zero QUANTAL1:def 2 :
for IT being non empty multMagma holds
( IT is with_left-zero iff ex a being Element of IT st
for b being Element of IT holds a * b = a );
:: deftheorem defines with_right-zero QUANTAL1:def 3 :
for IT being non empty multMagma holds
( IT is with_right-zero iff ex b being Element of IT st
for a being Element of IT holds a * b = b );
:: deftheorem Def4 defines with_zero QUANTAL1:def 4 :
for IT being non empty multMagma holds
( IT is with_zero iff ( IT is with_left-zero & IT is with_right-zero ) );
:: deftheorem Def5 defines right-distributive QUANTAL1:def 5 :
for IT being non empty QuantaleStr holds
( IT is right-distributive iff for a being Element of IT
for X being set holds a [*] ("\/" (X,IT)) = "\/" ( { (a [*] b) where b is Element of IT : b in X } ,IT) );
:: deftheorem Def6 defines left-distributive QUANTAL1:def 6 :
for IT being non empty QuantaleStr holds
( IT is left-distributive iff for a being Element of IT
for X being set holds ("\/" (X,IT)) [*] a = "\/" ( { (b [*] a) where b is Element of IT : b in X } ,IT) );
:: deftheorem defines times-additive QUANTAL1:def 7 :
for IT being non empty QuantaleStr holds
( IT is times-additive iff for a, b, c being Element of IT holds
( (a "\/" b) [*] c = (a [*] c) "\/" (b [*] c) & c [*] (a "\/" b) = (c [*] a) "\/" (c [*] b) ) );
:: deftheorem defines times-continuous QUANTAL1:def 8 :
for IT being non empty QuantaleStr holds
( IT is times-continuous iff for X1, X2 being Subset of IT st X1 is directed & X2 is directed holds
("\/" X1) [*] ("\/" X2) = "\/" ( { (a [*] b) where a, b is Element of IT : ( a in X1 & b in X2 ) } ,IT) );
theorem Th4:
theorem Th5:
theorem Th6:
registration
let A be non
empty set ;
let b1,
b2,
b3 be
BinOp of
A;
let e be
Element of
A;
cluster QuasiNetStr(#
A,
b1,
b2,
b3,
e #)
-> non
empty ;
coherence
not QuasiNetStr(# A,b1,b2,b3,e #) is empty
;
end;
theorem
theorem Th8:
theorem Th9:
:: deftheorem defines idempotent QUANTAL1:def 9 :
for f being Function holds
( f is idempotent iff f * f = f );
Lm1:
for A being non empty set
for f being UnOp of A st f is idempotent holds
for a being Element of A holds f . (f . a) = f . a
:: deftheorem defines inflationary QUANTAL1:def 10 :
for L being non empty LattStr
for IT being UnOp of L holds
( IT is inflationary iff for p being Element of L holds p [= IT . p );
:: deftheorem defines deflationary QUANTAL1:def 11 :
for L being non empty LattStr
for IT being UnOp of L holds
( IT is deflationary iff for p being Element of L holds IT . p [= p );
:: deftheorem Def12 defines monotone QUANTAL1:def 12 :
for L being non empty LattStr
for IT being UnOp of L holds
( IT is monotone iff for p, q being Element of L st p [= q holds
IT . p [= IT . q );
:: deftheorem defines \/-distributive QUANTAL1:def 13 :
for L being non empty LattStr
for IT being UnOp of L holds
( IT is \/-distributive iff for X being Subset of L holds IT . ("\/" X) [= "\/" ( { (IT . a) where a is Element of L : a in X } ,L) );
theorem Th10:
:: deftheorem defines times-monotone QUANTAL1:def 14 :
for Q being non empty QuantaleStr
for IT being UnOp of Q holds
( IT is times-monotone iff for a, b being Element of Q holds (IT . a) [*] (IT . b) [= IT . (a [*] b) );
:: deftheorem defines -r> QUANTAL1:def 15 :
for Q being non empty QuantaleStr
for a, b being Element of Q holds a -r> b = "\/" ( { c where c is Element of Q : c [*] a [= b } ,Q);
:: deftheorem defines -l> QUANTAL1:def 16 :
for Q being non empty QuantaleStr
for a, b being Element of Q holds a -l> b = "\/" ( { c where c is Element of Q : a [*] c [= b } ,Q);
theorem Th11:
theorem Th12:
theorem Th13:
theorem
:: deftheorem Def17 defines dualizing QUANTAL1:def 17 :
for Q being non empty QuantaleStr
for IT being Element of Q holds
( IT is dualizing iff for a being Element of Q holds
( (a -r> IT) -l> IT = a & (a -l> IT) -r> IT = a ) );
:: deftheorem Def18 defines cyclic QUANTAL1:def 18 :
for Q being non empty QuantaleStr
for IT being Element of Q holds
( IT is cyclic iff for a being Element of Q holds a -r> IT = a -l> IT );
theorem Th15:
theorem Th16:
theorem
theorem
theorem Th19:
theorem
:: deftheorem Def19 defines cyclic QUANTAL1:def 19 :
for IT being non empty Girard-QuantaleStr holds
( IT is cyclic iff the absurd of IT is cyclic );
:: deftheorem Def20 defines dualized QUANTAL1:def 20 :
for IT being non empty Girard-QuantaleStr holds
( IT is dualized iff the absurd of IT is dualizing );
theorem Th21:
registration
let A be non
empty set ;
let b1,
b2,
b3 be
BinOp of
A;
let e1,
e2 be
Element of
A;
cluster Girard-QuantaleStr(#
A,
b1,
b2,
b3,
e1,
e2 #)
-> non
empty ;
coherence
not Girard-QuantaleStr(# A,b1,b2,b3,e1,e2 #) is empty
;
end;
:: deftheorem defines Bottom QUANTAL1:def 21 :
for G being Girard-QuantaleStr holds Bottom G = the absurd of G;
:: deftheorem defines Top QUANTAL1:def 22 :
for G being non empty Girard-QuantaleStr holds Top G = (Bottom G) -r> (Bottom G);
:: deftheorem defines Bottom QUANTAL1:def 23 :
for G being non empty Girard-QuantaleStr
for a being Element of G holds Bottom a = a -r> (Bottom G);
:: deftheorem defines Negation QUANTAL1:def 24 :
for G being non empty Girard-QuantaleStr
for b2 being UnOp of G holds
( b2 = Negation G iff for a being Element of G holds b2 . a = Bottom a );
:: deftheorem defines Bottom QUANTAL1:def 25 :
for G being non empty Girard-QuantaleStr
for u being UnOp of G holds Bottom u = (Negation G) * u;
:: deftheorem defines Bottom QUANTAL1:def 26 :
for G being non empty Girard-QuantaleStr
for o being BinOp of G holds Bottom o = (Negation G) * o;
theorem Th22:
theorem
theorem Th24:
theorem Th25:
theorem Th26:
:: deftheorem defines delta QUANTAL1:def 27 :
for Q being Girard-Quantale
for a, b being Element of Q holds a delta b = Bottom ((Bottom a) [*] (Bottom b));
theorem
theorem
theorem
theorem
theorem
theorem Th32:
theorem
theorem