]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/ground_2/arith.ma
more symbols added for lambda_delta
[helm.git] / matita / matita / contribs / lambda_delta / ground_2 / arith.ma
index 0bdcd748f98ea9e0eacfef97f9db67c0e1180ca2..c51873baa09bb9d8f0c46e2f502e378cb05af921 100644 (file)
@@ -19,6 +19,9 @@ include "ground_2/star.ma".
 
 (* Equations ****************************************************************)
 
+lemma plus_n_2: ∀n. n + 2 = n + 1 + 1.
+// qed.
+
 lemma le_plus_minus: ∀m,n,p. p ≤ n → m + n - p = m + (n - p).
 /2 by plus_minus/ qed.
 
@@ -53,24 +56,34 @@ lemma lt_or_eq_or_gt: ∀m,n. ∨∨ m < n | n = m | n < m.
 #m #Hm * #H /2 width=1/ /3 width=1/
 qed-.
 
-lemma lt_refl_false: ∀n. n < n → False.
+lemma lt_refl_false: ∀n. n < n → .
 #n #H elim (lt_to_not_eq … H) -H /2 width=1/
 qed-.
 
-lemma lt_zero_false: ∀n. n < 0 → False.
+lemma lt_zero_false: ∀n. n < 0 → .
 #n #H elim (lt_to_not_le … H) -H /2 width=1/
 qed-.
 
-lemma false_lt_to_le: ∀x,y. (x < y → False) → y ≤ x.
+lemma false_lt_to_le: ∀x,y. (x < y → ) → y ≤ x.
 #x #y #H elim (decidable_lt x y) /2 width=1/
 #Hxy elim (H Hxy)
 qed-.
 
-(*
-lemma pippo: ∀x,y,z. x < z → y < z - x → x + y < z.
-/3 width=2/
+(* iterators ****************************************************************)
 
-lemma le_or_ge: ∀m,n. m ≤ n ∨ n ≤ m.
-#m #n elim (lt_or_ge m n) /2 width=1/ /3 width=2/
-qed-.
-*)
+(* Note: see also: lib/arithemetcs/bigops.ma *)
+let rec iter (n:nat) (B:Type[0]) (op: B → B) (nil: B) ≝
+  match n with
+   [ O   ⇒ nil
+   | S k ⇒ op (iter k B op nil)
+   ].
+
+interpretation "iterated function" 'exp op n = (iter n ? op).
+
+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.
+
+lemma iter_n_Sm: ∀B:Type[0]. ∀f:B→B. ∀b,l. f^l (f b) = f (f^l b).
+#B #f #b #l elim l -l normalize //
+qed.