begin
theorem
theorem
canceled;
theorem Th3:
theorem
canceled;
theorem Th5:
theorem Th6:
theorem Th7:
theorem Th8:
theorem Th9:
theorem Th10:
theorem Th11:
theorem Th12:
theorem Th13:
theorem Th14:
theorem
:: deftheorem defines Prefix FIB_NUM2:def 1 :
for f being Function of NAT,NAT
for A being with_non-empty_elements finite natural-membered set holds Prefix (f,A) = Seq (f | A);
theorem Th16:
theorem Th17:
theorem Th18:
theorem Th19:
theorem Th20:
theorem Th21:
theorem Th22:
begin
scheme
FibInd1{
P1[
set ] } :
provided
A1:
P1[1]
and A2:
P1[2]
and A3:
for
k being non
empty Nat st
P1[
k] &
P1[
k + 1] holds
P1[
k + 2]
scheme
FibInd2{
P1[
set ] } :
provided
A1:
P1[2]
and A2:
P1[3]
and A3:
for
k being non
trivial Nat st
P1[
k] &
P1[
k + 1] holds
P1[
k + 2]
theorem Th23:
theorem Th24:
theorem Th25:
theorem Th26:
theorem Th27:
theorem Th28:
theorem Th29:
Lm1:
for k being Nat holds Fib ((2 * (k + 2)) + 1) = (Fib ((2 * k) + 3)) + (Fib ((2 * k) + 4))
theorem Th30:
theorem Th31:
theorem Th32:
begin
theorem Th33:
theorem
theorem Th35:
theorem Th36:
theorem Th37:
theorem Th38:
theorem Th39:
theorem
theorem
theorem Th42:
theorem Th43:
theorem Th44:
theorem Th45:
theorem Th46:
theorem
theorem Th48:
theorem Th49:
for
k being
Nat holds
(
Fib k = 1 iff (
k = 1 or
k = 2 ) )
theorem Th50:
theorem Th51:
theorem
begin
:: deftheorem Def2 defines FIB FIB_NUM2:def 2 :
for b1 being Function of NAT,NAT holds
( b1 = FIB iff for k being Element of NAT holds b1 . k = Fib k );
:: deftheorem defines EvenNAT FIB_NUM2:def 3 :
EvenNAT = { (2 * k) where k is Element of NAT : verum } ;
:: deftheorem defines OddNAT FIB_NUM2:def 4 :
OddNAT = { ((2 * k) + 1) where k is Element of NAT : verum } ;
theorem Th53:
theorem Th54:
:: deftheorem defines EvenFibs FIB_NUM2:def 5 :
for n being Element of NAT holds EvenFibs n = Prefix (FIB,(EvenNAT /\ (Seg n)));
:: deftheorem defines OddFibs FIB_NUM2:def 6 :
for n being Element of NAT holds OddFibs n = Prefix (FIB,(OddNAT /\ (Seg n)));
theorem Th55:
theorem
theorem Th57:
theorem
theorem Th59:
theorem Th60:
theorem Th61:
theorem Th62:
theorem Th63:
theorem Th64:
theorem Th65:
theorem Th66:
theorem
theorem
begin
theorem Th69:
theorem Th70:
theorem
begin
theorem