(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) (* le y z -> le x z]. *) Inductive ord : Set := OEQ | OLT | OLE. Definition trans_ord o o' := match o, o' with | OEQ, _ => o' | _, OEQ => o | OLE, OLE => OLE | _, _ => OLT end. Local Infix "+" := trans_ord. (** ** The tactic requirements : a total order We need : - an equivalence [eq], - a strict order [lt] total and compatible with [eq], - a larger order [le] synonym for [lt\/eq]. This used to be provided here via a [TotalOrder], but for technical reasons related to extraction, we now ask for two sperate parts: relations in a [EqLtLe] + properties in [IsTotalOrder]. Note that [TotalOrder = EqLtLe <+ IsTotalOrder] *) Module Type IsTotalOrder (O:EqLtLe) := IsEq O <+ IsStrOrder O <+ LeIsLtEq O <+ LtIsTotal O. (** ** Properties that will be used by the [order] tactic *) Module OrderFacts (Import O:EqLtLe)(P:IsTotalOrder O). Include EqLtLeNotation O. (** Reflexivity rules *) Lemma eq_refl : forall x, x==x. Proof. reflexivity. Qed. Lemma le_refl : forall x, x<=x. Proof. intros; rewrite P.le_lteq; right; reflexivity. Qed. Lemma lt_irrefl : forall x, ~ x y==x. Proof. auto with *. Qed. Lemma le_antisym : forall x y, x<=y -> y<=x -> x==y. Proof. intros x y; rewrite 2 P.le_lteq. intuition. elim (StrictOrder_Irreflexive x); transitivity y; auto. Qed. Lemma neq_sym : forall x y, ~x==y -> ~y==x. Proof. auto using eq_sym. Qed. (** Transitivity rules : first, a generic formulation, then instances*) Ltac subst_eqns := match goal with | H : _==_ |- _ => (rewrite H || rewrite <- H); clear H; subst_eqns | _ => idtac end. Definition interp_ord o := match o with OEQ => O.eq | OLT => O.lt | OLE => O.le end. Local Notation "#" := interp_ord. Lemma trans : forall o o' x y z, #o x y -> #o' y z -> #(o+o') x z. Proof. destruct o, o'; simpl; intros x y z; rewrite ?P.le_lteq; intuition auto; subst_eqns; eauto using (StrictOrder_Transitive x y z) with *. Qed. Definition eq_trans x y z : x==y -> y==z -> x==z := @trans OEQ OEQ x y z. Definition le_trans x y z : x<=y -> y<=z -> x<=z := @trans OLE OLE x y z. Definition lt_trans x y z : x y x y x y<=z -> x y x y==z -> x y<=z -> x<=z := @trans OEQ OLE x y z. Definition le_eq x y z : x<=y -> y==z -> x<=z := @trans OLE OEQ x y z. Lemma eq_neq : forall x y z, x==y -> ~y==z -> ~x==z. Proof. eauto using eq_trans, eq_sym. Qed. Lemma neq_eq : forall x y z, ~x==y -> y==z -> ~x==z. Proof. eauto using eq_trans, eq_sym. Qed. (** (double) negation rules *) Lemma not_neq_eq : forall x y, ~~x==y -> x==y. Proof. intros x y H. destruct (P.lt_total x y) as [H'|[H'|H']]; auto; destruct H; intro H; rewrite H in H'; eapply lt_irrefl; eauto. Qed. Lemma not_ge_lt : forall x y, ~y<=x -> x x<=y. Proof. intros x y H. rewrite P.le_lteq. generalize (P.lt_total x y); intuition. Qed. Lemma le_neq_lt : forall x y, x<=y -> ~x==y -> x