]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/library/nat.ma
version 0.7.1
[helm.git] / helm / matita / library / nat.ma
index b58004b5a2fd789c05770fc6d53971eb14f98676..c125e6775b7a2f48604f5912dad0f7c2c5defde6 100644 (file)
 
 set "baseuri" "cic:/matita/nat/".
 
-alias id "eq" = "cic:/matita/equality/eq.ind#xpointer(1/1)".
-alias id "refl_equal" = "cic:/matita/equality/eq.ind#xpointer(1/1/1)".
-alias id "sym_eq" = "cic:/matita/equality/sym_eq.con".
-alias id "f_equal" = "cic:/matita/equality/f_equal.con".
-alias id "Not" = "cic:/matita/logic/Not.con".
-alias id "False" = "cic:/matita/logic/False.ind#xpointer(1/1)".
-alias id "True" = "cic:/matita/logic/True.ind#xpointer(1/1)".
-alias id "trans_eq" = "cic:/matita/equality/trans_eq.con".
-alias id "I" = "cic:/matita/logic/True.ind#xpointer(1/1/1)".
-alias id "f_equal2" = "cic:/matita/equality/f_equal2.con".
-alias id "False_ind" = "cic:/matita/logic/False_ind.con".
-alias id "false" = "cic:/matita/bool/bool.ind#xpointer(1/1/2)".
-alias id "true" = "cic:/matita/bool/bool.ind#xpointer(1/1/1)".
-alias id "if_then_else" = "cic:/matita/bool/if_then_else.con".
-alias id "EQ" = "cic:/matita/compare/compare.ind#xpointer(1/1/2)".
-alias id "GT" = "cic:/matita/compare/compare.ind#xpointer(1/1/3)".
-alias id "LT" = "cic:/matita/compare/compare.ind#xpointer(1/1/1)".
-alias id "compare" = "cic:/matita/compare/compare.ind#xpointer(1/1)".
-alias id "compare_invert" = "cic:/matita/compare/compare_invert.con".
+include "equality.ma".
+include "logic.ma".
+include "bool.ma".
+include "compare.ma".
 
 inductive nat : Set \def
   | O : nat
@@ -45,21 +30,21 @@ definition pred: nat \to nat \def
 
 theorem pred_Sn : \forall n:nat.
 (eq nat n (pred (S n))).
-intros.
-apply refl_equal.
+intros; reflexivity.
 qed.
 
 theorem injective_S : \forall n,m:nat. 
 (eq nat (S n) (S m)) \to (eq nat n m).
-intros.
-(elim (sym_eq ? ? ? (pred_Sn n))).(elim (sym_eq ? ? ? (pred_Sn m))).
-apply f_equal. assumption.
+intros;
+rewrite > pred_Sn;
+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
@@ -69,37 +54,39 @@ 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.elim (sym_eq ? ? ? 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)).
 intros.elim n.apply O_S.apply not_eq_S.assumption.
 qed.
 
-
 let rec plus n m \def 
  match n with 
  [ O \Rightarrow m
  | (S p) \Rightarrow S (plus p m) ].
 
 theorem plus_n_O: \forall n:nat. eq nat n (plus n O).
-intros.elim n.simplify.apply refl_equal.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)).
-intros.elim n.simplify.apply refl_equal.simplify.apply f_equal.assumption.
+intros.elim n.simplify.reflexivity.
+simplify.apply f_equal.assumption.
 qed.
 
 theorem sym_plus: \forall n,m:nat. eq nat (plus n m) (plus m n).
 intros.elim n.simplify.apply plus_n_O.
-simplify.elim (sym_eq ? ? ? H).apply plus_n_Sm.
+simplify.rewrite > H.apply plus_n_Sm.
 qed.
 
 theorem assoc_plus: 
 \forall n,m,p:nat. eq nat (plus (plus n m) p) (plus n (plus m p)).
-intros.elim n.simplify.apply refl_equal.simplify.apply f_equal.assumption.
+intros.elim n.simplify.reflexivity.
+simplify.apply f_equal.assumption.
 qed.
 
 let rec times n m \def 
@@ -108,23 +95,24 @@ let rec times n m \def
  | (S p) \Rightarrow (plus m (times p m)) ].
 
 theorem times_n_O: \forall n:nat. eq nat O (times n O).
-intros.elim n.simplify.apply refl_equal.simplify.assumption.
+intros.elim n.simplify.reflexivity.
+simplify.assumption.
 qed.
 
 theorem times_n_Sm : 
 \forall n,m:nat. eq nat (plus n (times n  m)) (times n (S m)).
-intros.elim n.simplify.apply refl_equal.
-simplify.apply f_equal.elim H.
-apply trans_eq ? ? (plus (plus e m) (times e m)).apply sym_eq.
-apply assoc_plus.apply trans_eq ? ? (plus (plus m e) (times e m)).
+intros.elim n.simplify.reflexivity.
+simplify.apply f_equal.rewrite < H.
+transitivity (plus (plus e1 m) (times e1 m)).symmetry.
+apply assoc_plus.transitivity (plus (plus m e1) (times e1 m)).
 apply f_equal2.
-apply sym_plus.apply refl_equal.apply assoc_plus.
+apply sym_plus.reflexivity.apply assoc_plus.
 qed.
 
 theorem sym_times : 
 \forall n,m:nat. eq nat (times n m) (times m n).
 intros.elim n.simplify.apply times_n_O.
-simplify.elim (sym_eq ? ? ? H).apply times_n_Sm.
+simplify.rewrite > H.apply times_n_Sm.
 qed.
 
 let rec minus n m \def 
@@ -146,8 +134,8 @@ theorem nat_double_ind :
 (\forall n:nat. R O n) \to 
 (\forall n:nat. R (S n) O) \to 
 (\forall n,m:nat. R n m \to R (S n) (S m)) \to \forall n,m:nat. R n m.
-intros.cut \forall m:nat.R n m.apply Hcut.elim n.apply H.
-apply nat_case m1.apply H1.intros.apply H2. apply H3.
+intros 5.elim n.apply H.
+apply nat_case m.apply H1.intros.apply H2. apply H3.
 qed.
 
 inductive le (n:nat) : nat \to Prop \def
@@ -188,19 +176,19 @@ qed.
 
 theorem le_n_O_eq : \forall n:nat. (le n O) \to (eq nat O n).
 intros.cut (le n O) \to (eq nat O n).apply Hcut. assumption.
-elim n.apply refl_equal.
+elim n.reflexivity.
 apply False_ind.apply  (le_Sn_O ? H2).
 qed.
 
 theorem le_S_n : \forall n,m:nat. le (S n) (S m) \to le n m.
-intros.cut le (pred (S n)) (pred (S m)).exact Hcut.
+intros.change with le (pred (S n)) (pred (S m)).
 elim H.apply le_n.apply trans_le ? (pred x).assumption.
 apply le_pred_n.
 qed.
 
 theorem le_Sn_n : \forall n:nat. Not (le (S n) n).
 intros.elim n.apply le_Sn_O.simplify.intros.
-cut le (S e) e.apply H.assumption.apply le_S_n.assumption.
+cut le (S e1) e1.apply H.assumption.apply le_S_n.assumption.
 qed.
 
 theorem le_antisym : \forall n,m:nat. (le n m) \to (le m n) \to (eq nat n m).
@@ -208,8 +196,8 @@ intros.cut (le n m) \to (le m n) \to (eq nat n m).exact Hcut H H1.
 apply nat_double_ind (\lambda n,m.((le n m) \to (le m n) \to eq nat n m)).
 intros.whd.intros.
 apply le_n_O_eq.assumption.
-intros.whd.intros.apply sym_eq.apply le_n_O_eq.assumption.
-intros.whd.intros.apply f_equal.apply H2.
+intros.symmetry.apply le_n_O_eq.assumption.
+intros.apply f_equal.apply H2.
 apply le_S_n.assumption.
 apply le_S_n.assumption.
 qed.
@@ -248,10 +236,10 @@ theorem nat_compare_invert: \forall n,m:nat.
 eq compare (nat_compare n m) (compare_invert (nat_compare m n)).
 intros. 
 apply nat_double_ind (\lambda n,m.eq compare (nat_compare n m) (compare_invert (nat_compare m n))).
-intros.elim n1.simplify.apply refl_equal.
-simplify.apply refl_equal.
-intro.elim n1.simplify.apply refl_equal.
-simplify.apply refl_equal.
-intros.simplify.elim H.apply refl_equal.
+intros.elim n1.simplify.reflexivity.
+simplify.reflexivity.
+intro.elim n1.simplify.reflexivity.
+simplify.reflexivity.
+intros.simplify.elim H.simplify.reflexivity.
 qed.