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.
22 alias num (instance 0) = "natural number".
24 definition pred: nat \to nat \def
25 \lambda n:nat. match n with
27 | (S p) \Rightarrow p ].
29 theorem pred_Sn : \forall n:nat.n=(pred (S n)).
30 intros. simplify. reflexivity.
33 theorem injective_S : injective nat nat S.
37 rewrite > (pred_Sn y).
38 apply eq_f. assumption.
41 theorem inj_S : \forall n,m:nat.(S n)=(S m) \to n=m \def
44 theorem not_eq_S : \forall n,m:nat.
45 \lnot n=m \to S n \neq S m.
46 intros. unfold Not. intros.
47 apply H. apply injective_S. assumption.
50 definition not_zero : nat \to Prop \def
54 | (S p) \Rightarrow True ].
56 theorem not_eq_O_S : \forall n:nat. O \neq S n.
57 intros. unfold Not. intros.
63 theorem not_eq_n_Sn : \forall n:nat. n \neq S n.
66 apply not_eq_S.assumption.
70 \forall n:nat.\forall P:nat \to Prop.
71 P O \to (\forall m:nat. P (S m)) \to P n.
78 \forall n:nat.\forall P:nat \to Prop.
79 (n=O \to P O) \to (\forall m:nat. (n=(S m) \to P (S m))) \to P n.
82 | apply H2;reflexivity ]
86 \forall R:nat \to nat \to Prop.
87 (\forall n:nat. R O n)
88 \to (\forall n:nat. R (S n) O)
89 \to (\forall n,m:nat. R n m \to R (S n) (S m))
90 \to \forall n,m:nat. R n m.
95 | intro; apply H2; apply H3 ] ]
98 theorem decidable_eq_nat : \forall n,m:nat.decidable (n=m).
99 intros.unfold decidable.
100 apply (nat_elim2 (\lambda n,m.(Or (n=m) ((n=m) \to False))))
103 | right; apply not_eq_O_S ]
104 | intro; right; intro; apply (not_eq_O_S n1); apply sym_eq; assumption
106 [ left; apply eq_f; assumption
107 | right; intro; apply H1; apply inj_S; assumption ] ]