]> matita.cs.unibo.it Git - helm.git/commitdiff
update in ground
authorFerruccio Guidi <ferruccio.guidi@unibo.it>
Wed, 18 Jan 2023 19:56:10 +0000 (20:56 +0100)
committerFerruccio Guidi <ferruccio.guidi@unibo.it>
Wed, 18 Jan 2023 19:56:10 +0000 (20:56 +0100)
+ some additions
+ some renaming

matita/matita/contribs/lambdadelta/ground/arith/arith_2a.ma
matita/matita/contribs/lambdadelta/ground/arith/nat_minus.ma
matita/matita/contribs/lambdadelta/ground/arith/nat_minus_plus.ma
matita/matita/contribs/lambdadelta/ground/arith/pnat_minus.ma [new file with mode: 0644]
matita/matita/contribs/lambdadelta/ground/relocation/tr_pap_lt.ma [new file with mode: 0644]
matita/matita/contribs/lambdadelta/ground/relocation/tr_pn.ma

index a136174bc8b209fc0c813c013b3382cdca587938..114eddb1e34837a19871dc0dffd802764e0abc54 100644 (file)
@@ -24,7 +24,7 @@ lemma plus_n_2: āˆ€n. (n + šŸ) = n + šŸ + šŸ.
 // 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.
index e4ccf83f4f4b5265f3e11c4874523dbdda5f98c0..6f11682723d94fdf4444c10e258363335538e9d8 100644 (file)
@@ -70,11 +70,11 @@ lemma nminus_succ_sn_refl (m): ninj (šŸ) = ā†‘m - m.
 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.
index e910fdaa0c94802be4767ba40f7e13a349951a46..97c24dbd470ae51b1eb17bb778a677f5df40dfa8 100644 (file)
@@ -38,7 +38,7 @@ 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 ******************************************)
diff --git a/matita/matita/contribs/lambdadelta/ground/arith/pnat_minus.ma b/matita/matita/contribs/lambdadelta/ground/arith/pnat_minus.ma
new file mode 100644 (file)
index 0000000..389510e
--- /dev/null
@@ -0,0 +1,67 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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.
diff --git a/matita/matita/contribs/lambdadelta/ground/relocation/tr_pap_lt.ma b/matita/matita/contribs/lambdadelta/ground/relocation/tr_pap_lt.ma
new file mode 100644 (file)
index 0000000..32ce73d
--- /dev/null
@@ -0,0 +1,25 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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.
index 2418afd112cb4f0cff5fac4a3ccaee22f5f8c66b..c38aad7123b995af280a7443ef6fc21b4c2d1399 100644 (file)
@@ -47,3 +47,12 @@ lemma tr_inj_push (f): ā«Æš­āØfā© = š­āØā«Æfā©.
 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-.