]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma
some renaming and some typos corrected ...
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / lib / arith.ma
index 1afc42ebcbccd6d0365a2869eb9035071a03916f..066d9a1f04a9e17f3d826ce4ae3e978fc270b2b8 100644 (file)
@@ -60,9 +60,13 @@ lemma arith_i: ∀x,y,z. y < x → x+z-y-1 = x-y-1+z.
 
 (* Properties ***************************************************************)
 
-axiom eq_nat_dec: ∀n1,n2:nat. Decidable (n1 = n2).
-
-axiom lt_dec: ∀n1,n2. Decidable (n1 < n2).
+lemma eq_nat_dec: ∀n1,n2:nat. Decidable (n1 = n2).
+#n1 elim n1 -n1 [| #n1 #IHn1 ] * [2,4: #n2 ]
+[1,4: @or_intror #H destruct
+| elim (IHn1 n2) -IHn1 /3 width=1 by or_intror, or_introl/
+| /2 width=1 by or_introl/
+]
+qed-.
 
 lemma lt_or_eq_or_gt: ∀m,n. ∨∨ m < n | n = m | n < m.
 #m #n elim (lt_or_ge m n) /2 width=1 by or3_intro0/
@@ -115,11 +119,6 @@ lemma lt_zero_false: ∀n. n < 0 → ⊥.
 #n #H elim (lt_to_not_le … H) -H /2 width=1 by/
 qed-.
 
-lemma false_lt_to_le: ∀x,y. (x < y → ⊥) → y ≤ x.
-#x #y #H elim (decidable_lt x y) /2 width=1 by not_lt_to_le/
-#Hxy elim (H Hxy)
-qed-.
-
 lemma pred_inv_refl: ∀m. pred m = m → m = 0.
 * // normalize #m #H elim (lt_refl_false m) //
 qed-.
@@ -141,7 +140,7 @@ lemma plus_xSy_x_false: ∀y,x. x + S y = x → ⊥.
 
 (* Iterators ****************************************************************)
 
-(* Note: see also: lib/arithemetcs/bigops.ma *)
+(* Note: see also: lib/arithemetics/bigops.ma *)
 let rec iter (n:nat) (B:Type[0]) (op: B → B) (nil: B) ≝
   match n with
    [ O   ⇒ nil