]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma
lift functions and identity map
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / lib / arith.ma
index 26898d348056f3367d1c7729c47eedf322d82c9a..fad49e616661cdfaa9d5f633b7252ba79a7192b8 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
+include "ground_2/notation/functions/successor_1.ma".
+include "ground_2/notation/functions/predecessor_1.ma".
 include "arithmetics/nat.ma".
 include "ground_2/lib/star.ma".
 
 (* ARITHMETICAL PROPERTIES **************************************************)
 
+interpretation "nat successor" 'Successor m = (S m).
+
+interpretation "nat predecessor" 'Predecessor m = (pred m).
+
+(* Iota equations ***********************************************************)
+
+lemma pred_O: pred 0 = 0.
+normalize // qed.
+
+lemma pred_S: ∀m. pred (S m) = m.
+// qed.
+
 (* Equations ****************************************************************)
 
+lemma minus_plus_m_m_commutative: ∀n,m:nat. n = m + n - m.
+// qed-.
+
 (* Note: uses minus_minus_comm, minus_plus_m_m, commutative_plus, plus_minus *)
 lemma plus_minus_minus_be: ∀x,y,z. y ≤ z → z ≤ x → (x - z) + (z - y) = x - y.
 #x #z #y #Hzy #Hyx >plus_minus // >commutative_plus >plus_minus //
@@ -89,6 +106,9 @@ lemma monotonic_lt_pred: ∀m,n. m < n → O < m → pred m < pred n.
 @le_S_S_to_le >S_pred /2 width=3 by transitive_lt/
 qed.
 
+lemma lt_S_S: ∀x,y. x < y → ⫯x < ⫯y.
+/2 width=1 by le_S_S/ qed.
+
 lemma arith_j: ∀x,y,z. x-y-1 ≤ x-(y-z)-1.
 /3 width=1 by monotonic_le_minus_l, monotonic_le_minus_r/ qed.
 
@@ -108,6 +128,9 @@ qed.
 
 (* Inversion & forward lemmas ***********************************************)
 
+lemma discr_plus_xy_y: ∀x,y. x + y = y → x = 0.
+// qed-.
+
 lemma lt_plus_SO_to_le: ∀x,y. x < y + 1 → x ≤ y.
 /2 width=1 by monotonic_pred/ qed-.
 
@@ -119,6 +142,9 @@ lemma lt_zero_false: ∀n. n < 0 → ⊥.
 #n #H elim (lt_to_not_le … H) -H /2 width=1 by/
 qed-.
 
+lemma lt_le_false: ∀x,y. x < y → y ≤ x → ⊥.
+/3 width=4 by lt_refl_false, lt_to_le_to_lt/ qed-.
+
 lemma pred_inv_refl: ∀m. pred m = m → m = 0.
 * // normalize #m #H elim (lt_refl_false m) //
 qed-.
@@ -147,6 +173,32 @@ lemma discr_x_minus_xy: ∀x,y. x = x - y → x = 0 ∨ y = 0.
 #H destruct
 qed-.
 
+lemma zero_eq_plus: ∀x,y. 0 = x + y → 0 = x ∧ 0 = y.
+* /2 width=1 by conj/ #x #y normalize #H destruct
+qed-.
+
+lemma lt_S_S_to_lt: ∀x,y. ⫯x < ⫯y → x < y.
+/2 width=1 by le_S_S_to_le/ qed-.
+
+lemma lt_elim: ∀R:relation nat.
+               (∀n2. R O (⫯n2)) →
+               (∀n1,n2. R n1 n2 → R (⫯n1) (⫯n2)) →
+               ∀n2,n1. n1 < n2 → R n1 n2.
+#R #IH1 #IH2 #n2 elim n2 -n2
+[ #n1 #H elim (lt_le_false … H) -H //
+| #n2 #IH * /4 width=1 by lt_S_S_to_lt/
+]
+qed-.
+
+lemma le_elim: ∀R:relation nat.
+               (∀n2. R O (n2)) →
+               (∀n1,n2. R n1 n2 → R (⫯n1) (⫯n2)) →
+               ∀n1,n2. n1 ≤ n2 → R n1 n2.
+#R #IH1 #IH2 #n1 #n2 @(nat_elim2 … n1 n2) -n1 -n2
+/4 width=1 by monotonic_pred/ -IH1 -IH2
+#n1 #H elim (lt_le_false … H) -H //
+qed-.
+
 (* Iterators ****************************************************************)
 
 (* Note: see also: lib/arithemetics/bigops.ma *)
@@ -158,6 +210,12 @@ let rec iter (n:nat) (B:Type[0]) (op: B → B) (nil: B) ≝
 
 interpretation "iterated function" 'exp op n = (iter n ? op).
 
+lemma iter_O: ∀B:Type[0]. ∀f:B→B.∀b. f^0 b = b.
+// qed.
+
+lemma iter_S: ∀B:Type[0]. ∀f:B→B.∀b,l. f^(S l) b = f (f^l b).
+// qed.
+
 lemma iter_SO: ∀B:Type[0]. ∀f:B→B. ∀b,l. f^(l+1) b = f (f^l b).
 #B #f #b #l >commutative_plus //
 qed.