begin
theorem Th1:
theorem Th2:
theorem Th3:
theorem Th4:
theorem Th5:
theorem
theorem
theorem
theorem Th9:
scheme
CQCF2FuncEx{
F1()
-> non
empty set ,
F2()
-> non
empty set ,
F3()
-> Element of
Funcs (
F1(),
F2()),
F4(
set ,
set ,
set )
-> Element of
Funcs (
F1(),
F2()),
F5(
set ,
set )
-> Element of
Funcs (
F1(),
F2()),
F6(
set ,
set ,
set ,
set )
-> Element of
Funcs (
F1(),
F2()),
F7(
set ,
set ,
set )
-> Element of
Funcs (
F1(),
F2()) } :
ex
F being
Function of
CQC-WFF,
(Funcs (F1(),F2())) st
(
F . VERUM = F3() & ( for
k being
Element of
NAT for
l being
CQC-variable_list of
k for
P being
QC-pred_symbol of
k holds
F . (P ! l) = F4(
k,
P,
l) ) & ( for
r,
s being
Element of
CQC-WFF for
x being
Element of
bound_QC-variables holds
(
F . ('not' r) = F5(
(F . r),
r) &
F . (r '&' s) = F6(
(F . r),
(F . s),
r,
s) &
F . (All (x,r)) = F7(
x,
(F . r),
r) ) ) )
scheme
CQCF2FUniq{
F1()
-> non
empty set ,
F2()
-> non
empty set ,
F3()
-> Function of
CQC-WFF,
(Funcs (F1(),F2())),
F4()
-> Function of
CQC-WFF,
(Funcs (F1(),F2())),
F5()
-> Function of
F1(),
F2(),
F6(
set ,
set ,
set )
-> Function of
F1(),
F2(),
F7(
set ,
set )
-> Function of
F1(),
F2(),
F8(
set ,
set ,
set ,
set )
-> Function of
F1(),
F2(),
F9(
set ,
set ,
set )
-> Function of
F1(),
F2() } :
provided
A1:
F3()
. VERUM = F5()
and A2:
for
k being
Element of
NAT for
ll being
CQC-variable_list of
k for
P being
QC-pred_symbol of
k holds
F3()
. (P ! ll) = F6(
k,
P,
ll)
and A3:
for
r,
s being
Element of
CQC-WFF for
x being
Element of
bound_QC-variables holds
(
F3()
. ('not' r) = F7(
(F3() . r),
r) &
F3()
. (r '&' s) = F8(
(F3() . r),
(F3() . s),
r,
s) &
F3()
. (All (x,r)) = F9(
x,
(F3() . r),
r) )
and A4:
F4()
. VERUM = F5()
and A5:
for
k being
Element of
NAT for
ll being
CQC-variable_list of
k for
P being
QC-pred_symbol of
k holds
F4()
. (P ! ll) = F6(
k,
P,
ll)
and A6:
for
r,
s being
Element of
CQC-WFF for
x being
Element of
bound_QC-variables holds
(
F4()
. ('not' r) = F7(
(F4() . r),
r) &
F4()
. (r '&' s) = F8(
(F4() . r),
(F4() . s),
r,
s) &
F4()
. (All (x,r)) = F9(
x,
(F4() . r),
r) )
theorem Th10:
theorem Th11:
theorem Th12:
theorem Th13:
:: deftheorem Def1 defines NEGATIVE CQC_SIM1:def 1 :
for D being non empty set
for f being Function of D,CQC-WFF
for b3 being Element of Funcs (D,CQC-WFF) holds
( b3 = NEGATIVE f iff for a being Element of D
for p being Element of CQC-WFF st p = f . a holds
b3 . a = 'not' p );
definition
let f,
g be
Function of
[:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):],
CQC-WFF;
let n be
Element of
NAT ;
func CON (
f,
g,
n)
-> Element of
Funcs (
[:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):],
CQC-WFF)
means :
Def2:
for
k being
Element of
NAT for
h being
Element of
Funcs (
bound_QC-variables,
bound_QC-variables)
for
p,
q being
Element of
CQC-WFF st
p = f . (
k,
h) &
q = g . (
(k + n),
h) holds
it . (
k,
h)
= p '&' q;
existence
ex b1 being Element of Funcs ([:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):],CQC-WFF) st
for k being Element of NAT
for h being Element of Funcs (bound_QC-variables,bound_QC-variables)
for p, q being Element of CQC-WFF st p = f . (k,h) & q = g . ((k + n),h) holds
b1 . (k,h) = p '&' q
uniqueness
for b1, b2 being Element of Funcs ([:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):],CQC-WFF) st ( for k being Element of NAT
for h being Element of Funcs (bound_QC-variables,bound_QC-variables)
for p, q being Element of CQC-WFF st p = f . (k,h) & q = g . ((k + n),h) holds
b1 . (k,h) = p '&' q ) & ( for k being Element of NAT
for h being Element of Funcs (bound_QC-variables,bound_QC-variables)
for p, q being Element of CQC-WFF st p = f . (k,h) & q = g . ((k + n),h) holds
b2 . (k,h) = p '&' q ) holds
b1 = b2
end;
:: deftheorem Def2 defines CON CQC_SIM1:def 2 :
for f, g being Function of [:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):],CQC-WFF
for n being Element of NAT
for b4 being Element of Funcs ([:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):],CQC-WFF) holds
( b4 = CON (f,g,n) iff for k being Element of NAT
for h being Element of Funcs (bound_QC-variables,bound_QC-variables)
for p, q being Element of CQC-WFF st p = f . (k,h) & q = g . ((k + n),h) holds
b4 . (k,h) = p '&' q );
Lm1:
for x being Element of bound_QC-variables
for k being Element of NAT
for h being Element of Funcs (bound_QC-variables,bound_QC-variables) holds h +* (x .--> (x. k)) is Function of bound_QC-variables,bound_QC-variables
definition
let f be
Function of
[:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):],
CQC-WFF;
let x be
bound_QC-variable;
func UNIVERSAL (
x,
f)
-> Element of
Funcs (
[:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):],
CQC-WFF)
means :
Def3:
for
k being
Element of
NAT for
h being
Element of
Funcs (
bound_QC-variables,
bound_QC-variables)
for
p being
Element of
CQC-WFF st
p = f . (
(k + 1),
(h +* (x .--> (x. k)))) holds
it . (
k,
h)
= All (
(x. k),
p);
existence
ex b1 being Element of Funcs ([:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):],CQC-WFF) st
for k being Element of NAT
for h being Element of Funcs (bound_QC-variables,bound_QC-variables)
for p being Element of CQC-WFF st p = f . ((k + 1),(h +* (x .--> (x. k)))) holds
b1 . (k,h) = All ((x. k),p)
uniqueness
for b1, b2 being Element of Funcs ([:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):],CQC-WFF) st ( for k being Element of NAT
for h being Element of Funcs (bound_QC-variables,bound_QC-variables)
for p being Element of CQC-WFF st p = f . ((k + 1),(h +* (x .--> (x. k)))) holds
b1 . (k,h) = All ((x. k),p) ) & ( for k being Element of NAT
for h being Element of Funcs (bound_QC-variables,bound_QC-variables)
for p being Element of CQC-WFF st p = f . ((k + 1),(h +* (x .--> (x. k)))) holds
b2 . (k,h) = All ((x. k),p) ) holds
b1 = b2
end;
:: deftheorem Def3 defines UNIVERSAL CQC_SIM1:def 3 :
for f being Function of [:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):],CQC-WFF
for x being bound_QC-variable
for b3 being Element of Funcs ([:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):],CQC-WFF) holds
( b3 = UNIVERSAL (x,f) iff for k being Element of NAT
for h being Element of Funcs (bound_QC-variables,bound_QC-variables)
for p being Element of CQC-WFF st p = f . ((k + 1),(h +* (x .--> (x. k)))) holds
b3 . (k,h) = All ((x. k),p) );
Lm2:
for k being Element of NAT
for f being CQC-variable_list of k holds f is FinSequence of bound_QC-variables
definition
let k be
Element of
NAT ;
let P be
QC-pred_symbol of
k;
let l be
CQC-variable_list of
k;
func ATOMIC (
P,
l)
-> Element of
Funcs (
[:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):],
CQC-WFF)
means :
Def4:
for
n being
Element of
NAT for
h being
Element of
Funcs (
bound_QC-variables,
bound_QC-variables) holds
it . (
n,
h)
= P ! (h * l);
existence
ex b1 being Element of Funcs ([:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):],CQC-WFF) st
for n being Element of NAT
for h being Element of Funcs (bound_QC-variables,bound_QC-variables) holds b1 . (n,h) = P ! (h * l)
uniqueness
for b1, b2 being Element of Funcs ([:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):],CQC-WFF) st ( for n being Element of NAT
for h being Element of Funcs (bound_QC-variables,bound_QC-variables) holds b1 . (n,h) = P ! (h * l) ) & ( for n being Element of NAT
for h being Element of Funcs (bound_QC-variables,bound_QC-variables) holds b2 . (n,h) = P ! (h * l) ) holds
b1 = b2
end;
:: deftheorem Def4 defines ATOMIC CQC_SIM1:def 4 :
for k being Element of NAT
for P being QC-pred_symbol of k
for l being CQC-variable_list of k
for b4 being Element of Funcs ([:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):],CQC-WFF) holds
( b4 = ATOMIC (P,l) iff for n being Element of NAT
for h being Element of Funcs (bound_QC-variables,bound_QC-variables) holds b4 . (n,h) = P ! (h * l) );
deffunc H1( set , set , set ) -> Element of NAT = 0 ;
deffunc H2( Element of NAT ) -> Element of NAT = $1;
deffunc H3( Element of NAT , Element of NAT ) -> Element of NAT = $1 + $2;
deffunc H4( set , Element of NAT ) -> Element of NAT = $2 + 1;
definition
let p be
Element of
CQC-WFF ;
func QuantNbr p -> Element of
NAT means :
Def5:
ex
F being
Function of
CQC-WFF,
NAT st
(
it = F . p &
F . VERUM = 0 & ( for
r,
s being
Element of
CQC-WFF for
x being
Element of
bound_QC-variables for
k being
Element of
NAT for
l being
CQC-variable_list of
k for
P being
QC-pred_symbol of
k holds
(
F . (P ! l) = 0 &
F . ('not' r) = F . r &
F . (r '&' s) = (F . r) + (F . s) &
F . (All (x,r)) = (F . r) + 1 ) ) );
correctness
existence
ex b1 being Element of NAT ex F being Function of CQC-WFF,NAT st
( b1 = F . p & F . VERUM = 0 & ( for r, s being Element of CQC-WFF
for x being Element of bound_QC-variables
for k being Element of NAT
for l being CQC-variable_list of k
for P being QC-pred_symbol of k holds
( F . (P ! l) = 0 & F . ('not' r) = F . r & F . (r '&' s) = (F . r) + (F . s) & F . (All (x,r)) = (F . r) + 1 ) ) );
uniqueness
for b1, b2 being Element of NAT st ex F being Function of CQC-WFF,NAT st
( b1 = F . p & F . VERUM = 0 & ( for r, s being Element of CQC-WFF
for x being Element of bound_QC-variables
for k being Element of NAT
for l being CQC-variable_list of k
for P being QC-pred_symbol of k holds
( F . (P ! l) = 0 & F . ('not' r) = F . r & F . (r '&' s) = (F . r) + (F . s) & F . (All (x,r)) = (F . r) + 1 ) ) ) & ex F being Function of CQC-WFF,NAT st
( b2 = F . p & F . VERUM = 0 & ( for r, s being Element of CQC-WFF
for x being Element of bound_QC-variables
for k being Element of NAT
for l being CQC-variable_list of k
for P being QC-pred_symbol of k holds
( F . (P ! l) = 0 & F . ('not' r) = F . r & F . (r '&' s) = (F . r) + (F . s) & F . (All (x,r)) = (F . r) + 1 ) ) ) holds
b1 = b2;
end;
:: deftheorem Def5 defines QuantNbr CQC_SIM1:def 5 :
for p being Element of CQC-WFF
for b2 being Element of NAT holds
( b2 = QuantNbr p iff ex F being Function of CQC-WFF,NAT st
( b2 = F . p & F . VERUM = 0 & ( for r, s being Element of CQC-WFF
for x being Element of bound_QC-variables
for k being Element of NAT
for l being CQC-variable_list of k
for P being QC-pred_symbol of k holds
( F . (P ! l) = 0 & F . ('not' r) = F . r & F . (r '&' s) = (F . r) + (F . s) & F . (All (x,r)) = (F . r) + 1 ) ) ) );
definition
let f be
Function of
CQC-WFF,
(Funcs ([:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):],CQC-WFF));
let x be
Element of
CQC-WFF ;
.redefine func f . x -> Element of
Funcs (
[:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):],
CQC-WFF);
coherence
f . x is Element of Funcs ([:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):],CQC-WFF)
end;
definition
func SepFunc -> Function of
CQC-WFF,
(Funcs ([:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):],CQC-WFF)) means :
Def6:
(
it . VERUM = [:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):] --> VERUM & ( for
k being
Element of
NAT for
l being
CQC-variable_list of
k for
P being
QC-pred_symbol of
k holds
it . (P ! l) = ATOMIC (
P,
l) ) & ( for
r,
s being
Element of
CQC-WFF for
x being
Element of
bound_QC-variables holds
(
it . ('not' r) = NEGATIVE (it . r) &
it . (r '&' s) = CON (
(it . r),
(it . s),
(QuantNbr r)) &
it . (All (x,r)) = UNIVERSAL (
x,
(it . r)) ) ) );
existence
ex b1 being Function of CQC-WFF,(Funcs ([:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):],CQC-WFF)) st
( b1 . VERUM = [:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):] --> VERUM & ( for k being Element of NAT
for l being CQC-variable_list of k
for P being QC-pred_symbol of k holds b1 . (P ! l) = ATOMIC (P,l) ) & ( for r, s being Element of CQC-WFF
for x being Element of bound_QC-variables holds
( b1 . ('not' r) = NEGATIVE (b1 . r) & b1 . (r '&' s) = CON ((b1 . r),(b1 . s),(QuantNbr r)) & b1 . (All (x,r)) = UNIVERSAL (x,(b1 . r)) ) ) )
uniqueness
for b1, b2 being Function of CQC-WFF,(Funcs ([:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):],CQC-WFF)) st b1 . VERUM = [:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):] --> VERUM & ( for k being Element of NAT
for l being CQC-variable_list of k
for P being QC-pred_symbol of k holds b1 . (P ! l) = ATOMIC (P,l) ) & ( for r, s being Element of CQC-WFF
for x being Element of bound_QC-variables holds
( b1 . ('not' r) = NEGATIVE (b1 . r) & b1 . (r '&' s) = CON ((b1 . r),(b1 . s),(QuantNbr r)) & b1 . (All (x,r)) = UNIVERSAL (x,(b1 . r)) ) ) & b2 . VERUM = [:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):] --> VERUM & ( for k being Element of NAT
for l being CQC-variable_list of k
for P being QC-pred_symbol of k holds b2 . (P ! l) = ATOMIC (P,l) ) & ( for r, s being Element of CQC-WFF
for x being Element of bound_QC-variables holds
( b2 . ('not' r) = NEGATIVE (b2 . r) & b2 . (r '&' s) = CON ((b2 . r),(b2 . s),(QuantNbr r)) & b2 . (All (x,r)) = UNIVERSAL (x,(b2 . r)) ) ) holds
b1 = b2
end;
:: deftheorem Def6 defines SepFunc CQC_SIM1:def 6 :
for b1 being Function of CQC-WFF,(Funcs ([:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):],CQC-WFF)) holds
( b1 = SepFunc iff ( b1 . VERUM = [:NAT,(Funcs (bound_QC-variables,bound_QC-variables)):] --> VERUM & ( for k being Element of NAT
for l being CQC-variable_list of k
for P being QC-pred_symbol of k holds b1 . (P ! l) = ATOMIC (P,l) ) & ( for r, s being Element of CQC-WFF
for x being Element of bound_QC-variables holds
( b1 . ('not' r) = NEGATIVE (b1 . r) & b1 . (r '&' s) = CON ((b1 . r),(b1 . s),(QuantNbr r)) & b1 . (All (x,r)) = UNIVERSAL (x,(b1 . r)) ) ) ) );
:: deftheorem defines SepFunc CQC_SIM1:def 7 :
for p being Element of CQC-WFF
for k being Element of NAT
for f being Element of Funcs (bound_QC-variables,bound_QC-variables) holds SepFunc (p,k,f) = (SepFunc . p) . [k,f];
deffunc H5( Element of CQC-WFF ) -> Element of NAT = QuantNbr $1;
theorem
theorem
theorem
theorem
theorem
theorem
canceled;
theorem Th20:
scheme
MaxFinDomElem{
F1()
-> non
empty set ,
F2()
-> set ,
P1[
set ,
set ] } :
ex
x being
Element of
F1() st
(
x in F2() & ( for
y being
Element of
F1() st
y in F2() holds
P1[
x,
y] ) )
provided
A1:
(
F2() is
finite &
F2()
<> {} &
F2()
c= F1() )
and A2:
for
x,
y being
Element of
F1() holds
(
P1[
x,
y] or
P1[
y,
x] )
and A3:
for
x,
y,
z being
Element of
F1() st
P1[
x,
y] &
P1[
y,
z] holds
P1[
x,
z]
:: deftheorem CQC_SIM1:def 8 :
canceled;
:: deftheorem defines NBI CQC_SIM1:def 9 :
for p being Element of CQC-WFF holds NBI p = { k where k is Element of NAT : for i being Element of NAT st k <= i holds
not x. i in still_not-bound_in p } ;
:: deftheorem defines index CQC_SIM1:def 10 :
for p being Element of CQC-WFF holds index p = min (NBI p);
theorem Th21:
theorem Th22:
theorem Th23:
theorem Th24:
theorem
:: deftheorem defines SepVar CQC_SIM1:def 11 :
for p being Element of CQC-WFF holds SepVar p = SepFunc (p,(index p),(id bound_QC-variables));
theorem
theorem Th27:
theorem
theorem Th29:
theorem
definition
let p be
Element of
CQC-WFF ;
let X be
Subset of
[:CQC-WFF,NAT,(Fin bound_QC-variables),(Funcs (bound_QC-variables,bound_QC-variables)):];
pred X is_Sep-closed_on p means :
Def12:
(
[p,(index p),({}. bound_QC-variables),(id bound_QC-variables)] in X & ( for
q being
Element of
CQC-WFF for
k being
Element of
NAT for
K being
Finite_Subset of
bound_QC-variables for
f being
Element of
Funcs (
bound_QC-variables,
bound_QC-variables) st
[('not' q),k,K,f] in X holds
[q,k,K,f] in X ) & ( for
q,
r being
Element of
CQC-WFF for
k being
Element of
NAT for
K being
Finite_Subset of
bound_QC-variables for
f being
Element of
Funcs (
bound_QC-variables,
bound_QC-variables) st
[(q '&' r),k,K,f] in X holds
(
[q,k,K,f] in X &
[r,(k + (QuantNbr q)),K,f] in X ) ) & ( for
q being
Element of
CQC-WFF for
x being
Element of
bound_QC-variables for
k being
Element of
NAT for
K being
Finite_Subset of
bound_QC-variables for
f being
Element of
Funcs (
bound_QC-variables,
bound_QC-variables) st
[(All (x,q)),k,K,f] in X holds
[q,(k + 1),(K \/ {x}),(f +* (x .--> (x. k)))] in X ) );
end;
:: deftheorem Def12 defines is_Sep-closed_on CQC_SIM1:def 12 :
for p being Element of CQC-WFF
for X being Subset of [:CQC-WFF,NAT,(Fin bound_QC-variables),(Funcs (bound_QC-variables,bound_QC-variables)):] holds
( X is_Sep-closed_on p iff ( [p,(index p),({}. bound_QC-variables),(id bound_QC-variables)] in X & ( for q being Element of CQC-WFF
for k being Element of NAT
for K being Finite_Subset of bound_QC-variables
for f being Element of Funcs (bound_QC-variables,bound_QC-variables) st [('not' q),k,K,f] in X holds
[q,k,K,f] in X ) & ( for q, r being Element of CQC-WFF
for k being Element of NAT
for K being Finite_Subset of bound_QC-variables
for f being Element of Funcs (bound_QC-variables,bound_QC-variables) st [(q '&' r),k,K,f] in X holds
( [q,k,K,f] in X & [r,(k + (QuantNbr q)),K,f] in X ) ) & ( for q being Element of CQC-WFF
for x being Element of bound_QC-variables
for k being Element of NAT
for K being Finite_Subset of bound_QC-variables
for f being Element of Funcs (bound_QC-variables,bound_QC-variables) st [(All (x,q)),k,K,f] in X holds
[q,(k + 1),(K \/ {x}),(f +* (x .--> (x. k)))] in X ) ) );
definition
let p be
Element of
CQC-WFF ;
func SepQuadruples p -> Subset of
[:CQC-WFF,NAT,(Fin bound_QC-variables),(Funcs (bound_QC-variables,bound_QC-variables)):] means :
Def13:
(
it is_Sep-closed_on p & ( for
D being
Subset of
[:CQC-WFF,NAT,(Fin bound_QC-variables),(Funcs (bound_QC-variables,bound_QC-variables)):] st
D is_Sep-closed_on p holds
it c= D ) );
existence
ex b1 being Subset of [:CQC-WFF,NAT,(Fin bound_QC-variables),(Funcs (bound_QC-variables,bound_QC-variables)):] st
( b1 is_Sep-closed_on p & ( for D being Subset of [:CQC-WFF,NAT,(Fin bound_QC-variables),(Funcs (bound_QC-variables,bound_QC-variables)):] st D is_Sep-closed_on p holds
b1 c= D ) )
uniqueness
for b1, b2 being Subset of [:CQC-WFF,NAT,(Fin bound_QC-variables),(Funcs (bound_QC-variables,bound_QC-variables)):] st b1 is_Sep-closed_on p & ( for D being Subset of [:CQC-WFF,NAT,(Fin bound_QC-variables),(Funcs (bound_QC-variables,bound_QC-variables)):] st D is_Sep-closed_on p holds
b1 c= D ) & b2 is_Sep-closed_on p & ( for D being Subset of [:CQC-WFF,NAT,(Fin bound_QC-variables),(Funcs (bound_QC-variables,bound_QC-variables)):] st D is_Sep-closed_on p holds
b2 c= D ) holds
b1 = b2
end;
:: deftheorem Def13 defines SepQuadruples CQC_SIM1:def 13 :
for p being Element of CQC-WFF
for b2 being Subset of [:CQC-WFF,NAT,(Fin bound_QC-variables),(Funcs (bound_QC-variables,bound_QC-variables)):] holds
( b2 = SepQuadruples p iff ( b2 is_Sep-closed_on p & ( for D being Subset of [:CQC-WFF,NAT,(Fin bound_QC-variables),(Funcs (bound_QC-variables,bound_QC-variables)):] st D is_Sep-closed_on p holds
b2 c= D ) ) );
theorem Th31:
theorem Th32:
theorem Th33:
for
p,
q,
r being
Element of
CQC-WFF for
k being
Element of
NAT for
K being
Finite_Subset of
bound_QC-variables for
f being
Element of
Funcs (
bound_QC-variables,
bound_QC-variables) st
[(q '&' r),k,K,f] in SepQuadruples p holds
(
[q,k,K,f] in SepQuadruples p &
[r,(k + (QuantNbr q)),K,f] in SepQuadruples p )
theorem Th34:
theorem Th35:
for
q,
p being
Element of
CQC-WFF for
k being
Element of
NAT for
f being
Element of
Funcs (
bound_QC-variables,
bound_QC-variables)
for
K being
Finite_Subset of
bound_QC-variables holds
( not
[q,k,K,f] in SepQuadruples p or
[q,k,K,f] = [p,(index p),({}. bound_QC-variables),(id bound_QC-variables)] or
[('not' q),k,K,f] in SepQuadruples p or ex
r being
Element of
CQC-WFF st
[(q '&' r),k,K,f] in SepQuadruples p or ex
r being
Element of
CQC-WFF ex
l being
Element of
NAT st
(
k = l + (QuantNbr r) &
[(r '&' q),l,K,f] in SepQuadruples p ) or ex
x being
Element of
bound_QC-variables ex
l being
Element of
NAT ex
h being
Element of
Funcs (
bound_QC-variables,
bound_QC-variables) st
(
l + 1
= k &
h +* ({x} --> (x. l)) = f & (
[(All (x,q)),l,K,h] in SepQuadruples p or
[(All (x,q)),l,(K \ {x}),h] in SepQuadruples p ) ) )
scheme
Sepregression{
F1()
-> Element of
CQC-WFF ,
P1[
set ,
set ,
set ,
set ] } :
provided
A1:
P1[
F1(),
index F1(),
{}. bound_QC-variables,
id bound_QC-variables]
and A2:
for
q being
Element of
CQC-WFF for
k being
Element of
NAT for
K being
Finite_Subset of
bound_QC-variables for
f being
Element of
Funcs (
bound_QC-variables,
bound_QC-variables) st
[('not' q),k,K,f] in SepQuadruples F1() &
P1[
'not' q,
k,
K,
f] holds
P1[
q,
k,
K,
f]
and A3:
for
q,
r being
Element of
CQC-WFF for
k being
Element of
NAT for
K being
Finite_Subset of
bound_QC-variables for
f being
Element of
Funcs (
bound_QC-variables,
bound_QC-variables) st
[(q '&' r),k,K,f] in SepQuadruples F1() &
P1[
q '&' r,
k,
K,
f] holds
(
P1[
q,
k,
K,
f] &
P1[
r,
k + (QuantNbr q),
K,
f] )
and A4:
for
q being
Element of
CQC-WFF for
x being
Element of
bound_QC-variables for
k being
Element of
NAT for
K being
Finite_Subset of
bound_QC-variables for
f being
Element of
Funcs (
bound_QC-variables,
bound_QC-variables) st
[(All (x,q)),k,K,f] in SepQuadruples F1() &
P1[
All (
x,
q),
k,
K,
f] holds
P1[
q,
k + 1,
K \/ {x},
f +* (x .--> (x. k))]
theorem Th36:
theorem
theorem
theorem Th39:
theorem Th40:
theorem
theorem Th42:
theorem Th43:
theorem
theorem Th45:
theorem
:: deftheorem Def14 defines are_similar CQC_SIM1:def 14 :
for p, q being Element of CQC-WFF holds
( p,q are_similar iff SepVar p = SepVar q );
theorem
canceled;
theorem
canceled;
theorem