Built with Alectryon, running Coq+SerAPI v8.10.0+0.7.0. Coq sources are in this panel; goals and messages will appear in the other. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑ Ctrl+↓ to navigate, Ctrl+🖱️ to focus.
(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) (* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) (* * (see LICENSE file for the text of the license) *) (************************************************************************) (* Finite sets library. * Authors: Pierre Letouzey and Jean-Christophe Filliâtre * Institution: LRI, CNRS UMR 8623 - Université Paris Sud * 91405 Orsay, France *) Require Import Orders PeanoNat POrderedType BinNat BinInt RelationPairs EqualitiesFacts.
Ordered Type for nat, Positive, N, Z with the usual order.
Module Nat_as_OT := PeanoNat.Nat. Module Positive_as_OT := BinPos.Pos. Module N_as_OT := BinNat.N. Module Z_as_OT := BinInt.Z.
An OrderedType can now directly be seen as a DecidableType
Module OT_as_DT (O:OrderedType) <: DecidableType := O.
(Usual) Decidable Type for nat, positive, N, Z
Module Nat_as_DT <: UsualDecidableType := Nat_as_OT. Module Positive_as_DT <: UsualDecidableType := Positive_as_OT. Module N_as_DT <: UsualDecidableType := N_as_OT. Module Z_as_DT <: UsualDecidableType := Z_as_OT.
From two ordered types, we can build a new OrderedType
over their cartesian product, using the lexicographic order.
Module PairOrderedType(O1 O2:OrderedType) <: OrderedType. Include PairDecidableType O1 O2. Definition lt := (relation_disjunction (O1.lt @@1) (O1.eq * O2.lt))%signature.StrictOrder ltStrictOrder lt(* irreflexive *)Irreflexive ltTransitive ltx1:O1.tx2:O2.tO1.lt x1 x1 \/ O1.eq x1 x1 /\ O2.lt x2 x2 -> FalseTransitive ltx1:O1.tx2:O2.tH:O1.lt x1 x1Falsex1:O1.tx2:O2.tH:O1.eq x1 x1 /\ O2.lt x2 x2FalseTransitive ltx1:O1.tx2:O2.tH:O1.eq x1 x1 /\ O2.lt x2 x2FalseTransitive lt(* transitive *)Transitive ltx1:O1.tx2:O2.ty1:O1.ty2:O2.tz1:O1.tz2:O2.tlt (x1, x2) (y1, y2) -> lt (y1, y2) (z1, z2) -> lt (x1, x2) (z1, z2)x1:O1.tx2:O2.ty1:O1.ty2:O2.tz1:O1.tz2:O2.tO1.lt x1 y1 \/ O1.eq x1 y1 /\ O2.lt x2 y2 -> O1.lt y1 z1 \/ O1.eq y1 z1 /\ O2.lt y2 z2 -> O1.lt x1 z1 \/ O1.eq x1 z1 /\ O2.lt x2 z2x1:O1.tx2:O2.ty1:O1.ty2:O2.tz1:O1.tz2:O2.tH1:O1.lt x1 y1H:O1.lt y1 z1O1.lt x1 z1 \/ O1.eq x1 z1 /\ O2.lt x2 z2x1:O1.tx2:O2.ty1:O1.ty2:O2.tz1:O1.tz2:O2.tH1:O1.lt x1 y1H0:O1.eq y1 z1H2:O2.lt y2 z2O1.lt x1 z1 \/ O1.eq x1 z1 /\ O2.lt x2 z2x1:O1.tx2:O2.ty1:O1.ty2:O2.tz1:O1.tz2:O2.tH:O1.eq x1 y1H2:O2.lt x2 y2H1:O1.lt y1 z1O1.lt x1 z1 \/ O1.eq x1 z1 /\ O2.lt x2 z2x1:O1.tx2:O2.ty1:O1.ty2:O2.tz1:O1.tz2:O2.tH:O1.eq x1 y1H2:O2.lt x2 y2H0:O1.eq y1 z1H3:O2.lt y2 z2O1.lt x1 z1 \/ O1.eq x1 z1 /\ O2.lt x2 z2x1:O1.tx2:O2.ty1:O1.ty2:O2.tz1:O1.tz2:O2.tH1:O1.lt x1 y1H0:O1.eq y1 z1H2:O2.lt y2 z2O1.lt x1 z1 \/ O1.eq x1 z1 /\ O2.lt x2 z2x1:O1.tx2:O2.ty1:O1.ty2:O2.tz1:O1.tz2:O2.tH:O1.eq x1 y1H2:O2.lt x2 y2H1:O1.lt y1 z1O1.lt x1 z1 \/ O1.eq x1 z1 /\ O2.lt x2 z2x1:O1.tx2:O2.ty1:O1.ty2:O2.tz1:O1.tz2:O2.tH:O1.eq x1 y1H2:O2.lt x2 y2H0:O1.eq y1 z1H3:O2.lt y2 z2O1.lt x1 z1 \/ O1.eq x1 z1 /\ O2.lt x2 z2x1:O1.tx2:O2.ty1:O1.ty2:O2.tz1:O1.tz2:O2.tH1:O1.lt x1 y1H0:O1.eq y1 z1H2:O2.lt y2 z2O1.lt x1 z1x1:O1.tx2:O2.ty1:O1.ty2:O2.tz1:O1.tz2:O2.tH:O1.eq x1 y1H2:O2.lt x2 y2H1:O1.lt y1 z1O1.lt x1 z1 \/ O1.eq x1 z1 /\ O2.lt x2 z2x1:O1.tx2:O2.ty1:O1.ty2:O2.tz1:O1.tz2:O2.tH:O1.eq x1 y1H2:O2.lt x2 y2H0:O1.eq y1 z1H3:O2.lt y2 z2O1.lt x1 z1 \/ O1.eq x1 z1 /\ O2.lt x2 z2x1:O1.tx2:O2.ty1:O1.ty2:O2.tz1:O1.tz2:O2.tH:O1.eq x1 y1H2:O2.lt x2 y2H1:O1.lt y1 z1O1.lt x1 z1 \/ O1.eq x1 z1 /\ O2.lt x2 z2x1:O1.tx2:O2.ty1:O1.ty2:O2.tz1:O1.tz2:O2.tH:O1.eq x1 y1H2:O2.lt x2 y2H0:O1.eq y1 z1H3:O2.lt y2 z2O1.lt x1 z1 \/ O1.eq x1 z1 /\ O2.lt x2 z2right; split; etransitivity; eauto. Qed.x1:O1.tx2:O2.ty1:O1.ty2:O2.tz1:O1.tz2:O2.tH:O1.eq x1 y1H2:O2.lt x2 y2H0:O1.eq y1 z1H3:O2.lt y2 z2O1.lt x1 z1 \/ O1.eq x1 z1 /\ O2.lt x2 z2Proper (eq ==> eq ==> iff) ltProper (eq ==> eq ==> iff) ltforall x y : O1.t * O2.t, O1.eq (let (x0, _) := x in x0) (let (x0, _) := y in x0) /\ O2.eq (let (_, y0) := x in y0) (let (_, y0) := y in y0) -> forall x0 y0 : O1.t * O2.t, O1.eq (let (x1, _) := x0 in x1) (let (x1, _) := y0 in x1) /\ O2.eq (let (_, y1) := x0 in y1) (let (_, y1) := y0 in y1) -> (O1.lt (let (x1, _) := x in x1) (let (x1, _) := x0 in x1) \/ O1.eq (let (x1, _) := x in x1) (let (x1, _) := x0 in x1) /\ O2.lt (let (_, y1) := x in y1) (let (_, y1) := x0 in y1) -> O1.lt (let (x1, _) := y in x1) (let (x1, _) := y0 in x1) \/ O1.eq (let (x1, _) := y in x1) (let (x1, _) := y0 in x1) /\ O2.lt (let (_, y1) := y in y1) (let (_, y1) := y0 in y1)) /\ (O1.lt (let (x1, _) := y in x1) (let (x1, _) := y0 in x1) \/ O1.eq (let (x1, _) := y in x1) (let (x1, _) := y0 in x1) /\ O2.lt (let (_, y1) := y in y1) (let (_, y1) := y0 in y1) -> O1.lt (let (x1, _) := x in x1) (let (x1, _) := x0 in x1) \/ O1.eq (let (x1, _) := x in x1) (let (x1, _) := x0 in x1) /\ O2.lt (let (_, y1) := x in y1) (let (_, y1) := x0 in y1))rewrite X1,X2,Y1,Y2; intuition. Qed. Definition compare x y := match O1.compare (fst x) (fst y) with | Eq => O2.compare (snd x) (snd y) | Lt => Lt | Gt => Gt end.x1:O1.tx2:O2.tx1':O1.tx2':O2.tX1:O1.eq x1 x1'X2:O2.eq x2 x2'y1:O1.ty2:O2.ty1':O1.ty2':O2.tY1:O1.eq y1 y1'Y2:O2.eq y2 y2'(O1.lt x1 y1 \/ O1.eq x1 y1 /\ O2.lt x2 y2 -> O1.lt x1' y1' \/ O1.eq x1' y1' /\ O2.lt x2' y2') /\ (O1.lt x1' y1' \/ O1.eq x1' y1' /\ O2.lt x2' y2' -> O1.lt x1 y1 \/ O1.eq x1 y1 /\ O2.lt x2 y2)forall x y : O1.t * O2.t, CompSpec eq lt x y (compare x y)forall x y : O1.t * O2.t, CompSpec eq lt x y (compare x y)x1:O1.tx2:O2.ty1:O1.ty2:O2.tCompSpec eq lt (x1, x2) (y1, y2) match O1.compare x1 y1 with | Eq => O2.compare x2 y2 | Lt => Lt | Gt => Gt enddestruct (O2.compare_spec x2 y2); constructor; compute; auto with relations. Qed. End PairOrderedType.x1:O1.tx2:O2.ty1:O1.ty2:O2.tH:O1.eq x1 y1CompSpec eq lt (x1, x2) (y1, y2) (O2.compare x2 y2)
Even if positive can be seen as an ordered type with respect to the
usual order (see above), we can also use a lexicographic order over bits
(lower bits are considered first). This is more natural when using
positive as indexes for sets or maps (see MSetPositive).
Local Open Scope positive. Module PositiveOrderedTypeBits <: UsualOrderedType. Definition t:=positive. Include HasUsualEq <+ UsualIsEq. Definition eqb := Pos.eqb. Definition eqb_eq := Pos.eqb_eq. Include HasEqBool2Dec. Fixpoint bits_lt (p q:positive) : Prop := match p, q with | xH, xI _ => True | xH, _ => False | xO p, xO q => bits_lt p q | xO _, _ => True | xI p, xI q => bits_lt p q | xI _, _ => False end. Definition lt:=bits_lt.forall x : positive, ~ bits_lt x xinduction x; simpl; auto. Qed.forall x : positive, ~ bits_lt x xforall x y z : positive, bits_lt x y -> bits_lt y z -> bits_lt x zinduction x; destruct y,z; simpl; eauto; intuition. Qed.forall x y z : positive, bits_lt x y -> bits_lt y z -> bits_lt x zProper (eq ==> eq ==> iff) ltProper (eq ==> eq ==> iff) ltrewrite Hx, Hy; intuition. Qed.x, x':positiveHx:eq x x'y, y':positiveHy:eq y y'lt x y <-> lt x' y'StrictOrder ltsplit; [ exact bits_lt_antirefl | exact bits_lt_trans ]. Qed. Fixpoint compare x y := match x, y with | x~1, y~1 => compare x y | x~1, _ => Gt | x~0, y~0 => compare x y | x~0, _ => Lt | 1, y~1 => Lt | 1, 1 => Eq | 1, y~0 => Gt end.StrictOrder ltforall x y : positive, CompSpec eq lt x y (compare x y)forall x y : positive, CompSpec eq lt x y (compare x y)forall x y : positive, CompSpec Logic.eq bits_lt x y (compare x y)x:positiveIHx:forall y0 : positive, CompSpec Logic.eq bits_lt x y0 (compare x y0)y:positiveCompSpec Logic.eq bits_lt x~1 y~1 (compare x y)x:positiveIHx:forall y0 : positive, CompSpec Logic.eq bits_lt x y0 (compare x y0)y:positiveCompSpec Logic.eq bits_lt x~0 y~0 (compare x y)destruct (IHx y); subst; auto. Qed. End PositiveOrderedTypeBits.x:positiveIHx:forall y0 : positive, CompSpec Logic.eq bits_lt x y0 (compare x y0)y:positiveCompSpec Logic.eq bits_lt x~0 y~0 (compare x y)