]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/RELATIONAL-ARITHMETICS/Plus_props.ma
new naming
[helm.git] / helm / software / matita / contribs / RELATIONAL-ARITHMETICS / Plus_props.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 set "baseuri" "cic:/matita/RELATIONAL-ARITHMETICsucc/Plus_props".
16
17 include "Plus_fwd.ma".
18
19 theorem Plus_zero_1: \forall q. Plus zero q q.
20  intros. elim q; clear q; auto.
21 qed.
22
23 theorem Plus_succ_1: \forall p,q,r. Plus p q r \to Plus (succ p) q (succ r).
24  intros 2. elim q; clear q;
25  [ lapply linear Plus_gen_zero_2 to H as H0.
26    rewrite > H0. clear H0 p
27  | lapply linear Plus_gen_succ_2 to H1 as H0.
28    decompose.
29    rewrite > H2. clear H2 r
30  ]; auto.
31 qed.
32
33 theorem Plus_sym: \forall p,q,r. Plus p q r \to Plus q p r.
34  intros 2. elim q; clear q;
35  [ lapply linear Plus_gen_zero_2 to H as H0.
36    rewrite > H0. clear H0 p
37  | lapply linear Plus_gen_succ_2 to H1 as H0.
38    decompose.
39    rewrite > H2. clear H2 r
40  ]; auto.
41 qed.
42
43 theorem Plus_shift_succ_sx: \forall p,q,r. 
44                             Plus p (succ q) r \to Plus (succ p) q r.
45  intros.
46  lapply linear Plus_gen_succ_2 to H as H0.
47  decompose.
48  rewrite > H1. clear H1 r.
49  auto.
50 qed.
51
52 theorem Plus_shift_succ_dx: \forall p,q,r. 
53                             Plus (succ p) q r \to Plus p (succ q) r.
54  intros.
55  lapply linear Plus_gen_succ_1 to H as H0.
56  decompose.
57  rewrite > H1. clear H1 r.
58  auto.
59 qed.
60
61 theorem Plus_trans_1: \forall p,q1,r1. Plus p q1 r1 \to 
62                       \forall q2,r2. Plus r1 q2 r2 \to
63                       \exists q. Plus q1 q2 q \land Plus p q r2.
64  intros 2; elim q1; clear q1; intros;
65  [ lapply linear Plus_gen_zero_2 to H as H0.
66    rewrite > H0. clear H0 p
67  | lapply linear Plus_gen_succ_2 to H1 as H0.
68    decompose.
69    rewrite > H3. rewrite > H3 in H2. clear H3 r1.
70    lapply linear Plus_gen_succ_1 to H2 as H0.
71    decompose.
72    rewrite > H2. clear H2 r2.
73    lapply linear H to H4, H3 as H0.
74    decompose.
75  ]; apply ex_intro; [| auto || auto ]. (**)
76 qed.
77
78 theorem Plus_trans_2: \forall p1,q,r1. Plus p1 q r1 \to 
79                       \forall p2,r2. Plus p2 r1 r2 \to
80                       \exists p. Plus p1 p2 p \land Plus p q r2.
81  intros 2; elim q; clear q; intros;
82  [ lapply linear Plus_gen_zero_2 to H as H0.
83    rewrite > H0. clear H0 p1
84  | lapply linear Plus_gen_succ_2 to H1 as H0.
85    decompose.
86    rewrite > H3. rewrite > H3 in H2. clear H3 r1.
87    lapply linear Plus_gen_succ_2 to H2 as H0.
88    decompose.
89    rewrite > H2. clear H2 r2.
90    lapply linear H to H4, H3 as H0.
91    decompose.
92  ]; apply ex_intro; [| auto || auto ]. (**)
93 qed.
94
95 theorem Plus_conf: \forall p,q,r1. Plus p q r1 \to 
96                    \forall r2. Plus p q r2 \to r1 = r2.
97  intros 2. elim q; clear q; intros;
98  [ lapply linear Plus_gen_zero_2 to H as H0.
99    rewrite > H0 in H1. clear H0 p
100  | lapply linear Plus_gen_succ_2 to H1 as H0.
101    decompose.
102    rewrite > H3. clear H3 r1.
103    lapply linear Plus_gen_succ_2 to H2 as H0.
104    decompose.
105    rewrite > H2. clear H2 r2.
106  ]; auto.
107 qed.