(************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) (* Logic.eq==>iff) lt. Proof. repeat red; intros; subst; auto. Qed. Definition le_lteq := le_lt_or_eq_iff. Definition compare_spec := nat_compare_spec. End Nat_as_OT. (** Note that [Nat_as_OT] can also be seen as a [UsualOrderedType] and a [OrderedType] (and also as a [DecidableType]). *) (** * An [order] tactic for Peano numbers *) Module NatOrder := OTF_to_OrderTac Nat_as_OT. Ltac nat_order := NatOrder.order. (** Note that [nat_order] is domain-agnostic: it will not prove [1<=2] or [x<=x+x], but rather things like [x<=y -> y<=x -> x=y]. *) Section Test. Let test : forall x y : nat, x<=y -> y<=x -> x=y. Proof. nat_order. Qed. End Test.