]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/tests/fguidi.ma
Bugs fixed:
[helm.git] / helm / software / matita / tests / fguidi.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "logic/connectives.ma".
16 include "nat/nat.ma".
17
18 definition is_S: nat \to Prop \def
19    \lambda n. match n with 
20       [ O     \Rightarrow False
21       | (S n) \Rightarrow True
22       ].
23
24 definition pred: nat \to nat \def
25    \lambda n. match n with
26       [ O     \Rightarrow O
27       | (S n) \Rightarrow n
28       ]. 
29
30 theorem eq_gen_S_O: \forall x. (S x = O) \to \forall P:Prop. P.
31 intros. apply False_ind. cut (is_S O). elim Hcut. rewrite < H. apply I.
32 qed.
33
34 theorem eq_gen_S_O_cc: (\forall P:Prop. P) \to \forall x. (S x = O).
35 intros. apply H.
36 qed.
37
38 theorem eq_gen_S_S: \forall m,n. (S m) = (S n) \to m = n. 
39 intros. cut ((pred (S m)) = (pred (S n))). 
40 assumption. elim H. autobatch.
41 qed.
42
43 theorem eq_gen_S_S_cc: \forall m,n. m = n \to (S m) = (S n).
44 intros. elim H. autobatch.
45 qed.
46
47 inductive le: nat \to nat \to Prop \def
48      le_zero: \forall n. (le O n)
49    | le_succ: \forall m, n. (le m n) \to (le (S m) (S n)).
50
51 theorem le_refl: \forall x. (le x x).
52 intros. elim x; autobatch.
53 qed.
54
55 theorem le_gen_x_O_aux: \forall x, y. (le x y) \to (y =O) \to 
56                         (x = O).
57 intros 3. elim H. autobatch. apply eq_gen_S_O. exact n1. autobatch.
58 qed.
59
60 theorem le_gen_x_O: \forall x. (le x O) \to (x = O).
61 intros. apply le_gen_x_O_aux. exact O. autobatch. autobatch.
62 qed.
63
64 theorem le_gen_x_O_cc: \forall x. (x = O) \to (le x O).
65 intros. elim H. autobatch.
66 qed.
67
68 theorem le_gen_S_x_aux: \forall m,x,y. (le y x) \to (y = S m) \to 
69                         (\exists n. x = (S n) \land (le m n)).
70 intros 4. elim H; clear H x y.
71 apply eq_gen_S_O. exact m. elim H1. autobatch.
72 clear H2. cut (n = m).
73 elim Hcut. apply ex_intro. exact n1. split; autobatch.
74 apply eq_gen_S_S. autobatch.
75 qed.
76
77 theorem le_gen_S_x: \forall m,x. (le (S m) x) \to 
78                     (\exists n. x = (S n) \land (le m n)).
79 intros. apply le_gen_S_x_aux. exact (S m). autobatch. autobatch.
80 qed.
81
82 theorem le_gen_S_x_cc: \forall m,x. (\exists n. x = (S n) \land (le m n)) \to
83                        (le (S m) x).
84 intros. elim H. elim H1. cut ((S a) = x). elim Hcut. autobatch.
85 elim H2. autobatch.
86 qed.
87
88 theorem le_gen_S_S: \forall m,n. (le (S m) (S n)) \to (le m n).
89 intros.
90 lapply le_gen_S_x to H as H0. elim H0. elim H1. 
91 lapply eq_gen_S_S to H2 as H4. rewrite > H4. assumption.
92 qed.
93
94 theorem le_gen_S_S_cc: \forall m,n. (le m n) \to (le (S m) (S n)).
95 intros. autobatch.
96 qed.
97
98 (*
99 theorem le_trans: \forall x,y. (le x y) \to \forall z. (le y z) \to (le x z).
100 intros 1. elim x; clear H. clear x. 
101 autobatch.
102 fwd H1 [H]. decompose.
103 *)