interpretation "natural 'not less than'" 'nless x y = (Not (lt x y)).
+(* nlemma eq_lt: ∀n,m. (n < m) = (S n ≤ m).
+//; nqed. *)
+
ndefinition ge: nat \to nat \to Prop \def
\lambda n,m:nat.m \leq n.
ntheorem trans_le: \forall n,m,p:nat. n \leq m \to m \leq p \to n \leq p
\def transitive_le. *)
-ntheorem transitive_lt: transitive nat lt.
-#a; #b; #c; #ltab; #ltbc;nelim ltbc;/2/;nqed.
+
+naxiom transitive_lt: transitive nat lt.
+(* #a; #b; #c; #ltab; #ltbc;nelim ltbc;/2/;nqed.*)
(*
theorem trans_lt: \forall n,m,p:nat. lt n m \to lt m p \to lt n p
//; nqed.
ntheorem le_plus_to_le: ∀a,n,m. a + n ≤ a + m → n ≤ m.
-#a; nelim a; /3/; nqed.
+#a; nelim a; nnormalize; /3/; nqed.
ntheorem le_plus_to_le_r: ∀a,n,m. n + a ≤ m +a → n ≤ m.
/2/; nqed.