From 97fefe1fcca363a1317e066e7f4b99b9c1e9987b Mon Sep 17 00:00:00 2001 From: Stephane Glondu Date: Thu, 12 Jan 2012 16:02:20 +0100 Subject: Imported Upstream version 8.4~beta --- theories/MSets/MSetAVL.v | 15 ++- theories/MSets/MSetDecide.v | 2 - theories/MSets/MSetEqProperties.v | 2 - theories/MSets/MSetFacts.v | 2 - theories/MSets/MSetInterface.v | 219 +++++++++++++++++++++++++++++++++++++- theories/MSets/MSetList.v | 5 +- theories/MSets/MSetProperties.v | 27 ++--- theories/MSets/MSetToFiniteSet.v | 2 - theories/MSets/MSetWeakList.v | 4 +- theories/MSets/MSets.v | 2 - 10 files changed, 237 insertions(+), 43 deletions(-) (limited to 'theories/MSets') diff --git a/theories/MSets/MSetAVL.v b/theories/MSets/MSetAVL.v index 96580749..bdada486 100644 --- a/theories/MSets/MSetAVL.v +++ b/theories/MSets/MSetAVL.v @@ -7,8 +7,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (***********************************************************************) -(* $Id$ *) - (** * MSetAVL : Implementation of MSetInterface via AVL trees *) (** This module implements finite sets using AVL trees. @@ -48,6 +46,7 @@ Local Open Scope Int_scope. Local Open Scope lazy_bool_scope. Definition elt := X.t. +Hint Transparent elt. (** ** Trees @@ -376,7 +375,7 @@ Fixpoint fold (A : Type) (f : elt -> A -> A)(s : t) : A -> A := | Leaf => a | Node l x r _ => fold f r (f x (fold f l a)) end. -Implicit Arguments fold [A]. +Arguments fold [A] f s _. (** ** Subset *) @@ -877,12 +876,12 @@ Open Scope Int_scope. Ltac join_tac := intro l; induction l as [| ll _ lx lr Hlr lh]; [ | intros x r; induction r as [| rl Hrl rx rr _ rh]; unfold join; - [ | destruct (gt_le_dec lh (rh+2)); + [ | destruct (gt_le_dec lh (rh+2)) as [GT|LE]; [ match goal with |- context b [ bal ?a ?b ?c] => replace (bal a b c) with (bal ll lx (join lr x (Node rl rx rr rh))); [ | auto] end - | destruct (gt_le_dec rh (lh+2)); + | destruct (gt_le_dec rh (lh+2)) as [GT'|LE']; [ match goal with |- context b [ bal ?a ?b ?c] => replace (bal a b c) with (bal (join (Node ll lx lr lh) x rl) rx rr); [ | auto] @@ -905,7 +904,7 @@ Instance join_ok : forall l x r `(Ok l, Ok r, lt_tree x l, gt_tree x r), Ok (join l x r). Proof. join_tac; auto with *; inv; apply bal_ok; auto; - clear Hrl Hlr z; intro; intros; rewrite join_spec in *. + clear Hrl Hlr; intro; intros; rewrite join_spec in *. intuition; [ setoid_replace y with x | ]; eauto. intuition; [ setoid_replace y with x | ]; eauto. Qed. @@ -1691,7 +1690,7 @@ Proof. Qed. Definition lt (s1 s2 : t) : Prop := - exists s1', exists s2', Ok s1' /\ Ok s2' /\ eq s1 s1' /\ eq s2 s2' + exists s1' s2', Ok s1' /\ Ok s2' /\ eq s1 s1' /\ eq s2 s2' /\ L.lt (elements s1') (elements s2'). Instance lt_strorder : StrictOrder lt. @@ -1768,7 +1767,7 @@ Lemma compare_more_Cmp : forall x1 cont x2 r2 e2 l, Cmp (compare_more x1 cont (More x2 r2 e2)) (x1::l) (flatten_e (More x2 r2 e2)). Proof. - simpl; intros; elim_compare x1 x2; simpl; auto. + simpl; intros; elim_compare x1 x2; simpl; red; auto. Qed. Lemma compare_cont_Cmp : forall s1 cont e2 l, diff --git a/theories/MSets/MSetDecide.v b/theories/MSets/MSetDecide.v index 4ec050bd..eefd2951 100644 --- a/theories/MSets/MSetDecide.v +++ b/theories/MSets/MSetDecide.v @@ -6,8 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (***********************************************************************) -(* $Id$ *) - (**************************************************************) (* MSetDecide.v *) (* *) diff --git a/theories/MSets/MSetEqProperties.v b/theories/MSets/MSetEqProperties.v index fe6c3c79..2e7da404 100644 --- a/theories/MSets/MSetEqProperties.v +++ b/theories/MSets/MSetEqProperties.v @@ -6,8 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (***********************************************************************) -(* $Id$ *) - (** * Finite sets library *) (** This module proves many properties of finite sets that diff --git a/theories/MSets/MSetFacts.v b/theories/MSets/MSetFacts.v index 6d38b696..4e17618f 100644 --- a/theories/MSets/MSetFacts.v +++ b/theories/MSets/MSetFacts.v @@ -6,8 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (***********************************************************************) -(* $Id$ *) - (** * Finite sets library *) (** This functor derives additional facts from [MSetInterface.S]. These diff --git a/theories/MSets/MSetInterface.v b/theories/MSets/MSetInterface.v index 194cb904..f2b908af 100644 --- a/theories/MSets/MSetInterface.v +++ b/theories/MSets/MSetInterface.v @@ -6,8 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (***********************************************************************) -(* $Id$ *) - (** * Finite set library *) (** Set interfaces, inspired by the one of Ocaml. When compared with @@ -439,7 +437,7 @@ Module WRaw2SetsOn (E:DecidableType)(M:WRawSets E) <: WSetsOn E. Record t_ := Mkt {this :> M.t; is_ok : M.Ok this}. Definition t := t_. - Implicit Arguments Mkt [ [is_ok] ]. + Arguments Mkt this {is_ok}. Hint Resolve is_ok : typeclass_instances. Definition In (x : elt)(s : t) := M.In x s.(this). @@ -653,7 +651,218 @@ Module Raw2Sets (O:OrderedType)(M:RawSets O) <: Sets with Module E := O. End Raw2Sets. -(** We provide an ordering for sets-as-sorted-lists *) +(** It is in fact possible to provide an ordering on sets with + very little information on them (more or less only the [In] + predicate). This generic build of ordering is in fact not + used for the moment, we rather use a simplier version + dedicated to sets-as-sorted-lists, see [MakeListOrdering]. +*) + +Module Type IN (O:OrderedType). + Parameter Inline t : Type. + Parameter Inline In : O.t -> t -> Prop. + Declare Instance In_compat : Proper (O.eq==>eq==>iff) In. + Definition Equal s s' := forall x, In x s <-> In x s'. + Definition Empty s := forall x, ~In x s. +End IN. + +Module MakeSetOrdering (O:OrderedType)(Import M:IN O). + Module Import MO := OrderedTypeFacts O. + + Definition eq : t -> t -> Prop := Equal. + + Instance eq_equiv : Equivalence eq. + Proof. firstorder. Qed. + + Instance : Proper (O.eq==>eq==>iff) In. + Proof. + intros x x' Ex s s' Es. rewrite Ex. apply Es. + Qed. + + Definition Below x s := forall y, In y s -> O.lt y x. + Definition Above x s := forall y, In y s -> O.lt x y. + + Definition EquivBefore x s s' := + forall y, O.lt y x -> (In y s <-> In y s'). + + Definition EmptyBetween x y s := + forall z, In z s -> O.lt z y -> O.lt z x. + + Definition lt s s' := exists x, EquivBefore x s s' /\ + ((In x s' /\ Below x s) \/ + (In x s /\ exists y, In y s' /\ O.lt x y /\ EmptyBetween x y s')). + + Instance : Proper (O.eq==>eq==>eq==>iff) EquivBefore. + Proof. + unfold EquivBefore. intros x x' E s1 s1' E1 s2 s2' E2. + setoid_rewrite E; setoid_rewrite E1; setoid_rewrite E2; intuition. + Qed. + + Instance : Proper (O.eq==>eq==>iff) Below. + Proof. + unfold Below. intros x x' Ex s s' Es. + setoid_rewrite Ex; setoid_rewrite Es; intuition. + Qed. + + Instance : Proper (O.eq==>eq==>iff) Above. + Proof. + unfold Above. intros x x' Ex s s' Es. + setoid_rewrite Ex; setoid_rewrite Es; intuition. + Qed. + + Instance : Proper (O.eq==>O.eq==>eq==>iff) EmptyBetween. + Proof. + unfold EmptyBetween. intros x x' Ex y y' Ey s s' Es. + setoid_rewrite Ex; setoid_rewrite Ey; setoid_rewrite Es; intuition. + Qed. + + Instance lt_compat : Proper (eq==>eq==>iff) lt. + Proof. + unfold lt. intros s1 s1' E1 s2 s2' E2. + setoid_rewrite E1; setoid_rewrite E2; intuition. + Qed. + + Instance lt_strorder : StrictOrder lt. + Proof. + split. + (* irreflexive *) + intros s (x & _ & [(IN,Em)|(IN & y & IN' & LT & Be)]). + specialize (Em x IN); order. + specialize (Be x IN LT); order. + (* transitive *) + intros s1 s2 s3 (x & EQ & [(IN,Pre)|(IN,Lex)]) + (x' & EQ' & [(IN',Pre')|(IN',Lex')]). + (* 1) Pre / Pre --> Pre *) + assert (O.lt x x') by (specialize (Pre' x IN); auto). + exists x; split. + intros y Hy; rewrite <- (EQ' y); auto; order. + left; split; auto. + rewrite <- (EQ' x); auto. + (* 2) Pre / Lex *) + elim_compare x x'. + (* 2a) x=x' --> Pre *) + destruct Lex' as (y & INy & LT & Be). + exists y; split. + intros z Hz. split; intros INz. + specialize (Pre z INz). rewrite <- (EQ' z), <- (EQ z); auto; order. + specialize (Be z INz Hz). rewrite (EQ z), (EQ' z); auto; order. + left; split; auto. + intros z Hz. transitivity x; auto; order. + (* 2b) x Pre *) + exists x; split. + intros z Hz. rewrite <- (EQ' z) by order; auto. + left; split; auto. + rewrite <- (EQ' x); auto. + (* 2c) x>x' --> Lex *) + exists x'; split. + intros z Hz. rewrite (EQ z) by order; auto. + right; split; auto. + rewrite (EQ x'); auto. + (* 3) Lex / Pre --> Lex *) + destruct Lex as (y & INy & LT & Be). + specialize (Pre' y INy). + exists x; split. + intros z Hz. rewrite <- (EQ' z) by order; auto. + right; split; auto. + exists y; repeat split; auto. + rewrite <- (EQ' y); auto. + intros z Hz LTz; apply Be; auto. rewrite (EQ' z); auto; order. + (* 4) Lex / Lex *) + elim_compare x x'. + (* 4a) x=x' --> impossible *) + destruct Lex as (y & INy & LT & Be). + setoid_replace x with x' in LT; auto. + specialize (Be x' IN' LT); order. + (* 4b) x Lex *) + exists x; split. + intros z Hz. rewrite <- (EQ' z) by order; auto. + right; split; auto. + destruct Lex as (y & INy & LT & Be). + elim_compare y x'. + (* 4ba *) + destruct Lex' as (y' & Iny' & LT' & Be'). + exists y'; repeat split; auto. order. + intros z Hz LTz. specialize (Be' z Hz LTz). + rewrite <- (EQ' z) in Hz by order. + apply Be; auto. order. + (* 4bb *) + exists y; repeat split; auto. + rewrite <- (EQ' y); auto. + intros z Hz LTz. apply Be; auto. rewrite (EQ' z); auto; order. + (* 4bc*) + assert (O.lt x' x) by auto. order. + (* 4c) x>x' --> Lex *) + exists x'; split. + intros z Hz. rewrite (EQ z) by order; auto. + right; split; auto. + rewrite (EQ x'); auto. + Qed. + + Lemma lt_empty_r : forall s s', Empty s' -> ~ lt s s'. + Proof. + intros s s' Hs' (x & _ & [(IN,_)|(_ & y & IN & _)]). + elim (Hs' x IN). + elim (Hs' y IN). + Qed. + + Definition Add x s s' := forall y, In y s' <-> O.eq x y \/ In y s. + + Lemma lt_empty_l : forall x s1 s2 s2', + Empty s1 -> Above x s2 -> Add x s2 s2' -> lt s1 s2'. + Proof. + intros x s1 s2 s2' Em Ab Ad. + exists x; split. + intros y Hy; split; intros IN. + elim (Em y IN). + rewrite (Ad y) in IN; destruct IN as [EQ|IN]. order. + specialize (Ab y IN). order. + left; split. + rewrite (Ad x). now left. + intros y Hy. elim (Em y Hy). + Qed. + + Lemma lt_add_lt : forall x1 x2 s1 s1' s2 s2', + Above x1 s1 -> Above x2 s2 -> Add x1 s1 s1' -> Add x2 s2 s2' -> + O.lt x1 x2 -> lt s1' s2'. + Proof. + intros x1 x2 s1 s1' s2 s2' Ab1 Ab2 Ad1 Ad2 LT. + exists x1; split; [ | right; split]; auto. + intros y Hy. rewrite (Ad1 y), (Ad2 y). + split; intros [U|U]; try order. + specialize (Ab1 y U). order. + specialize (Ab2 y U). order. + rewrite (Ad1 x1); auto with *. + exists x2; repeat split; auto. + rewrite (Ad2 x2); now left. + intros y. rewrite (Ad2 y). intros [U|U]. order. + specialize (Ab2 y U). order. + Qed. + + Lemma lt_add_eq : forall x1 x2 s1 s1' s2 s2', + Above x1 s1 -> Above x2 s2 -> Add x1 s1 s1' -> Add x2 s2 s2' -> + O.eq x1 x2 -> lt s1 s2 -> lt s1' s2'. + Proof. + intros x1 x2 s1 s1' s2 s2' Ab1 Ab2 Ad1 Ad2 Hx (x & EQ & Disj). + assert (O.lt x1 x). + destruct Disj as [(IN,_)|(IN,_)]; auto. rewrite Hx; auto. + exists x; split. + intros z Hz. rewrite (Ad1 z), (Ad2 z). + split; intros [U|U]; try (left; order); right. + rewrite <- (EQ z); auto. + rewrite (EQ z); auto. + destruct Disj as [(IN,Em)|(IN & y & INy & LTy & Be)]. + left; split; auto. + rewrite (Ad2 x); auto. + intros z. rewrite (Ad1 z); intros [U|U]; try specialize (Ab1 z U); auto; order. + right; split; auto. + rewrite (Ad1 x); auto. + exists y; repeat split; auto. + rewrite (Ad2 y); auto. + intros z. rewrite (Ad2 z). intros [U|U]; try specialize (Ab2 z U); auto; order. + Qed. + +End MakeSetOrdering. + Module MakeListOrdering (O:OrderedType). Module MO:=OrderedTypeFacts O. @@ -663,7 +872,7 @@ Module MakeListOrdering (O:OrderedType). Definition eq s s' := forall x, In x s <-> In x s'. - Instance eq_equiv : Equivalence eq. + Instance eq_equiv : Equivalence eq := _. Inductive lt_list : t -> t -> Prop := | lt_nil : forall x s, lt_list nil (x :: s) diff --git a/theories/MSets/MSetList.v b/theories/MSets/MSetList.v index 48af7e6a..bcf68f1d 100644 --- a/theories/MSets/MSetList.v +++ b/theories/MSets/MSetList.v @@ -6,8 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (***********************************************************************) -(* $Id$ *) - (** * Finite sets library *) (** This file proposes an implementation of the non-dependant @@ -788,8 +786,7 @@ Module MakeRaw (X: OrderedType) <: RawSets X. Definition eq := L.eq. Definition eq_equiv := L.eq_equiv. Definition lt l1 l2 := - exists l1', exists l2', Ok l1' /\ Ok l2' /\ - eq l1 l1' /\ eq l2 l2' /\ L.lt l1' l2'. + exists l1' l2', Ok l1' /\ Ok l2' /\ eq l1 l1' /\ eq l2 l2' /\ L.lt l1' l2'. Instance lt_strorder : StrictOrder lt. Proof. diff --git a/theories/MSets/MSetProperties.v b/theories/MSets/MSetProperties.v index c0038a4f..0f24d76a 100644 --- a/theories/MSets/MSetProperties.v +++ b/theories/MSets/MSetProperties.v @@ -6,8 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (***********************************************************************) -(* $Id$ *) - (** * Finite sets library *) (** This functor derives additional properties from [MSetInterface.S]. @@ -339,6 +337,14 @@ Module WPropertiesOn (Import E : DecidableType)(M : WSetsOn E). Notation NoDup := (NoDupA E.eq). Notation InA := (InA E.eq). + (** Alternative specification via [fold_right] *) + + Lemma fold_spec_right (s:t)(A:Type)(i:A)(f : elt -> A -> A) : + fold f s i = List.fold_right f i (rev (elements s)). + Proof. + rewrite fold_spec. symmetry. apply fold_left_rev_right. + Qed. + (** ** Induction principles for fold (contributed by S. Lescuyer) *) (** In the following lemma, the step hypothesis is deliberately restricted @@ -352,8 +358,7 @@ Module WPropertiesOn (Import E : DecidableType)(M : WSetsOn E). P s (fold f s i). Proof. intros A P f i s Pempty Pstep. - rewrite fold_1; unfold flip; rewrite <- fold_left_rev_right. - set (l:=rev (elements s)). + rewrite fold_spec_right. set (l:=rev (elements s)). assert (Pstep' : forall x a s' s'', InA x l -> ~In x s' -> Add x s' s'' -> P s' a -> P s'' (f x a)). intros; eapply Pstep; eauto. @@ -425,8 +430,7 @@ Module WPropertiesOn (Import E : DecidableType)(M : WSetsOn E). R (fold f s i) (fold g s j). Proof. intros A B R f g i j s Rempty Rstep. - do 2 (rewrite fold_1; unfold flip; rewrite <- fold_left_rev_right). - set (l:=rev (elements s)). + rewrite 2 fold_spec_right. set (l:=rev (elements s)). assert (Rstep' : forall x a b, InA x l -> R a b -> R (f x a) (g x b)) by (intros; apply Rstep; auto; rewrite elements_iff, <- InA_rev; auto with *). clearbody l; clear Rstep s. @@ -484,8 +488,7 @@ Module WPropertiesOn (Import E : DecidableType)(M : WSetsOn E). split; intros. rewrite elements_iff; do 2 rewrite InA_alt. split; destruct 1; generalize (In_rev (elements s) x0); exists x0; intuition. - rewrite fold_left_rev_right. - apply fold_1. + apply fold_spec_right. Qed. (** An alternate (and previous) specification for [fold] was based on @@ -1095,8 +1098,7 @@ Module OrdProperties (M:Sets). Above x s -> Add x s s' -> eqA (fold f s' i) (f x (fold f s i)). Proof. intros. - rewrite !FM.fold_1. - unfold flip; rewrite <-!fold_left_rev_right. + rewrite 2 fold_spec_right. change (f x (fold_right f i (rev (elements s)))) with (fold_right f i (rev (x::nil)++rev (elements s))). apply (@fold_right_eqlistA E.t E.eq A eqA st); auto with *. @@ -1112,7 +1114,7 @@ Module OrdProperties (M:Sets). Below x s -> Add x s s' -> eqA (fold f s' i) (fold f s (f x i)). Proof. intros. - rewrite !FM.fold_1. + rewrite !fold_spec. change (eqA (fold_left (flip f) (elements s') i) (fold_left (flip f) (x::elements s) i)). unfold flip; rewrite <-!fold_left_rev_right. @@ -1133,8 +1135,7 @@ Module OrdProperties (M:Sets). forall i s s', s[=]s' -> eqA (fold f s i) (fold f s' i). Proof. intros. - rewrite !FM.fold_1. - unfold flip; rewrite <- !fold_left_rev_right. + rewrite 2 fold_spec_right. apply (@fold_right_eqlistA E.t E.eq A eqA st); auto. apply eqlistA_rev. apply sort_equivlistA_eqlistA; auto with set. diff --git a/theories/MSets/MSetToFiniteSet.v b/theories/MSets/MSetToFiniteSet.v index f0b964cf..e8087bc5 100644 --- a/theories/MSets/MSetToFiniteSet.v +++ b/theories/MSets/MSetToFiniteSet.v @@ -6,8 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (***********************************************************************) -(* $Id$ *) - (** * Finite sets library : conversion to old [Finite_sets] *) Require Import Ensembles Finite_sets. diff --git a/theories/MSets/MSetWeakList.v b/theories/MSets/MSetWeakList.v index 945cb2dd..76f09c76 100644 --- a/theories/MSets/MSetWeakList.v +++ b/theories/MSets/MSetWeakList.v @@ -6,8 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (***********************************************************************) -(* $Id$ *) - (** * Finite sets library *) (** This file proposes an implementation of the non-dependant @@ -517,7 +515,7 @@ Module MakeRaw (X:DecidableType) <: WRawSets X. Definition In := InA X.eq. Definition eq := Equal. - Instance eq_equiv : Equivalence eq. + Instance eq_equiv : Equivalence eq := _. End MakeRaw. diff --git a/theories/MSets/MSets.v b/theories/MSets/MSets.v index 958e9861..f179bcd1 100644 --- a/theories/MSets/MSets.v +++ b/theories/MSets/MSets.v @@ -6,8 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (***********************************************************************) -(* $Id$ *) - Require Export Orders. Require Export OrdersEx. Require Export OrdersAlt. -- cgit v1.2.3