]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/RELATIONAL/NPlus/fun.ma
propagating the arithmetics library, partial commit
[helm.git] / matita / matita / contribs / RELATIONAL / NPlus / fun.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
16
17 include "NPlus/inv.ma".
18
19 (* Functional properties ****************************************************)
20
21 theorem nplus_total: ∀p,q. ∃r. p ⊕ q ≍ r.
22  intros; elim q; clear q;
23  [ autobatch | decompose; autobatch ].
24 qed.
25
26 theorem nplus_mono: ∀p,q,r1. p ⊕ q ≍ r1 → 
27                     ∀r2. p ⊕ q ≍ r2 → r1 = r2.
28  intros 4; elim H; clear H q r1;
29  [ lapply linear nplus_inv_zero_2 to H1
30  | lapply linear nplus_inv_succ_2 to H3. decompose
31  ]; destruct; autobatch.
32 qed.
33
34 theorem nplus_inj_1: ∀p1, q, r. p1 ⊕ q ≍ r →
35                      ∀p2. p2 ⊕ q ≍ r → p2 = p1.
36  intros 4; elim H; clear H q r;
37  [ lapply linear nplus_inv_zero_2 to H1
38  | lapply linear nplus_inv_succ_2_3 to H3
39  ]; autobatch.
40 qed.