]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground/arith/nat_pred.ma
milestone update in ground, partial commit
[helm.git] / matita / matita / contribs / lambdadelta / ground / arith / nat_pred.ma
index 4fcab387d0bd8d5248fef6c98e41d0abc2a78f9e..4acdf37b5f1997531e980ac9b9195b333c28c71e 100644 (file)
@@ -13,7 +13,7 @@
 (**************************************************************************)
 
 include "ground/notation/functions/downarrow_1.ma".
-include "ground/arith/pnat_dis.ma".
+include "ground/arith/pnat_split.ma".
 include "ground/arith/nat.ma".
 
 (* PREDECESSOR FOR NON-NEGATIVE INTEGERS ************************************)
@@ -21,14 +21,35 @@ include "ground/arith/nat.ma".
 (*** pred *)
 definition npred (m): nat ≝ match m with
 [ nzero  ⇒ 𝟎
-| ninj p ⇒ pdis … (𝟎) ninj p
+| ninj p ⇒ psplit … (𝟎) ninj p
 ].
 
 interpretation
-  "predecessor (non-negative integers"
+  "predecessor (non-negative integers)"
   'DownArrow m = (npred m).
 
-(* Basic rewrites ***********************************************************)
+(* Basic constructions ******************************************************)
 
+(*** pred_O *)
 lemma npred_zero: 𝟎 = ↓𝟎.
 // qed.
+
+lemma npred_one: 𝟎 = ↓𝟏.
+// qed.
+
+lemma npred_psucc (p): ninj p = ↓↑p.
+// qed.
+
+(* Basic inversions *********************************************************)
+
+lemma npred_pnat_inv_refl (p): ninj p = ↓p → ⊥.
+*
+[ <npred_one #H destruct
+| #p /3 width=2 by psucc_inv_refl, eq_inv_ninj_bi/
+]
+qed-.
+
+(*** pred_inv_fix_sn *)
+lemma npred_inv_refl (n): n = ↓n → 𝟎 = n.
+* // #p #H elim (npred_pnat_inv_refl … H)
+qed-.