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 theorem esempio: \forall A,B,C:Prop.(A \to B \to C) \to (A \to B)
22 inductive nat : Set \def
26 definition pred: nat \to nat \def
27 \lambda n:nat. match n with
29 | (S p) \Rightarrow p ].
31 theorem pred_Sn : \forall n:nat.n=(pred (S n)).
32 intros. simplify. reflexivity.
35 theorem injective_S : injective nat nat S.
39 rewrite > (pred_Sn y).
40 apply eq_f. assumption.
43 theorem inj_S : \forall n,m:nat.(S n)=(S m) \to n=m \def
46 theorem not_eq_S : \forall n,m:nat.
47 \lnot n=m \to S n \neq S m.
48 intros. unfold Not. intros.
49 apply H. apply injective_S. assumption.
52 definition not_zero : nat \to Prop \def
56 | (S p) \Rightarrow True ].
58 theorem not_eq_O_S : \forall n:nat. O \neq S n.
59 intros. unfold Not. intros.
65 theorem not_eq_n_Sn : \forall n:nat. n \neq S n.
68 apply not_eq_S.assumption.
72 \forall n:nat.\forall P:nat \to Prop.
73 P O \to (\forall m:nat. P (S m)) \to P n.
80 \forall n:nat.\forall P:nat \to Prop.
81 (n=O \to P O) \to (\forall m:nat. (n=(S m) \to P (S m))) \to P n.
84 | apply H2;reflexivity ]
88 \forall R:nat \to nat \to Prop.
89 (\forall n:nat. R O n)
90 \to (\forall n:nat. R (S n) O)
91 \to (\forall n,m:nat. R n m \to R (S n) (S m))
92 \to \forall n,m:nat. R n m.
97 | intro; apply H2; apply H3 ] ]
100 theorem decidable_eq_nat : \forall n,m:nat.decidable (n=m).
101 intros.unfold decidable.
102 apply (nat_elim2 (\lambda n,m.(Or (n=m) ((n=m) \to False))))
105 | right; apply not_eq_O_S ]
106 | intro; right; intro; apply (not_eq_O_S n1); apply sym_eq; assumption
108 [ left; apply eq_f; assumption
109 | right; intro; apply H1; apply inj_S; assumption ] ]