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 set "baseuri" "cic:/matita/nat/nat".
17 include "higher_order_defs/functions.ma".
19 inductive nat : Set \def
23 definition pred: nat \to nat \def
24 \lambda n:nat. match n with
26 | (S p) \Rightarrow p ].
28 theorem pred_Sn : \forall n:nat.n=(pred (S n)).
32 theorem injective_S : injective nat nat S.
36 rewrite > (pred_Sn y).
37 apply eq_f. assumption.
40 theorem inj_S : \forall n,m:nat.(S n)=(S m) \to n=m \def
43 theorem not_eq_S : \forall n,m:nat.
44 \lnot n=m \to S n \neq S m.
45 intros. unfold Not. intros.
46 apply H. apply injective_S. assumption.
49 definition not_zero : nat \to Prop \def
53 | (S p) \Rightarrow True ].
55 theorem not_eq_O_S : \forall n:nat. O \neq S n.
56 intros. unfold Not. intros.
62 theorem not_eq_n_Sn : \forall n:nat. n \neq S n.
65 apply not_eq_S.assumption.
69 \forall n:nat.\forall P:nat \to Prop.
70 P O \to (\forall m:nat. P (S m)) \to P n.
77 \forall n:nat.\forall P:nat \to Prop.
78 (n=O \to P O) \to (\forall m:nat. (n=(S m) \to P (S m))) \to P n.
81 | apply H2;reflexivity ]
85 \forall R:nat \to nat \to Prop.
86 (\forall n:nat. R O n)
87 \to (\forall n:nat. R (S n) O)
88 \to (\forall n,m:nat. R n m \to R (S n) (S m))
89 \to \forall n,m:nat. R n m.
94 | intro; apply H2; apply H3 ] ]
97 theorem decidable_eq_nat : \forall n,m:nat.decidable (n=m).
98 intros.unfold decidable.
99 apply (nat_elim2 (\lambda n,m.(Or (n=m) ((n=m) \to False))))
102 | right; apply not_eq_O_S ]
103 | intro; right; intro; apply (not_eq_O_S n1); apply sym_eq; assumption
105 [ left; apply eq_f; assumption
106 | right; intro; apply H1; apply inj_S; assumption ] ]