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