#
"cic:/matita/arithmetics/nat/minus_pred_pred.con"
"cic:/matita/arithmetics/nat/monotonic_lt_minus_l.con"
-
-"cic:/matita/arithmetics/nat/lt_minus_to_plus.con"
-"cic:/matita/arithmetics/nat/lt_minus_to_plus_r.con"
+#
"cic:/matita/arithmetics/nat/lt_plus_to_minus.con"
"cic:/matita/arithmetics/nat/lt_plus_to_minus_r.con"
"cic:/matita/arithmetics/nat/lt_inv_plus_l.con"
+"cic:/matita/arithmetics/nat/lt_minus_to_plus.con"
+"cic:/matita/arithmetics/nat/lt_minus_to_plus_r.con"
####################################
#n #IH /2 width=3 by nle_trans/
qed.
+lemma nle_minus_succ_sn (m) (n): ↑n - m ≤ ↑(n - m).
+// qed.
+
(*** inv_eq_minus_O *)
lemma nle_eq_minus_O (m) (n): 𝟎 = m - n → m ≤ n.
#m #n @(nat_ind_2 … m n) //
(* ORDER FOR NON-NEGATIVE INTEGERS ******************************************)
+(* Destructions with npred **************************************************)
+
+lemma nle_inv_pred_sn (m) (n): ↓m ≤ n → m ≤ ↑n.
+#m #n elim m -m
+/2 width=1 by nle_succ_bi/
+qed-.
+
(* Constructions with npred *************************************************)
+lemma nle_succ_pred_dx_refl (m): m ≤ ↑↓m.
+#m @nle_inv_pred_sn // qed.
+
(*** le_pred_n *)
lemma nle_pred_sn_refl (m): ↓m ≤ m.
#m elim m -m //
#m #n elim m -m //
/2 width=1 by nle_pred_bi/
qed-.
-
-(* Destructions with npred **************************************************)
-
-lemma nle_inv_pred_sn (m) (n): ↓m ≤ n → m ≤ ↑n.
-#m #n elim m -m
-/2 width=1 by nle_succ_bi/
-qed-.
lemma nlt_zero_succ (m): 𝟎 < ↑m.
/2 width=1 by nle_succ_bi/ qed.
+lemma nlt_succ_bi (m) (n): m < n → ↑m < ↑n.
+/2 width=1 by nle_succ_bi/ qed.
+
(*** le_to_or_lt_eq *)
lemma nle_lt_eq_e (m) (n): m ≤ n → ∨∨ m < n | m = n.
#m #n * -n /3 width=1 by nle_succ_bi, or_introl/
lapply (nle_minus_sn_bi … o Hmn) -Hmn
<(nminus_succ_sn … Hom) //
qed.
+
+(* Destructions with nminus *************************************************)
+
+lemma nlt_fwd_minus_dx (o) (m) (n): m < n - o → o < n.
+#o elim o -o
+[ #m #n <nminus_zero_dx
+ /2 width=3 by le_nlt_trans/
+| #o #IH #m #n <nminus_succ_dx_pred_sn #H
+ /3 width=2 by nlt_inv_pred_dx/
+]
+qed-.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "ground/arith/nat_le_minus_plus.ma".
+include "ground/arith/nat_lt_minus.ma".
+
+(* STRICT ORDER FOR NON-NEGATIVE INTEGERS ***********************************)
+
+(* Constructions with nminus and nplus **************************************)
+
+(*** lt_plus_to_minus *)
+lemma nlt_minus_sn (o) (m) (n): m ≤ n → n < o + m → n - m < o.
+#o #m #n #Hmn #Ho
+lapply (nle_minus_sn … Ho) -Ho
+<nminus_succ_sn //
+qed.
+
+(*** lt_plus_to_minus_r *)
+lemma nlt_minus_dx (o) (m) (n): m + o < n → m < n - o.
+/2 width=1 by nle_minus_dx/ qed.
+
+(*** lt_inv_plus_l *)
+lemma nlt_minus_dx_full (o) (m) (n): m + o < n → ∧∧ o < n & m < n - o.
+/3 width=3 by nlt_minus_dx, le_nlt_trans, conj/ qed-.
+
+(* Inversions with nminus and nplus *****************************************)
+
+(*** lt_minus_to_plus *)
+lemma nlt_inv_minus_sn (o) (m) (n): m - o < n → m < n + o.
+#o #m #n #Ho
+@nle_inv_minus_sn
+@(nle_trans … Ho) -Ho //
+qed-.
+
+(*** lt_minus_to_plus_r *)
+lemma nlt_inv_minus_dx (o) (m) (n): m < n - o → m + o < n.
+#o #m #n #Ho
+lapply (nle_inv_minus_dx ???? Ho) //
+/3 width=2 by nlt_fwd_minus_dx, nlt_des_le/
+qed-.
(* *)
(**************************************************************************)
-include "ground/arith/nat_pred_succ.ma".
+include "ground/arith/nat_le_pred.ma".
include "ground/arith/nat_lt.ma".
(* STRICT ORDER FOR NON-NEGATIVE INTEGERS ***********************************)
#m @(nat_ind … m) -m //
#H elim (nlt_inv_refl … H)
qed-.
+
+lemma nlt_inv_pred_dx (m) (n): m < ↓n → ↑m < n.
+#m #n #H >(nlt_inv_zero_sn n)
+[ /2 width=1 by nlt_succ_bi/
+| /3 width=3 by le_nlt_trans, nlt_le_trans/
+]
+qed-.
qed.
(*** minus_plus *)
-theorem nminus_assoc (o) (m) (n): o-m-n = o-(m+n).
+theorem nminus_plus_assoc (o) (m) (n): o-m-n = o-(m+n).
#o #m #n elim n -n //
#n #IH <nplus_succ_dx <nminus_succ_dx <nminus_succ_dx //
qed-.
(*** minus_plus_plus_l *)
lemma nminus_plus_dx_bi (m) (n) (o): m - n = (m + o) - (n + o).
-#m #n #o <nminus_assoc <nminus_minus_comm //
+#m #n #o <nminus_plus_assoc <nminus_minus_comm //
qed.
(*** plus_minus_plus_plus_l *) (**)