begin
:: deftheorem defines seq CALCUL_2:def 1 :
for m, n being natural number holds seq (m,n) = { k where k is Element of NAT : ( 1 + m <= k & k <= n + m ) } ;
theorem Th1:
theorem Th2:
theorem Th3:
theorem Th4:
theorem Th5:
theorem Th6:
theorem Th7:
theorem
theorem
theorem Th10:
theorem Th11:
theorem Th12:
theorem Th13:
theorem Th14:
theorem Th15:
theorem Th16:
theorem Th17:
theorem Th18:
:: deftheorem defines Per CALCUL_2:def 2 :
for D being non empty set
for f being FinSequence of D
for P being Permutation of (dom f) holds Per (f,P) = P * f;
theorem Th19:
theorem Th20:
begin
:: deftheorem Def3 defines Begin CALCUL_2:def 3 :
for f being FinSequence of CQC-WFF
for b2 being Element of CQC-WFF holds
( ( 1 <= len f implies ( b2 = Begin f iff b2 = f . 1 ) ) & ( not 1 <= len f implies ( b2 = Begin f iff b2 = VERUM ) ) );
:: deftheorem Def4 defines Impl CALCUL_2:def 4 :
for f being FinSequence of CQC-WFF st 1 <= len f holds
for b2 being Element of CQC-WFF holds
( b2 = Impl f iff ex F being FinSequence of CQC-WFF st
( b2 = F . (len f) & len F = len f & ( F . 1 = Begin f or len f = 0 ) & ( for n being Element of NAT st 1 <= n & n < len f holds
ex p, q being Element of CQC-WFF st
( p = f . (n + 1) & q = F . n & F . (n + 1) = p => q ) ) ) );
theorem Th21:
theorem Th22:
theorem Th23:
theorem Th24:
theorem Th25:
theorem Th26:
theorem Th27:
theorem Th28:
theorem Th29:
theorem
begin
theorem Th31:
theorem