]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/ground/arith/arith_2a.ma
milestone update in ground, partial commit
[helm.git] / matita / matita / contribs / lambdadelta / ground / arith / arith_2a.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 "ground/notation/functions/two_0.ma".
16 include "ground/arith/nat_le_minus_plus.ma".
17 include "ground/arith/nat_lt.ma".
18
19 (* ARITHMETICAL PROPERTIES FOR λδ-2A ****************************************)
20
21 interpretation
22   "zero (non-negative integers)"
23   'Two = (ninj (psucc punit)).
24
25 (* Equalities ***************************************************************)
26
27 lemma plus_n_2: ∀n. (n + 𝟐) = n + 𝟏 + 𝟏.
28 // qed.
29
30 lemma arith_b1: ∀a,b,c1. c1 ≤ b → a - c1 - (b - c1) = a - b.
31 #a #b #c1 #H >nminus_comm <nminus_assoc_comm_23 //
32 qed-.
33
34 lemma arith_b2: ∀a,b,c1,c2. c1 + c2 ≤ b → a - c1 - c2 - (b - c1 - c2) = a - b.
35 #a #b #c1 #c2 #H
36 >(nminus_plus_assoc ? c1 c2) >(nminus_plus_assoc ? c1 c2)
37 /2 width=1 by arith_b1/
38 qed-.
39
40 lemma arith_c1x: ∀x,a,b,c1. x + c1 + a - (b + c1) = x + a - b.
41 #x #a #b #c1
42 <nplus_plus_comm_23 //
43 qed.
44
45 lemma arith_h1: ∀a1,a2,b,c1. c1 ≤ a1 → c1 ≤ b →
46                 a1 - c1 + a2 - (b - c1) = a1 + a2 - b.
47 #a1 #a2 #b #c1 #H1 #H2
48 >nminus_plus_comm_23
49 /2 width=1 by arith_b1/
50 qed-.
51
52 lemma arith_i: ∀x,y,z. y < x → x+z-y-(𝟏) = x-y-(𝟏)+z.
53 /2 width=1 by nminus_plus_comm_23/ qed-.
54
55 (* Constructions ************************************************************)
56
57 fact le_repl_sn_conf_aux: ∀x,y,z:nat. x ≤ z → x = y → y ≤ z.
58 // qed-.
59
60 fact le_repl_sn_trans_aux: ∀x,y,z:nat. x ≤ z → y = x → y ≤ z.
61 // qed-.
62
63 lemma monotonic_le_minus_l2: ∀x1,x2,y,z. x1 ≤ x2 → x1 - y - z ≤ x2 - y - z.
64 /3 width=1 by nle_minus_bi_dx/ qed.
65
66 lemma arith_j: ∀x,y,z. x-y-(𝟏) ≤ x-(y-z)-𝟏.
67 /3 width=1 by nle_minus_bi_dx, nle_minus_bi_sn/ qed.
68
69 lemma arith_k_sn: ∀z,x,y,n. z < x → x+n ≤ y → x-z-(𝟏)+n ≤ y-z-𝟏.
70 #z #x #y #n #Hzx #Hxny
71 >nminus_plus_comm_23 [2: /2 width=1 by nle_minus_bi_sn/ ]
72 >nminus_plus_comm_23 [2: /2 width=1 by nlt_des_le/ ]
73 /2 width=1 by monotonic_le_minus_l2/
74 qed.
75
76 lemma arith_k_dx: ∀z,x,y,n. z < x → y ≤ x+n → y-z-(𝟏) ≤ x-z-(𝟏)+n.
77 #z #x #y #n #Hzx #Hyxn
78 >nminus_plus_comm_23 [2: /2 width=1 by nle_minus_bi_sn/ ]
79 >nminus_plus_comm_23 [2: /2 width=1 by nlt_des_le/ ]
80 /2 width=1 by monotonic_le_minus_l2/
81 qed.
82
83 (* Inversions ***************************************************************)
84
85 lemma lt_plus_SO_to_le: ∀x,y. x < y + (𝟏) → x ≤ y.
86 /2 width=1 by nle_inv_succ_bi/ qed-.
87
88 (* Iterators ****************************************************************)
89
90 lemma iter_SO: ∀B:Type[0]. ∀f:B→B. ∀b,l. f^(l+𝟏) b = f (f^l b).
91 #B #f #b #l
92 <niter_succ //
93 qed.