begin
Lm1:
for k being Nat
for j being Integer st k <> 0 & [\(j / k)/] + 1 >= j / k holds
k >= j - ([\(j / k)/] * k)
Lm2:
not 1 is prime
by INT_2:def 5;
theorem
canceled;
theorem Th2:
theorem Th3:
theorem Th4:
theorem Th5:
theorem Th6:
for
a,
b being
Nat st
a gcd b = 1 holds
for
c being
Nat holds
(a * c) gcd (b * c) = c
theorem Th7:
theorem Th8:
theorem Th9:
theorem Th10:
theorem Th11:
theorem Th12:
theorem Th13:
theorem Th14:
theorem Th15:
theorem Th16:
theorem Th17:
begin
:: deftheorem defines Euler EULER_1:def 1 :
for n being Nat holds Euler n = card { k where k is Element of NAT : ( n,k are_relative_prime & k >= 1 & k <= n ) } ;
set X = { k where k is Element of NAT : ( 1,k are_relative_prime & k >= 1 & k <= 1 ) } ;
for x being set holds
( x in { k where k is Element of NAT : ( 1,k are_relative_prime & k >= 1 & k <= 1 ) } iff x = 1 )
then Lm3:
card {1} = Euler 1
by TARSKI:def 1;
theorem
set X = { k where k is Element of NAT : ( 2,k are_relative_prime & k >= 1 & k <= 2 ) } ;
for x being set holds
( x in { k where k is Element of NAT : ( 2,k are_relative_prime & k >= 1 & k <= 2 ) } iff x = 1 )
then Lm4:
card {1} = Euler 2
by TARSKI:def 1;
theorem
theorem Th20:
theorem
theorem