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