]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/library/nat.ma
1. tactical "try_tacticals" renamed to "first"
[helm.git] / helm / matita / library / nat.ma
index 066018d9e0defaa009cb7ea7c359a7b4b3bf1689..3ab39d9e485883e51868263d0b47c64cb42aebff 100644 (file)
@@ -30,21 +30,21 @@ definition pred: nat \to nat \def
 
 theorem pred_Sn : \forall n:nat.
 (eq nat n (pred (S n))).
-intros.reflexivity.
+introsreflexivity.
 qed.
 
 theorem injective_S : \forall n,m:nat. 
 (eq nat (S n) (S m)) \to (eq nat n m).
-intros.
-rewrite > pred_Sn n.
-rewrite > pred_Sn m.
-apply f_equal. assumption.
+intros;
+rewrite > pred_Sn n;
+rewrite > pred_Sn m;
+apply f_equal; assumption.
 qed.
 
 theorem not_eq_S  : \forall n,m:nat. 
 Not (eq nat n m) \to Not (eq nat (S n) (S m)).
-intros. simplify.intros.
-apply H.apply injective_S.assumption.
+intros; simplify; intros;
+apply H; apply injective_S; assumption.
 qed.
 
 definition not_zero : nat \to Prop \def
@@ -54,9 +54,8 @@ definition not_zero : nat \to Prop \def
   | (S p) \Rightarrow True ].
 
 theorem O_S : \forall n:nat. Not (eq nat O (S n)).
-intros.simplify.intros.
-cut (not_zero O).exact Hcut.rewrite > H.
-exact I.
+intros; simplify; intros;
+cut (not_zero O); [ exact Hcut | rewrite > H; exact I ].
 qed.
 
 theorem n_Sn : \forall n:nat. Not (eq nat n (S n)).
@@ -69,8 +68,9 @@ let rec plus n m \def
  | (S p) \Rightarrow S (plus p m) ].
 
 theorem plus_n_O: \forall n:nat. eq nat n (plus n O).
-intros.elim n.simplify.reflexivity.
-simplify.apply f_equal.assumption.
+intros;elim n;
+ [ simplify;reflexivity
+ | simplify;apply f_equal;assumption ].
 qed.
 
 theorem plus_n_Sm : \forall n,m:nat. eq nat (S (plus n  m)) (plus n (S m)).