]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/nat/euler_theorem.ma
- transcript: bugfix
[helm.git] / helm / software / matita / library / nat / euler_theorem.ma
index ab7d8242e1e5752573c61e187807a4aeeecfdbab..e5933ad95d4dd16263c99d7b036869c5db12e93f 100644 (file)
@@ -12,8 +12,6 @@
 (*                                                                        *)
 (**************************************************************************)
 
-set "baseuri" "cic:/matita/nat/euler_theorem".
-
 include "nat/map_iter_p.ma".
 include "nat/totient.ma".
 
@@ -55,7 +53,6 @@ intro.apply (nat_case n)
 qed.
 *)
 
-
 (*this obvious property is useful because simplify, sometimes,
   "simplifies too much", and doesn't allow to obtain this simple result.
  *)
@@ -86,7 +83,7 @@ elim (m)
         reflexivity
       | rewrite > gcd_O_n.
         apply not_eq_to_eqb_false.
-        apply cic:/matita/nat/nat/not_eq_S.con.                
+        apply not_eq_S.             
         unfold Not.
         intro.        
         cut ( (S n) \le O)
@@ -136,7 +133,6 @@ elim (m)
 ]
 qed.
   
-
 lemma totient_card: \forall n.
 totient n = card n (\lambda i.eqb (gcd i n) (S O)).
 intros.
@@ -163,7 +159,7 @@ elim n
       apply (totient_card_aux n2 n2).
       reflexivity    
     | apply not_eq_to_eqb_false.
-      apply cic:/matita/nat/nat/not_eq_S.con.
+      apply not_eq_S.
       unfold Not.
       intro.
       cut ( (S n2) \le O)