// qed.
lemma arith_b1: āa,b,c1. c1 ⤠b ā a - c1 - (b - c1) = a - b.
-#a #b #c1 #H >nminus_comm <nminus_assoc_comm_23 //
+#a #b #c1 #H >nminus_comm_21 <nminus_assoc_comm_23 //
qed-.
lemma arith_b2: āa,b,c1,c2. c1 + c2 ⤠b ā a - c1 - c2 - (b - c1 - c2) = a - b.
qed.
(*** minus_minus_comm *)
-lemma nminus_comm (o) (m) (n): o - m - n = o - n - m.
-#o #m #n @(nat_ind_succ ⦠n) -n //
+lemma nminus_comm_21 (m) (n1) (n2): m - n1 - n2 = m - n2 - n1.
+#m #n1 #n2 @(nat_ind_succ ⦠n2) -n2 //
qed.
(*** minus_minus_comm3 *)
-lemma nminus_comm_231 (n) (m1) (m2) (m3):
- n-m1-m2-m3 = n-m2-m3-m1.
+lemma nminus_comm_231 (m) (n1) (n2) (n3):
+ m-n1-n2-n3 = m-n2-n3-n1.
// qed.
(*** minus_plus_plus_l *)
lemma nminus_plus_dx_bi (m) (n) (o): m - n = (m + o) - (n + o).
-#m #n #o <nminus_plus_assoc <nminus_comm //
+#m #n #o <nminus_plus_assoc <nminus_comm_21 //
qed.
(* Helper constructions with nplus ******************************************)
--- /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/pnat_pred.ma".
+include "ground/arith/pnat_iter.ma".
+
+(* SUBTRACTION FOR NON-NEGATIVE INTEGERS ************************************)
+
+definition pminus: pnat ā pnat ā pnat ā
+ λp,q. (ppred^q) p.
+
+interpretation
+ "minus (positive integers)"
+ 'minus p q = (pminus p q).
+
+(* Basic constructions ******************************************************)
+
+lemma pminus_unit_dx (p): āp = p - š.
+// qed.
+
+lemma pminus_succ_dx (p) (q): ā(p - q) = p - āq.
+#p #q @(piter_succ ⦠ppred)
+qed.
+
+(* Advanced constructions ***************************************************)
+
+lemma pminus_pred_sn (p) (q): ā(p - q) = āp - q.
+#p #q @(piter_appl ⦠ppred)
+qed.
+
+lemma pminus_unit_sn (q): š = š - q.
+#q elim q -q //
+qed.
+
+lemma pminus_succ_bi (p) (q): p - q = āp - āq.
+#p #q elim q -q //
+qed.
+
+lemma pminus_succ_dx_pred_sn (p) (q): āp - q = p - āq.
+// qed-.
+
+lemma pminus_refl (p): š = p - p.
+#p elim p -p //
+qed.
+
+lemma pminus_succ_sn_refl (p): š = āp - p.
+#p elim p -p //
+qed.
+
+lemma pminus_comm_21 (p) (q1) (q2): p - q1 - q2 = p - q2 - q1.
+#p #q1 #q2 elim q2 -q2 //
+qed.
+
+lemma pminus_comm_231 (p) (q1) (q2) (q3):
+ p-q1-q2-q3 = p-q2-q3-q1.
+// 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/relocation/tr_pap_pat.ma".
+include "ground/relocation/pr_pat_lt.ma".
+
+(* POSITIVE APPLICATION FOR TOTAL RELOCATION MAPS ***************************)
+
+(* Constructions with plt ***************************************************)
+
+lemma tr_pap_increasing (f) (p):
+ p ⤠fļ¼ ā§£āØpā©.
+/2 width=2 by pr_pat_increasing/
+qed.
lemma tr_inj_next (f): āšāØfā© = šāØāfā©.
* //
qed.
+
+(* Basic eliminations *******************************************************)
+
+lemma tr_map_split (f:tr_map):
+ āØāØ āg. ⫯g = f
+ | āg. āg = f.
+* *
+/3 width=2 by ex_intro, or_introl, or_intror/
+qed-.