]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/RELATIONAL/NPlus/props.ma
43966337b8fa052f0c71d0394d77b83978069e2d
[helm.git] / helm / software / matita / contribs / RELATIONAL / NPlus / 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/NPlus/props".
16
17 include "NPlus/inv.ma".
18
19 (* Monoidal properties *)
20
21 theorem nplus_conf: \forall p,q,r1. (p + q == r1) \to 
22                     \forall r2. (p + q == r2) \to r1 = r2.
23  intros 4. elim H; clear H q r1;
24  [ lapply linear nplus_gen_zero_2 to H1
25  | lapply linear nplus_gen_succ_2 to H3. decompose
26  ]; subst; auto.
27 qed.
28
29 theorem nplus_zero_1: \forall q. zero + q == q.
30  intros. elim q; clear q; auto.
31 qed.
32
33 theorem nplus_succ_1: \forall p,q,r. NPlus p q r \to 
34                       (succ p) + q == (succ r).
35  intros. elim H; clear H q r; auto.
36 qed.
37
38 theorem nplus_comm: \forall p,q,r. (p + q == r) \to q + p == r.
39  intros. elim H; clear H q r; auto.
40 qed.
41
42 (* Corollaries *)
43
44 theorem nplus_comm_1: \forall p1,q,r1. (p1 + q == r1) \to
45                       \forall p2,r2. (p2 + q == r2) \to
46                       \forall s. (p1 + r2 == s) \to (p2 + r1 == s).
47  intros 4. elim H; clear H q r1;
48  [ lapply linear nplus_gen_zero_2 to H1. subst
49  | lapply linear nplus_gen_succ_2 to H3. decompose. subst.
50    lapply linear nplus_gen_succ_2 to H4. decompose. subst
51  ]; auto.
52 qed.
53
54 (*                      
55 theorem nplus_shift_succ_sx: \forall p,q,r. 
56                              (p + (succ q) == r) \to (succ p) + q == r.
57  intros.
58  lapply linear nplus_gen_succ_2 to H as H0.
59  decompose. subst. auto new timeout=100.
60 qed.
61
62 theorem nplus_shift_succ_dx: \forall p,q,r. 
63                              ((succ p) + q == r) \to p + (succ q) == r.
64  intros.
65  lapply linear nplus_gen_succ_1 to H as H0.
66  decompose. subst. auto new timeout=100.
67 qed.
68
69 theorem nplus_trans_1: \forall p,q1,r1. (p + q1 == r1) \to 
70                        \forall q2,r2. (r1 + q2 == r2) \to
71                        \exists q. (q1 + q2 == q) \land p + q == r2.
72  intros 2; elim q1; clear q1; intros;
73  [ lapply linear nplus_gen_zero_2 to H as H0.
74    subst.
75  | lapply linear nplus_gen_succ_2 to H1 as H0.
76    decompose. subst.
77    lapply linear nplus_gen_succ_1 to H2 as H0.
78    decompose. subst.
79    lapply linear H to H4, H3 as H0.
80    decompose.
81  ]; apply ex_intro; [| auto new timeout=100 || auto new timeout=100 ]. (**)
82 qed.
83
84 theorem nplus_trans_2: \forall p1,q,r1. (p1 + q == r1) \to 
85                        \forall p2,r2. (p2 + r1 == r2) \to
86                        \exists p. (p1 + p2 == p) \land p + q == r2.
87  intros 2; elim q; clear q; intros;
88  [ lapply linear nplus_gen_zero_2 to H as H0.
89    subst
90  | lapply linear nplus_gen_succ_2 to H1 as H0.
91    decompose. subst.
92    lapply linear nplus_gen_succ_2 to H2 as H0.
93    decompose. subst.
94    lapply linear H to H4, H3 as H0.
95    decompose.
96  ]; apply ex_intro; [| auto new timeout=100 || auto new timeout=100 ]. (**)
97 qed.
98 *)
99