1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| A.Asperti, C.Sacerdoti Coen, *)
8 (* ||A|| E.Tassi, S.Zacchiroli *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU Lesser General Public License Version 2.1 *)
13 (**************************************************************************)
15 include "higher_order_defs/functions.ma".
17 inductive nat : Set \def
21 interpretation "Natural numbers" 'N = nat.
23 default "natural numbers" cic:/matita/nat/nat/nat.ind.
25 alias num (instance 0) = "natural number".
27 definition pred: nat \to nat \def
28 \lambda n:nat. match n with
30 | (S p) \Rightarrow p ].
32 theorem pred_Sn : \forall n:nat.n=(pred (S n)).
33 intros. simplify. reflexivity.
36 theorem injective_S : injective nat nat S.
40 rewrite > (pred_Sn y).
41 apply eq_f. assumption.
44 theorem inj_S : \forall n,m:nat.(S n)=(S m) \to n=m \def
47 theorem not_eq_S : \forall n,m:nat.
48 \lnot n=m \to S n \neq S m.
49 intros. unfold Not. intros.
50 apply H. apply injective_S. assumption.
53 definition not_zero : nat \to Prop \def
57 | (S p) \Rightarrow True ].
59 theorem not_eq_O_S : \forall n:nat. O \neq S n.
60 intros. unfold Not. intros.
66 theorem not_eq_n_Sn : \forall n:nat. n \neq S n.
69 apply not_eq_S.assumption.
73 \forall n:nat.\forall P:nat \to Prop.
74 P O \to (\forall m:nat. P (S m)) \to P n.
81 \forall n:nat.\forall P:nat \to Prop.
82 (n=O \to P O) \to (\forall m:nat. (n=(S m) \to P (S m))) \to P n.
85 | apply H2;reflexivity ]
89 \forall R:nat \to nat \to Prop.
90 (\forall n:nat. R O n)
91 \to (\forall n:nat. R (S n) O)
92 \to (\forall n,m:nat. R n m \to R (S n) (S m))
93 \to \forall n,m:nat. R n m.
98 | intro; apply H2; apply H3 ] ]
101 theorem decidable_eq_nat : \forall n,m:nat.decidable (n=m).
102 intros.unfold decidable.
103 apply (nat_elim2 (\lambda n,m.(Or (n=m) ((n=m) \to False))))
106 | right; apply not_eq_O_S ]
107 | intro; right; intro; apply (not_eq_O_S n1); apply sym_eq; assumption
109 [ left; apply eq_f; assumption
110 | right; intro; apply H1; apply inj_S; assumption ] ]