diff options
author | Stephane Glondu <steph@glondu.net> | 2010-07-21 09:46:51 +0200 |
---|---|---|
committer | Stephane Glondu <steph@glondu.net> | 2010-07-21 09:46:51 +0200 |
commit | 5b7eafd0f00a16d78f99a27f5c7d5a0de77dc7e6 (patch) | |
tree | 631ad791a7685edafeb1fb2e8faeedc8379318ae /theories/MSets | |
parent | da178a880e3ace820b41d38b191d3785b82991f5 (diff) |
Imported Upstream snapshot 8.3~beta0+13298
Diffstat (limited to 'theories/MSets')
-rw-r--r-- | theories/MSets/MSetAVL.v | 1842 | ||||
-rw-r--r-- | theories/MSets/MSetDecide.v | 880 | ||||
-rw-r--r-- | theories/MSets/MSetEqProperties.v | 936 | ||||
-rw-r--r-- | theories/MSets/MSetFacts.v | 528 | ||||
-rw-r--r-- | theories/MSets/MSetInterface.v | 732 | ||||
-rw-r--r-- | theories/MSets/MSetList.v | 899 | ||||
-rw-r--r-- | theories/MSets/MSetPositive.v | 1149 | ||||
-rw-r--r-- | theories/MSets/MSetProperties.v | 1176 | ||||
-rw-r--r-- | theories/MSets/MSetToFiniteSet.v | 158 | ||||
-rw-r--r-- | theories/MSets/MSetWeakList.v | 533 | ||||
-rw-r--r-- | theories/MSets/MSets.v | 23 | ||||
-rw-r--r-- | theories/MSets/vo.itarget | 11 |
12 files changed, 8867 insertions, 0 deletions
diff --git a/theories/MSets/MSetAVL.v b/theories/MSets/MSetAVL.v new file mode 100644 index 00000000..c41df7c2 --- /dev/null +++ b/theories/MSets/MSetAVL.v @@ -0,0 +1,1842 @@ +(* -*- coding: utf-8 -*- *) +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) +(* \VV/ *************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(***********************************************************************) + +(* $Id$ *) + +(** * MSetAVL : Implementation of MSetInterface via AVL trees *) + +(** This module implements finite sets using AVL trees. + It follows the implementation from Ocaml's standard library, + + All operations given here expect and produce well-balanced trees + (in the ocaml sense: heigths of subtrees shouldn't differ by more + than 2), and hence has low complexities (e.g. add is logarithmic + in the size of the set). But proving these balancing preservations + is in fact not necessary for ensuring correct operational behavior + and hence fulfilling the MSet interface. As a consequence, + balancing results are not part of this file anymore, they can + now be found in [MSetFullAVL]. + + Four operations ([union], [subset], [compare] and [equal]) have + been slightly adapted in order to have only structural recursive + calls. The precise ocaml versions of these operations have also + been formalized (thanks to Function+measure), see [ocaml_union], + [ocaml_subset], [ocaml_compare] and [ocaml_equal] in + [MSetFullAVL]. The structural variants compute faster in Coq, + whereas the other variants produce nicer and/or (slightly) faster + code after extraction. +*) + +Require Import MSetInterface ZArith Int. + +Set Implicit Arguments. +Unset Strict Implicit. +(* for nicer extraction, we create only logical inductive principles *) +Local Unset Elimination Schemes. +Local Unset Case Analysis Schemes. + +(** * Ops : the pure functions *) + +Module Ops (Import I:Int)(X:OrderedType) <: WOps X. +Local Open Scope Int_scope. +Local Open Scope lazy_bool_scope. + +Definition elt := X.t. + +(** ** Trees + + The fourth field of [Node] is the height of the tree *) + +Inductive tree := + | Leaf : tree + | Node : tree -> X.t -> tree -> int -> tree. + +Definition t := tree. + +(** ** Basic functions on trees: height and cardinal *) + +Definition height (s : t) : int := + match s with + | Leaf => 0 + | Node _ _ _ h => h + end. + +Fixpoint cardinal (s : t) : nat := + match s with + | Leaf => 0%nat + | Node l _ r _ => S (cardinal l + cardinal r) + end. + +(** ** Empty Set *) + +Definition empty := Leaf. + +(** ** Emptyness test *) + +Definition is_empty s := + match s with Leaf => true | _ => false end. + +(** ** Appartness *) + +(** The [mem] function is deciding appartness. It exploits the + binary search tree invariant to achieve logarithmic complexity. *) + +Fixpoint mem x s := + match s with + | Leaf => false + | Node l y r _ => match X.compare x y with + | Lt => mem x l + | Eq => true + | Gt => mem x r + end + end. + +(** ** Singleton set *) + +Definition singleton x := Node Leaf x Leaf 1. + +(** ** Helper functions *) + +(** [create l x r] creates a node, assuming [l] and [r] + to be balanced and [|height l - height r| <= 2]. *) + +Definition create l x r := + Node l x r (max (height l) (height r) + 1). + +(** [bal l x r] acts as [create], but performs one step of + rebalancing if necessary, i.e. assumes [|height l - height r| <= 3]. *) + +Definition assert_false := create. + +Definition bal l x r := + let hl := height l in + let hr := height r in + if gt_le_dec hl (hr+2) then + match l with + | Leaf => assert_false l x r + | Node ll lx lr _ => + if ge_lt_dec (height ll) (height lr) then + create ll lx (create lr x r) + else + match lr with + | Leaf => assert_false l x r + | Node lrl lrx lrr _ => + create (create ll lx lrl) lrx (create lrr x r) + end + end + else + if gt_le_dec hr (hl+2) then + match r with + | Leaf => assert_false l x r + | Node rl rx rr _ => + if ge_lt_dec (height rr) (height rl) then + create (create l x rl) rx rr + else + match rl with + | Leaf => assert_false l x r + | Node rll rlx rlr _ => + create (create l x rll) rlx (create rlr rx rr) + end + end + else + create l x r. + +(** ** Insertion *) + +Fixpoint add x s := match s with + | Leaf => Node Leaf x Leaf 1 + | Node l y r h => + match X.compare x y with + | Lt => bal (add x l) y r + | Eq => Node l y r h + | Gt => bal l y (add x r) + end + end. + +(** ** Join + + Same as [bal] but does not assume anything regarding heights + of [l] and [r]. +*) + +Fixpoint join l : elt -> t -> t := + match l with + | Leaf => add + | Node ll lx lr lh => fun x => + fix join_aux (r:t) : t := match r with + | Leaf => add x l + | Node rl rx rr rh => + if gt_le_dec lh (rh+2) then bal ll lx (join lr x r) + else if gt_le_dec rh (lh+2) then bal (join_aux rl) rx rr + else create l x r + end + end. + +(** ** Extraction of minimum element + + Morally, [remove_min] is to be applied to a non-empty tree + [t = Node l x r h]. Since we can't deal here with [assert false] + for [t=Leaf], we pre-unpack [t] (and forget about [h]). +*) + +Fixpoint remove_min l x r : t*elt := + match l with + | Leaf => (r,x) + | Node ll lx lr lh => + let (l',m) := remove_min ll lx lr in (bal l' x r, m) + end. + +(** ** Merging two trees + + [merge t1 t2] builds the union of [t1] and [t2] assuming all elements + of [t1] to be smaller than all elements of [t2], and + [|height t1 - height t2| <= 2]. +*) + +Definition merge s1 s2 := match s1,s2 with + | Leaf, _ => s2 + | _, Leaf => s1 + | _, Node l2 x2 r2 h2 => + let (s2',m) := remove_min l2 x2 r2 in bal s1 m s2' +end. + +(** ** Deletion *) + +Fixpoint remove x s := match s with + | Leaf => Leaf + | Node l y r h => + match X.compare x y with + | Lt => bal (remove x l) y r + | Eq => merge l r + | Gt => bal l y (remove x r) + end + end. + +(** ** Minimum element *) + +Fixpoint min_elt s := match s with + | Leaf => None + | Node Leaf y _ _ => Some y + | Node l _ _ _ => min_elt l +end. + +(** ** Maximum element *) + +Fixpoint max_elt s := match s with + | Leaf => None + | Node _ y Leaf _ => Some y + | Node _ _ r _ => max_elt r +end. + +(** ** Any element *) + +Definition choose := min_elt. + +(** ** Concatenation + + Same as [merge] but does not assume anything about heights. +*) + +Definition concat s1 s2 := + match s1, s2 with + | Leaf, _ => s2 + | _, Leaf => s1 + | _, Node l2 x2 r2 _ => + let (s2',m) := remove_min l2 x2 r2 in + join s1 m s2' + end. + +(** ** Splitting + + [split x s] returns a triple [(l, present, r)] where + - [l] is the set of elements of [s] that are [< x] + - [r] is the set of elements of [s] that are [> x] + - [present] is [true] if and only if [s] contains [x]. +*) + +Record triple := mktriple { t_left:t; t_in:bool; t_right:t }. +Notation "<< l , b , r >>" := (mktriple l b r) (at level 9). + +Fixpoint split x s : triple := match s with + | Leaf => << Leaf, false, Leaf >> + | Node l y r h => + match X.compare x y with + | Lt => let (ll,b,rl) := split x l in << ll, b, join rl y r >> + | Eq => << l, true, r >> + | Gt => let (rl,b,rr) := split x r in << join l y rl, b, rr >> + end + end. + +(** ** Intersection *) + +Fixpoint inter s1 s2 := match s1, s2 with + | Leaf, _ => Leaf + | _, Leaf => Leaf + | Node l1 x1 r1 h1, _ => + let (l2',pres,r2') := split x1 s2 in + if pres then join (inter l1 l2') x1 (inter r1 r2') + else concat (inter l1 l2') (inter r1 r2') + end. + +(** ** Difference *) + +Fixpoint diff s1 s2 := match s1, s2 with + | Leaf, _ => Leaf + | _, Leaf => s1 + | Node l1 x1 r1 h1, _ => + let (l2',pres,r2') := split x1 s2 in + if pres then concat (diff l1 l2') (diff r1 r2') + else join (diff l1 l2') x1 (diff r1 r2') +end. + +(** ** Union *) + +(** In ocaml, heights of [s1] and [s2] are compared each time in order + to recursively perform the split on the smaller set. + Unfortunately, this leads to a non-structural algorithm. The + following code is a simplification of the ocaml version: no + comparison of heights. It might be slightly slower, but + experimentally all the tests I've made in ocaml have shown this + potential slowdown to be non-significant. Anyway, the exact code + of ocaml has also been formalized thanks to Function+measure, see + [ocaml_union] in [MSetFullAVL]. +*) + +Fixpoint union s1 s2 := + match s1, s2 with + | Leaf, _ => s2 + | _, Leaf => s1 + | Node l1 x1 r1 h1, _ => + let (l2',_,r2') := split x1 s2 in + join (union l1 l2') x1 (union r1 r2') + end. + +(** ** Elements *) + +(** [elements_tree_aux acc t] catenates the elements of [t] in infix + order to the list [acc] *) + +Fixpoint elements_aux (acc : list X.t) (s : t) : list X.t := + match s with + | Leaf => acc + | Node l x r _ => elements_aux (x :: elements_aux acc r) l + end. + +(** then [elements] is an instanciation with an empty [acc] *) + +Definition elements := elements_aux nil. + +(** ** Filter *) + +Fixpoint filter_acc (f:elt->bool) acc s := match s with + | Leaf => acc + | Node l x r h => + filter_acc f (filter_acc f (if f x then add x acc else acc) l) r + end. + +Definition filter f := filter_acc f Leaf. + + +(** ** Partition *) + +Fixpoint partition_acc (f:elt->bool)(acc : t*t)(s : t) : t*t := + match s with + | Leaf => acc + | Node l x r _ => + let (acct,accf) := acc in + partition_acc f + (partition_acc f + (if f x then (add x acct, accf) else (acct, add x accf)) l) r + end. + +Definition partition f := partition_acc f (Leaf,Leaf). + +(** ** [for_all] and [exists] *) + +Fixpoint for_all (f:elt->bool) s := match s with + | Leaf => true + | Node l x r _ => f x &&& for_all f l &&& for_all f r +end. + +Fixpoint exists_ (f:elt->bool) s := match s with + | Leaf => false + | Node l x r _ => f x ||| exists_ f l ||| exists_ f r +end. + +(** ** Fold *) + +Fixpoint fold (A : Type) (f : elt -> A -> A)(s : t) : A -> A := + fun a => match s with + | Leaf => a + | Node l x r _ => fold f r (f x (fold f l a)) + end. +Implicit Arguments fold [A]. + + +(** ** Subset *) + +(** In ocaml, recursive calls are made on "half-trees" such as + (Node l1 x1 Leaf _) and (Node Leaf x1 r1 _). Instead of these + non-structural calls, we propose here two specialized functions for + these situations. This version should be almost as efficient as + the one of ocaml (closures as arguments may slow things a bit), + it is simply less compact. The exact ocaml version has also been + formalized (thanks to Function+measure), see [ocaml_subset] in + [MSetFullAVL]. + *) + +Fixpoint subsetl (subset_l1:t->bool) x1 s2 : bool := + match s2 with + | Leaf => false + | Node l2 x2 r2 h2 => + match X.compare x1 x2 with + | Eq => subset_l1 l2 + | Lt => subsetl subset_l1 x1 l2 + | Gt => mem x1 r2 &&& subset_l1 s2 + end + end. + +Fixpoint subsetr (subset_r1:t->bool) x1 s2 : bool := + match s2 with + | Leaf => false + | Node l2 x2 r2 h2 => + match X.compare x1 x2 with + | Eq => subset_r1 r2 + | Lt => mem x1 l2 &&& subset_r1 s2 + | Gt => subsetr subset_r1 x1 r2 + end + end. + +Fixpoint subset s1 s2 : bool := match s1, s2 with + | Leaf, _ => true + | Node _ _ _ _, Leaf => false + | Node l1 x1 r1 h1, Node l2 x2 r2 h2 => + match X.compare x1 x2 with + | Eq => subset l1 l2 &&& subset r1 r2 + | Lt => subsetl (subset l1) x1 l2 &&& subset r1 s2 + | Gt => subsetr (subset r1) x1 r2 &&& subset l1 s2 + end + end. + +(** ** A new comparison algorithm suggested by Xavier Leroy + + Transformation in C.P.S. suggested by Benjamin GrĂ©goire. + The original ocaml code (with non-structural recursive calls) + has also been formalized (thanks to Function+measure), see + [ocaml_compare] in [MSetFullAVL]. The following code with + continuations computes dramatically faster in Coq, and + should be almost as efficient after extraction. +*) + +(** Enumeration of the elements of a tree *) + +Inductive enumeration := + | End : enumeration + | More : elt -> t -> enumeration -> enumeration. + + +(** [cons t e] adds the elements of tree [t] on the head of + enumeration [e]. *) + +Fixpoint cons s e : enumeration := + match s with + | Leaf => e + | Node l x r h => cons l (More x r e) + end. + +(** One step of comparison of elements *) + +Definition compare_more x1 (cont:enumeration->comparison) e2 := + match e2 with + | End => Gt + | More x2 r2 e2 => + match X.compare x1 x2 with + | Eq => cont (cons r2 e2) + | Lt => Lt + | Gt => Gt + end + end. + +(** Comparison of left tree, middle element, then right tree *) + +Fixpoint compare_cont s1 (cont:enumeration->comparison) e2 := + match s1 with + | Leaf => cont e2 + | Node l1 x1 r1 _ => + compare_cont l1 (compare_more x1 (compare_cont r1 cont)) e2 + end. + +(** Initial continuation *) + +Definition compare_end e2 := + match e2 with End => Eq | _ => Lt end. + +(** The complete comparison *) + +Definition compare s1 s2 := compare_cont s1 compare_end (cons s2 End). + +(** ** Equality test *) + +Definition equal s1 s2 : bool := + match compare s1 s2 with + | Eq => true + | _ => false + end. + +End Ops. + + + +(** * MakeRaw + + Functor of pure functions + a posteriori proofs of invariant + preservation *) + +Module MakeRaw (Import I:Int)(X:OrderedType) <: RawSets X. +Include Ops I X. + +(** * Invariants *) + +(** ** Occurrence in a tree *) + +Inductive InT (x : elt) : tree -> Prop := + | IsRoot : forall l r h y, X.eq x y -> InT x (Node l y r h) + | InLeft : forall l r h y, InT x l -> InT x (Node l y r h) + | InRight : forall l r h y, InT x r -> InT x (Node l y r h). + +Definition In := InT. + +(** ** Some shortcuts *) + +Definition Equal s s' := forall a : elt, InT a s <-> InT a s'. +Definition Subset s s' := forall a : elt, InT a s -> InT a s'. +Definition Empty s := forall a : elt, ~ InT a s. +Definition For_all (P : elt -> Prop) s := forall x, InT x s -> P x. +Definition Exists (P : elt -> Prop) s := exists x, InT x s /\ P x. + +(** ** Binary search trees *) + +(** [lt_tree x s]: all elements in [s] are smaller than [x] + (resp. greater for [gt_tree]) *) + +Definition lt_tree x s := forall y, InT y s -> X.lt y x. +Definition gt_tree x s := forall y, InT y s -> X.lt x y. + +(** [bst t] : [t] is a binary search tree *) + +Inductive bst : tree -> Prop := + | BSLeaf : bst Leaf + | BSNode : forall x l r h, bst l -> bst r -> + lt_tree x l -> gt_tree x r -> bst (Node l x r h). + +(** [bst] is the (decidable) invariant our trees will have to satisfy. *) + +Definition IsOk := bst. + +Class Ok (s:t) : Prop := ok : bst s. + +Instance bst_Ok s (Hs : bst s) : Ok s := { ok := Hs }. + +Fixpoint ltb_tree x s := + match s with + | Leaf => true + | Node l y r _ => + match X.compare x y with + | Gt => ltb_tree x l && ltb_tree x r + | _ => false + end + end. + +Fixpoint gtb_tree x s := + match s with + | Leaf => true + | Node l y r _ => + match X.compare x y with + | Lt => gtb_tree x l && gtb_tree x r + | _ => false + end + end. + +Fixpoint isok s := + match s with + | Leaf => true + | Node l x r _ => isok l && isok r && ltb_tree x l && gtb_tree x r + end. + + +(** * Correctness proofs *) + +Module Import MX := OrderedTypeFacts X. + +(** * Automation and dedicated tactics *) + +Scheme tree_ind := Induction for tree Sort Prop. + +Local Hint Resolve MX.eq_refl MX.eq_trans MX.lt_trans @ok. +Local Hint Immediate MX.eq_sym. +Local Hint Unfold In lt_tree gt_tree. +Local Hint Constructors InT bst. +Local Hint Unfold Ok. + +Tactic Notation "factornode" ident(l) ident(x) ident(r) ident(h) + "as" ident(s) := + set (s:=Node l x r h) in *; clearbody s; clear l x r h. + +(** Automatic treatment of [Ok] hypothesis *) + +Ltac inv_ok := match goal with + | H:Ok (Node _ _ _ _) |- _ => inversion_clear H; inv_ok + | H:Ok Leaf |- _ => clear H; inv_ok + | H:bst ?x |- _ => change (Ok x) in H; inv_ok + | _ => idtac +end. + +(** A tactic to repeat [inversion_clear] on all hyps of the + form [(f (Node _ _ _ _))] *) + +Ltac is_tree_constr c := + match c with + | Leaf => idtac + | Node _ _ _ _ => idtac + | _ => fail + end. + +Ltac invtree f := + match goal with + | H:f ?s |- _ => is_tree_constr s; inversion_clear H; invtree f + | H:f _ ?s |- _ => is_tree_constr s; inversion_clear H; invtree f + | H:f _ _ ?s |- _ => is_tree_constr s; inversion_clear H; invtree f + | _ => idtac + end. + +Ltac inv := inv_ok; invtree InT. + +Ltac intuition_in := repeat progress (intuition; inv). + +(** Helper tactic concerning order of elements. *) + +Ltac order := match goal with + | U: lt_tree _ ?s, V: InT _ ?s |- _ => generalize (U _ V); clear U; order + | U: gt_tree _ ?s, V: InT _ ?s |- _ => generalize (U _ V); clear U; order + | _ => MX.order +end. + + +(** [isok] is indeed a decision procedure for [Ok] *) + +Lemma ltb_tree_iff : forall x s, lt_tree x s <-> ltb_tree x s = true. +Proof. + induction s as [|l IHl y r IHr h]; simpl. + unfold lt_tree; intuition_in. + elim_compare x y. + split; intros; try discriminate. assert (X.lt y x) by auto. order. + split; intros; try discriminate. assert (X.lt y x) by auto. order. + rewrite !andb_true_iff, <-IHl, <-IHr. + unfold lt_tree; intuition_in; order. +Qed. + +Lemma gtb_tree_iff : forall x s, gt_tree x s <-> gtb_tree x s = true. +Proof. + induction s as [|l IHl y r IHr h]; simpl. + unfold gt_tree; intuition_in. + elim_compare x y. + split; intros; try discriminate. assert (X.lt x y) by auto. order. + rewrite !andb_true_iff, <-IHl, <-IHr. + unfold gt_tree; intuition_in; order. + split; intros; try discriminate. assert (X.lt x y) by auto. order. +Qed. + +Lemma isok_iff : forall s, Ok s <-> isok s = true. +Proof. + induction s as [|l IHl y r IHr h]; simpl. + intuition_in. + rewrite !andb_true_iff, <- IHl, <-IHr, <- ltb_tree_iff, <- gtb_tree_iff. + intuition_in. +Qed. + +Instance isok_Ok s : isok s = true -> Ok s | 10. +Proof. intros; apply <- isok_iff; auto. Qed. + + +(** * Basic results about [In], [lt_tree], [gt_tree], [height] *) + +(** [In] is compatible with [X.eq] *) + +Lemma In_1 : + forall s x y, X.eq x y -> InT x s -> InT y s. +Proof. + induction s; simpl; intuition_in; eauto. +Qed. +Local Hint Immediate In_1. + +Instance In_compat : Proper (X.eq==>eq==>iff) InT. +Proof. +apply proper_sym_impl_iff_2; auto with *. +repeat red; intros; subst. apply In_1 with x; auto. +Qed. + +Lemma In_node_iff : + forall l x r h y, + InT y (Node l x r h) <-> InT y l \/ X.eq y x \/ InT y r. +Proof. + intuition_in. +Qed. + +(** Results about [lt_tree] and [gt_tree] *) + +Lemma lt_leaf : forall x : elt, lt_tree x Leaf. +Proof. + red; inversion 1. +Qed. + +Lemma gt_leaf : forall x : elt, gt_tree x Leaf. +Proof. + red; inversion 1. +Qed. + +Lemma lt_tree_node : + forall (x y : elt) (l r : tree) (h : int), + lt_tree x l -> lt_tree x r -> X.lt y x -> lt_tree x (Node l y r h). +Proof. + unfold lt_tree; intuition_in; order. +Qed. + +Lemma gt_tree_node : + forall (x y : elt) (l r : tree) (h : int), + gt_tree x l -> gt_tree x r -> X.lt x y -> gt_tree x (Node l y r h). +Proof. + unfold gt_tree; intuition_in; order. +Qed. + +Local Hint Resolve lt_leaf gt_leaf lt_tree_node gt_tree_node. + +Lemma lt_tree_not_in : + forall (x : elt) (t : tree), lt_tree x t -> ~ InT x t. +Proof. + intros; intro; order. +Qed. + +Lemma lt_tree_trans : + forall x y, X.lt x y -> forall t, lt_tree x t -> lt_tree y t. +Proof. + eauto. +Qed. + +Lemma gt_tree_not_in : + forall (x : elt) (t : tree), gt_tree x t -> ~ InT x t. +Proof. + intros; intro; order. +Qed. + +Lemma gt_tree_trans : + forall x y, X.lt y x -> forall t, gt_tree x t -> gt_tree y t. +Proof. + eauto. +Qed. + +Local Hint Resolve lt_tree_not_in lt_tree_trans gt_tree_not_in gt_tree_trans. + +(** * Inductions principles for some of the set operators *) + +Functional Scheme bal_ind := Induction for bal Sort Prop. +Functional Scheme remove_min_ind := Induction for remove_min Sort Prop. +Functional Scheme merge_ind := Induction for merge Sort Prop. +Functional Scheme min_elt_ind := Induction for min_elt Sort Prop. +Functional Scheme max_elt_ind := Induction for max_elt Sort Prop. +Functional Scheme concat_ind := Induction for concat Sort Prop. +Functional Scheme inter_ind := Induction for inter Sort Prop. +Functional Scheme diff_ind := Induction for diff Sort Prop. +Functional Scheme union_ind := Induction for union Sort Prop. + +Ltac induct s x := + induction s as [|l IHl x' r IHr h]; simpl; intros; + [|elim_compare x x'; intros; inv]. + + +(** * Notations and helper lemma about pairs and triples *) + +Notation "s #1" := (fst s) (at level 9, format "s '#1'") : pair_scope. +Notation "s #2" := (snd s) (at level 9, format "s '#2'") : pair_scope. +Notation "t #l" := (t_left t) (at level 9, format "t '#l'") : pair_scope. +Notation "t #b" := (t_in t) (at level 9, format "t '#b'") : pair_scope. +Notation "t #r" := (t_right t) (at level 9, format "t '#r'") : pair_scope. + +Open Local Scope pair_scope. + + +(** * Empty set *) + +Lemma empty_spec : Empty empty. +Proof. + intro; intro. + inversion H. +Qed. + +Instance empty_ok : Ok empty. +Proof. + auto. +Qed. + +(** * Emptyness test *) + +Lemma is_empty_spec : forall s, is_empty s = true <-> Empty s. +Proof. + destruct s as [|r x l h]; simpl; auto. + split; auto. red; red; intros; inv. + split; auto. try discriminate. intro H; elim (H x); auto. +Qed. + +(** * Appartness *) + +Lemma mem_spec : forall s x `{Ok s}, mem x s = true <-> InT x s. +Proof. + split. + induct s x; auto; try discriminate. + induct s x; intuition_in; order. +Qed. + + +(** * Singleton set *) + +Lemma singleton_spec : forall x y, InT y (singleton x) <-> X.eq y x. +Proof. + unfold singleton; intuition_in. +Qed. + +Instance singleton_ok x : Ok (singleton x). +Proof. + unfold singleton; auto. +Qed. + + + +(** * Helper functions *) + +Lemma create_spec : + forall l x r y, InT y (create l x r) <-> X.eq y x \/ InT y l \/ InT y r. +Proof. + unfold create; split; [ inversion_clear 1 | ]; intuition. +Qed. + +Instance create_ok l x r `(Ok l, Ok r, lt_tree x l, gt_tree x r) : + Ok (create l x r). +Proof. + unfold create; auto. +Qed. + +Lemma bal_spec : forall l x r y, + InT y (bal l x r) <-> X.eq y x \/ InT y l \/ InT y r. +Proof. + intros l x r; functional induction bal l x r; intros; try clear e0; + rewrite !create_spec; intuition_in. +Qed. + +Instance bal_ok l x r `(Ok l, Ok r, lt_tree x l, gt_tree x r) : + Ok (bal l x r). +Proof. + functional induction bal l x r; intros; + inv; repeat apply create_ok; auto; unfold create; + (apply lt_tree_node || apply gt_tree_node); auto; + (eapply lt_tree_trans || eapply gt_tree_trans); eauto. +Qed. + + +(** * Insertion *) + +Lemma add_spec' : forall s x y, + InT y (add x s) <-> X.eq y x \/ InT y s. +Proof. + induct s x; try rewrite ?bal_spec, ?IHl, ?IHr; intuition_in. + setoid_replace y with x'; eauto. +Qed. + +Lemma add_spec : forall s x y `{Ok s}, + InT y (add x s) <-> X.eq y x \/ InT y s. +Proof. intros; apply add_spec'. Qed. + +Instance add_ok s x `(Ok s) : Ok (add x s). +Proof. + induct s x; auto; apply bal_ok; auto; + intros y; rewrite add_spec'; intuition; order. +Qed. + + +Open Scope Int_scope. + +(** * Join *) + +(* Function/Functional Scheme can't deal with internal fix. + Let's do its job by hand: *) + +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)); + [ 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)); + [ 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] + end + | ] ] ] ]; intros. + +Lemma join_spec : forall l x r y, + InT y (join l x r) <-> X.eq y x \/ InT y l \/ InT y r. +Proof. + join_tac. + simpl. + rewrite add_spec'; intuition_in. + rewrite add_spec'; intuition_in. + rewrite bal_spec, Hlr; clear Hlr Hrl; intuition_in. + rewrite bal_spec, Hrl; clear Hlr Hrl; intuition_in. + apply create_spec. +Qed. + +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 *. + intuition; [ setoid_replace y with x | ]; eauto. + intuition; [ setoid_replace y with x | ]; eauto. +Qed. + + +(** * Extraction of minimum element *) + +Lemma remove_min_spec : forall l x r h y, + InT y (Node l x r h) <-> + X.eq y (remove_min l x r)#2 \/ InT y (remove_min l x r)#1. +Proof. + intros l x r; functional induction (remove_min l x r); simpl in *; intros. + intuition_in. + rewrite bal_spec, In_node_iff, IHp, e0; simpl; intuition. +Qed. + +Instance remove_min_ok l x r : forall h `(Ok (Node l x r h)), + Ok (remove_min l x r)#1. +Proof. + functional induction (remove_min l x r); simpl; intros. + inv; auto. + assert (O : Ok (Node ll lx lr _x)) by (inv; auto). + assert (L : lt_tree x (Node ll lx lr _x)) by (inv; auto). + specialize IHp with (1:=O); rewrite e0 in IHp; auto; simpl in *. + apply bal_ok; auto. + inv; auto. + intro y; specialize (L y). + rewrite remove_min_spec, e0 in L; simpl in L; intuition. + inv; auto. +Qed. + +Lemma remove_min_gt_tree : forall l x r h `{Ok (Node l x r h)}, + gt_tree (remove_min l x r)#2 (remove_min l x r)#1. +Proof. + intros l x r; functional induction (remove_min l x r); simpl; intros. + inv; auto. + assert (O : Ok (Node ll lx lr _x)) by (inv; auto). + assert (L : lt_tree x (Node ll lx lr _x)) by (inv; auto). + specialize IHp with (1:=O); rewrite e0 in IHp; simpl in IHp. + intro y; rewrite bal_spec; intuition; + specialize (L m); rewrite remove_min_spec, e0 in L; simpl in L; + [setoid_replace y with x|inv]; eauto. +Qed. +Local Hint Resolve remove_min_gt_tree. + + + +(** * Merging two trees *) + +Lemma merge_spec : forall s1 s2 y, + InT y (merge s1 s2) <-> InT y s1 \/ InT y s2. +Proof. + intros s1 s2; functional induction (merge s1 s2); intros; + try factornode _x _x0 _x1 _x2 as s1. + intuition_in. + intuition_in. + rewrite bal_spec, remove_min_spec, e1; simpl; intuition. +Qed. + +Instance merge_ok s1 s2 : forall `(Ok s1, Ok s2) + `(forall y1 y2 : elt, InT y1 s1 -> InT y2 s2 -> X.lt y1 y2), + Ok (merge s1 s2). +Proof. + functional induction (merge s1 s2); intros; auto; + try factornode _x _x0 _x1 _x2 as s1. + apply bal_ok; auto. + change s2' with ((s2',m)#1); rewrite <-e1; eauto with *. + intros y Hy. + apply H1; auto. + rewrite remove_min_spec, e1; simpl; auto. + change (gt_tree (s2',m)#2 (s2',m)#1); rewrite <-e1; eauto. +Qed. + + + +(** * Deletion *) + +Lemma remove_spec : forall s x y `{Ok s}, + (InT y (remove x s) <-> InT y s /\ ~ X.eq y x). +Proof. + induct s x. + intuition_in. + rewrite merge_spec; intuition; [order|order|intuition_in]. + elim H6; eauto. + rewrite bal_spec, IHl; clear IHl IHr; intuition; [order|order|intuition_in]. + rewrite bal_spec, IHr; clear IHl IHr; intuition; [order|order|intuition_in]. +Qed. + +Instance remove_ok s x `(Ok s) : Ok (remove x s). +Proof. + induct s x. + auto. + (* EQ *) + apply merge_ok; eauto. + (* LT *) + apply bal_ok; auto. + intro z; rewrite remove_spec; auto; destruct 1; eauto. + (* GT *) + apply bal_ok; auto. + intro z; rewrite remove_spec; auto; destruct 1; eauto. +Qed. + + +(** * Minimum element *) + +Lemma min_elt_spec1 : forall s x, min_elt s = Some x -> InT x s. +Proof. + intro s; functional induction (min_elt s); auto; inversion 1; auto. +Qed. + +Lemma min_elt_spec2 : forall s x y `{Ok s}, + min_elt s = Some x -> InT y s -> ~ X.lt y x. +Proof. + intro s; functional induction (min_elt s); + try rename _x1 into l1, _x2 into x1, _x3 into r1, _x4 into h1. + discriminate. + intros x y0 U V W. + inversion V; clear V; subst. + inv; order. + intros; inv; auto. + assert (X.lt x y) by (apply H4; apply min_elt_spec1; auto). + order. + assert (X.lt x1 y) by auto. + assert (~X.lt x1 x) by auto. + order. +Qed. + +Lemma min_elt_spec3 : forall s, min_elt s = None -> Empty s. +Proof. + intro s; functional induction (min_elt s). + red; red; inversion 2. + inversion 1. + intro H0. + destruct (IHo H0 _x2); auto. +Qed. + + + +(** * Maximum element *) + +Lemma max_elt_spec1 : forall s x, max_elt s = Some x -> InT x s. +Proof. + intro s; functional induction (max_elt s); auto; inversion 1; auto. +Qed. + +Lemma max_elt_spec2 : forall s x y `{Ok s}, + max_elt s = Some x -> InT y s -> ~ X.lt x y. +Proof. + intro s; functional induction (max_elt s); + try rename _x1 into l1, _x2 into x1, _x3 into r1, _x4 into h1. + discriminate. + intros x y0 U V W. + inversion V; clear V; subst. + inv; order. + intros; inv; auto. + assert (X.lt y x1) by auto. + assert (~ X.lt x x1) by auto. + order. + assert (X.lt y x) by (apply H5; apply max_elt_spec1; auto). + order. +Qed. + +Lemma max_elt_spec3 : forall s, max_elt s = None -> Empty s. +Proof. + intro s; functional induction (max_elt s). + red; auto. + inversion 1. + intros H0; destruct (IHo H0 _x2); auto. +Qed. + + + +(** * Any element *) + +Lemma choose_spec1 : forall s x, choose s = Some x -> InT x s. +Proof. + exact min_elt_spec1. +Qed. + +Lemma choose_spec2 : forall s, choose s = None -> Empty s. +Proof. + exact min_elt_spec3. +Qed. + +Lemma choose_spec3 : forall s s' x x' `{Ok s, Ok s'}, + choose s = Some x -> choose s' = Some x' -> + Equal s s' -> X.eq x x'. +Proof. + unfold choose, Equal; intros s s' x x' Hb Hb' Hx Hx' H. + assert (~X.lt x x'). + apply min_elt_spec2 with s'; auto. + rewrite <-H; auto using min_elt_spec1. + assert (~X.lt x' x). + apply min_elt_spec2 with s; auto. + rewrite H; auto using min_elt_spec1. + elim_compare x x'; intuition. +Qed. + + +(** * Concatenation *) + +Lemma concat_spec : forall s1 s2 y, + InT y (concat s1 s2) <-> InT y s1 \/ InT y s2. +Proof. + intros s1 s2; functional induction (concat s1 s2); intros; + try factornode _x _x0 _x1 _x2 as s1. + intuition_in. + intuition_in. + rewrite join_spec, remove_min_spec, e1; simpl; intuition. +Qed. + +Instance concat_ok s1 s2 : forall `(Ok s1, Ok s2) + `(forall y1 y2 : elt, InT y1 s1 -> InT y2 s2 -> X.lt y1 y2), + Ok (concat s1 s2). +Proof. + functional induction (concat s1 s2); intros; auto; + try factornode _x _x0 _x1 _x2 as s1. + apply join_ok; auto. + change (Ok (s2',m)#1); rewrite <-e1; eauto with *. + intros y Hy. + apply H1; auto. + rewrite remove_min_spec, e1; simpl; auto. + change (gt_tree (s2',m)#2 (s2',m)#1); rewrite <-e1; eauto. +Qed. + + + +(** * Splitting *) + +Lemma split_spec1 : forall s x y `{Ok s}, + (InT y (split x s)#l <-> InT y s /\ X.lt y x). +Proof. + induct s x. + intuition_in. + intuition_in; order. + specialize (IHl x y). + destruct (split x l); simpl in *. rewrite IHl; intuition_in; order. + specialize (IHr x y). + destruct (split x r); simpl in *. rewrite join_spec, IHr; intuition_in; order. +Qed. + +Lemma split_spec2 : forall s x y `{Ok s}, + (InT y (split x s)#r <-> InT y s /\ X.lt x y). +Proof. + induct s x. + intuition_in. + intuition_in; order. + specialize (IHl x y). + destruct (split x l); simpl in *. rewrite join_spec, IHl; intuition_in; order. + specialize (IHr x y). + destruct (split x r); simpl in *. rewrite IHr; intuition_in; order. +Qed. + +Lemma split_spec3 : forall s x `{Ok s}, + ((split x s)#b = true <-> InT x s). +Proof. + induct s x. + intuition_in; try discriminate. + intuition. + specialize (IHl x). + destruct (split x l); simpl in *. rewrite IHl; intuition_in; order. + specialize (IHr x). + destruct (split x r); simpl in *. rewrite IHr; intuition_in; order. +Qed. + +Lemma split_ok : forall s x `{Ok s}, Ok (split x s)#l /\ Ok (split x s)#r. +Proof. + induct s x; simpl; auto. + specialize (IHl x). + generalize (fun y => @split_spec2 _ x y H1). + destruct (split x l); simpl in *; intuition. apply join_ok; auto. + intros y; rewrite H; intuition. + specialize (IHr x). + generalize (fun y => @split_spec1 _ x y H2). + destruct (split x r); simpl in *; intuition. apply join_ok; auto. + intros y; rewrite H; intuition. +Qed. + +Instance split_ok1 s x `(Ok s) : Ok (split x s)#l. +Proof. intros; destruct (@split_ok s x); auto. Qed. + +Instance split_ok2 s x `(Ok s) : Ok (split x s)#r. +Proof. intros; destruct (@split_ok s x); auto. Qed. + + +(** * Intersection *) + +Ltac destruct_split := match goal with + | H : split ?x ?s = << ?u, ?v, ?w >> |- _ => + assert ((split x s)#l = u) by (rewrite H; auto); + assert ((split x s)#b = v) by (rewrite H; auto); + assert ((split x s)#r = w) by (rewrite H; auto); + clear H; subst u w + end. + +Lemma inter_spec_ok : forall s1 s2 `{Ok s1, Ok s2}, + Ok (inter s1 s2) /\ (forall y, InT y (inter s1 s2) <-> InT y s1 /\ InT y s2). +Proof. + intros s1 s2; functional induction inter s1 s2; intros B1 B2; + [intuition_in|intuition_in | | ]; + factornode _x0 _x1 _x2 _x3 as s2; destruct_split; inv; + destruct IHt0 as (IHo1,IHi1), IHt1 as (IHo2,IHi2); auto with *; + split; intros. + (* Ok join *) + apply join_ok; auto with *; intro y; rewrite ?IHi1, ?IHi2; intuition. + (* InT join *) + rewrite join_spec, IHi1, IHi2, split_spec1, split_spec2; intuition_in. + setoid_replace y with x1; auto. rewrite <- split_spec3; auto. + (* Ok concat *) + apply concat_ok; auto with *; intros y1 y2; rewrite IHi1, IHi2; intuition; order. + (* InT concat *) + rewrite concat_spec, IHi1, IHi2, split_spec1, split_spec2; auto. + intuition_in. + absurd (InT x1 s2). + rewrite <- split_spec3; auto; congruence. + setoid_replace x1 with y; auto. +Qed. + +Lemma inter_spec : forall s1 s2 y `{Ok s1, Ok s2}, + (InT y (inter s1 s2) <-> InT y s1 /\ InT y s2). +Proof. intros; destruct (@inter_spec_ok s1 s2); auto. Qed. + +Instance inter_ok s1 s2 `(Ok s1, Ok s2) : Ok (inter s1 s2). +Proof. intros; destruct (@inter_spec_ok s1 s2); auto. Qed. + + +(** * Difference *) + +Lemma diff_spec_ok : forall s1 s2 `{Ok s1, Ok s2}, + Ok (diff s1 s2) /\ (forall y, InT y (diff s1 s2) <-> InT y s1 /\ ~InT y s2). +Proof. + intros s1 s2; functional induction diff s1 s2; intros B1 B2; + [intuition_in|intuition_in | | ]; + factornode _x0 _x1 _x2 _x3 as s2; destruct_split; inv; + destruct IHt0 as (IHb1,IHi1), IHt1 as (IHb2,IHi2); auto with *; + split; intros. + (* Ok concat *) + apply concat_ok; auto; intros y1 y2; rewrite IHi1, IHi2; intuition; order. + (* InT concat *) + rewrite concat_spec, IHi1, IHi2, split_spec1, split_spec2; intuition_in. + absurd (InT x1 s2). + setoid_replace x1 with y; auto. + rewrite <- split_spec3; auto; congruence. + (* Ok join *) + apply join_ok; auto; intro y; rewrite ?IHi1, ?IHi2; intuition. + (* InT join *) + rewrite join_spec, IHi1, IHi2, split_spec1, split_spec2; auto with *. + intuition_in. + absurd (InT x1 s2); auto. + rewrite <- split_spec3; auto; congruence. + setoid_replace x1 with y; auto. +Qed. + +Lemma diff_spec : forall s1 s2 y `{Ok s1, Ok s2}, + (InT y (diff s1 s2) <-> InT y s1 /\ ~InT y s2). +Proof. intros; destruct (@diff_spec_ok s1 s2); auto. Qed. + +Instance diff_ok s1 s2 `(Ok s1, Ok s2) : Ok (diff s1 s2). +Proof. intros; destruct (@diff_spec_ok s1 s2); auto. Qed. + + +(** * Union *) + +Lemma union_spec : forall s1 s2 y `{Ok s1, Ok s2}, + (InT y (union s1 s2) <-> InT y s1 \/ InT y s2). +Proof. + intros s1 s2; functional induction union s1 s2; intros y B1 B2. + intuition_in. + intuition_in. + factornode _x0 _x1 _x2 _x3 as s2; destruct_split; inv. + rewrite join_spec, IHt0, IHt1, split_spec1, split_spec2; auto with *. + elim_compare y x1; intuition_in. +Qed. + +Instance union_ok s1 s2 : forall `(Ok s1, Ok s2), Ok (union s1 s2). +Proof. + functional induction union s1 s2; intros B1 B2; auto. + factornode _x0 _x1 _x2 _x3 as s2; destruct_split; inv. + apply join_ok; auto with *. + intro y; rewrite union_spec, split_spec1; intuition_in. + intro y; rewrite union_spec, split_spec2; intuition_in. +Qed. + + +(** * Elements *) + +Lemma elements_spec1' : forall s acc x, + InA X.eq x (elements_aux acc s) <-> InT x s \/ InA X.eq x acc. +Proof. + induction s as [ | l Hl x r Hr h ]; simpl; auto. + intuition. + inversion H0. + intros. + rewrite Hl. + destruct (Hr acc x0); clear Hl Hr. + intuition; inversion_clear H3; intuition. +Qed. + +Lemma elements_spec1 : forall s x, InA X.eq x (elements s) <-> InT x s. +Proof. + intros; generalize (elements_spec1' s nil x); intuition. + inversion_clear H0. +Qed. + +Lemma elements_spec2' : forall s acc `{Ok s}, sort X.lt acc -> + (forall x y : elt, InA X.eq x acc -> InT y s -> X.lt y x) -> + sort X.lt (elements_aux acc s). +Proof. + induction s as [ | l Hl y r Hr h]; simpl; intuition. + inv. + apply Hl; auto. + constructor. + apply Hr; auto. + eapply InA_InfA; eauto with *. + intros. + destruct (elements_spec1' r acc y0); intuition. + intros. + inversion_clear H. + order. + destruct (elements_spec1' r acc x); intuition eauto. +Qed. + +Lemma elements_spec2 : forall s `(Ok s), sort X.lt (elements s). +Proof. + intros; unfold elements; apply elements_spec2'; auto. + intros; inversion H0. +Qed. +Local Hint Resolve elements_spec2. + +Lemma elements_spec2w : forall s `(Ok s), NoDupA X.eq (elements s). +Proof. + intros. eapply SortA_NoDupA; eauto with *. +Qed. + +Lemma elements_aux_cardinal : + forall s acc, (length acc + cardinal s)%nat = length (elements_aux acc s). +Proof. + simple induction s; simpl in |- *; intuition. + rewrite <- H. + simpl in |- *. + rewrite <- H0; omega. +Qed. + +Lemma elements_cardinal : forall s : tree, cardinal s = length (elements s). +Proof. + exact (fun s => elements_aux_cardinal s nil). +Qed. + +Definition cardinal_spec (s:t)(Hs:Ok s) := elements_cardinal s. + +Lemma elements_app : + forall s acc, elements_aux acc s = elements s ++ acc. +Proof. + induction s; simpl; intros; auto. + rewrite IHs1, IHs2. + unfold elements; simpl. + rewrite 2 IHs1, IHs2, <- !app_nil_end, !app_ass; auto. +Qed. + +Lemma elements_node : + forall l x r h acc, + elements l ++ x :: elements r ++ acc = + elements (Node l x r h) ++ acc. +Proof. + unfold elements; simpl; intros; auto. + rewrite !elements_app, <- !app_nil_end, !app_ass; auto. +Qed. + + +(** * Filter *) + +Lemma filter_spec' : forall s x acc f, + Proper (X.eq==>eq) f -> + (InT x (filter_acc f acc s) <-> InT x acc \/ InT x s /\ f x = true). +Proof. + induction s; simpl; intros. + intuition_in. + rewrite IHs2, IHs1 by (destruct (f t0); auto). + case_eq (f t0); intros. + rewrite add_spec'; auto. + intuition_in. + rewrite (H _ _ H2). + intuition. + intuition_in. + rewrite (H _ _ H2) in H3. + rewrite H0 in H3; discriminate. +Qed. + +Instance filter_ok' : forall s acc f `(Ok s, Ok acc), + Ok (filter_acc f acc s). +Proof. + induction s; simpl; auto. + intros. inv. + destruct (f t0); auto with *. +Qed. + +Lemma filter_spec : forall s x f, + Proper (X.eq==>eq) f -> + (InT x (filter f s) <-> InT x s /\ f x = true). +Proof. + unfold filter; intros; rewrite filter_spec'; intuition_in. +Qed. + +Instance filter_ok s f `(Ok s) : Ok (filter f s). +Proof. + unfold filter; intros; apply filter_ok'; auto. +Qed. + + +(** * Partition *) + +Lemma partition_spec1' : forall s acc f, + Proper (X.eq==>eq) f -> forall x : elt, + InT x (partition_acc f acc s)#1 <-> + InT x acc#1 \/ InT x s /\ f x = true. +Proof. + induction s; simpl; intros. + intuition_in. + destruct acc as [acct accf]; simpl in *. + rewrite IHs2 by + (destruct (f t0); auto; apply partition_acc_avl_1; simpl; auto). + rewrite IHs1 by (destruct (f t0); simpl; auto). + case_eq (f t0); simpl; intros. + rewrite add_spec'; auto. + intuition_in. + rewrite (H _ _ H2). + intuition. + intuition_in. + rewrite (H _ _ H2) in H3. + rewrite H0 in H3; discriminate. +Qed. + +Lemma partition_spec2' : forall s acc f, + Proper (X.eq==>eq) f -> forall x : elt, + InT x (partition_acc f acc s)#2 <-> + InT x acc#2 \/ InT x s /\ f x = false. +Proof. + induction s; simpl; intros. + intuition_in. + destruct acc as [acct accf]; simpl in *. + rewrite IHs2 by + (destruct (f t0); auto; apply partition_acc_avl_2; simpl; auto). + rewrite IHs1 by (destruct (f t0); simpl; auto). + case_eq (f t0); simpl; intros. + intuition. + intuition_in. + rewrite (H _ _ H2) in H3. + rewrite H0 in H3; discriminate. + rewrite add_spec'; auto. + intuition_in. + rewrite (H _ _ H2). + intuition. +Qed. + +Lemma partition_spec1 : forall s f, + Proper (X.eq==>eq) f -> + Equal (partition f s)#1 (filter f s). +Proof. + unfold partition; intros s f P x. + rewrite partition_spec1', filter_spec; simpl; intuition_in. +Qed. + +Lemma partition_spec2 : forall s f, + Proper (X.eq==>eq) f -> + Equal (partition f s)#2 (filter (fun x => negb (f x)) s). +Proof. + unfold partition; intros s f P x. + rewrite partition_spec2', filter_spec; simpl; intuition_in. + rewrite H1; auto. + right; split; auto. + rewrite negb_true_iff in H1; auto. + intros u v H; rewrite H; auto. +Qed. + +Instance partition_ok1' : forall s acc f `(Ok s, Ok acc#1), + Ok (partition_acc f acc s)#1. +Proof. + induction s; simpl; auto. + destruct acc as [acct accf]; simpl in *. + intros. inv. + destruct (f t0); auto. + apply IHs2; simpl; auto. + apply IHs1; simpl; auto with *. +Qed. + +Instance partition_ok2' : forall s acc f `(Ok s, Ok acc#2), + Ok (partition_acc f acc s)#2. +Proof. + induction s; simpl; auto. + destruct acc as [acct accf]; simpl in *. + intros. inv. + destruct (f t0); auto. + apply IHs2; simpl; auto. + apply IHs1; simpl; auto with *. +Qed. + +Instance partition_ok1 s f `(Ok s) : Ok (partition f s)#1. +Proof. apply partition_ok1'; auto. Qed. + +Instance partition_ok2 s f `(Ok s) : Ok (partition f s)#2. +Proof. apply partition_ok2'; auto. Qed. + + + +(** * [for_all] and [exists] *) + +Lemma for_all_spec : forall s f, Proper (X.eq==>eq) f -> + (for_all f s = true <-> For_all (fun x => f x = true) s). +Proof. + split. + induction s; simpl; auto; intros; red; intros; inv. + destruct (andb_prop _ _ H0); auto. + destruct (andb_prop _ _ H1); eauto. + apply IHs1; auto. + destruct (andb_prop _ _ H0); auto. + destruct (andb_prop _ _ H1); auto. + apply IHs2; auto. + destruct (andb_prop _ _ H0); auto. + (* <- *) + induction s; simpl; auto. + intros. red in H0. + rewrite IHs1; try red; auto. + rewrite IHs2; try red; auto. + generalize (H0 t0). + destruct (f t0); simpl; auto. +Qed. + +Lemma exists_spec : forall s f, Proper (X.eq==>eq) f -> + (exists_ f s = true <-> Exists (fun x => f x = true) s). +Proof. + split. + induction s; simpl; intros; rewrite <- ?orb_lazy_alt in *. + discriminate. + destruct (orb_true_elim _ _ H0) as [H1|H1]. + destruct (orb_true_elim _ _ H1) as [H2|H2]. + exists t0; auto. + destruct (IHs1 H2); auto; exists x; intuition. + destruct (IHs2 H1); auto; exists x; intuition. + (* <- *) + induction s; simpl; destruct 1 as (x,(U,V)); inv; rewrite <- ?orb_lazy_alt. + rewrite (H _ _ (MX.eq_sym H0)); rewrite V; auto. + apply orb_true_intro; left. + apply orb_true_intro; right; apply IHs1; auto; exists x; auto. + apply orb_true_intro; right; apply IHs2; auto; exists x; auto. +Qed. + + +(** * Fold *) + +Lemma fold_spec' : + forall (A : Type) (f : elt -> A -> A) (s : tree) (i : A) (acc : list elt), + fold_left (flip f) (elements_aux acc s) i = fold_left (flip f) acc (fold f s i). +Proof. + induction s as [|l IHl x r IHr h]; simpl; intros; auto. + rewrite IHl. + simpl. unfold flip at 2. + apply IHr. +Qed. + +Lemma fold_spec : + forall (s:t) (A : Type) (i : A) (f : elt -> A -> A), + fold f s i = fold_left (flip f) (elements s) i. +Proof. + unfold elements. + induction s as [|l IHl x r IHr h]; simpl; intros; auto. + rewrite fold_spec'. + rewrite IHr. + simpl; auto. +Qed. + + +(** * Subset *) + +Lemma subsetl_spec : forall subset_l1 l1 x1 h1 s2 + `{Ok (Node l1 x1 Leaf h1), Ok s2}, + (forall s `{Ok s}, (subset_l1 s = true <-> Subset l1 s)) -> + (subsetl subset_l1 x1 s2 = true <-> Subset (Node l1 x1 Leaf h1) s2 ). +Proof. + induction s2 as [|l2 IHl2 x2 r2 IHr2 h2]; simpl; intros. + unfold Subset; intuition; try discriminate. + assert (H': InT x1 Leaf) by auto; inversion H'. + specialize (IHl2 H). + specialize (IHr2 H). + inv. + elim_compare x1 x2. + + rewrite H1 by auto; clear H1 IHl2 IHr2. + unfold Subset. intuition_in. + assert (X.eq a x2) by order; intuition_in. + assert (InT a (Node l2 x2 r2 h2)) by auto; intuition_in; order. + + rewrite IHl2 by auto; clear H1 IHl2 IHr2. + unfold Subset. intuition_in. + assert (InT a (Node l2 x2 r2 h2)) by auto; intuition_in; order. + assert (InT a (Node l2 x2 r2 h2)) by auto; intuition_in; order. + + rewrite <-andb_lazy_alt, andb_true_iff, H1 by auto; clear H1 IHl2 IHr2. + unfold Subset. intuition_in. + constructor 3. setoid_replace a with x1; auto. rewrite <- mem_spec; auto. + rewrite mem_spec; auto. + assert (InT x1 (Node l2 x2 r2 h2)) by auto; intuition_in; order. +Qed. + + +Lemma subsetr_spec : forall subset_r1 r1 x1 h1 s2, + bst (Node Leaf x1 r1 h1) -> bst s2 -> + (forall s, bst s -> (subset_r1 s = true <-> Subset r1 s)) -> + (subsetr subset_r1 x1 s2 = true <-> Subset (Node Leaf x1 r1 h1) s2). +Proof. + induction s2 as [|l2 IHl2 x2 r2 IHr2 h2]; simpl; intros. + unfold Subset; intuition; try discriminate. + assert (H': InT x1 Leaf) by auto; inversion H'. + specialize (IHl2 H). + specialize (IHr2 H). + inv. + elim_compare x1 x2. + + rewrite H1 by auto; clear H1 IHl2 IHr2. + unfold Subset. intuition_in. + assert (X.eq a x2) by order; intuition_in. + assert (InT a (Node l2 x2 r2 h2)) by auto; intuition_in; order. + + rewrite <-andb_lazy_alt, andb_true_iff, H1 by auto; clear H1 IHl2 IHr2. + unfold Subset. intuition_in. + constructor 2. setoid_replace a with x1; auto. rewrite <- mem_spec; auto. + rewrite mem_spec; auto. + assert (InT x1 (Node l2 x2 r2 h2)) by auto; intuition_in; order. + + rewrite IHr2 by auto; clear H1 IHl2 IHr2. + unfold Subset. intuition_in. + assert (InT a (Node l2 x2 r2 h2)) by auto; intuition_in; order. + assert (InT a (Node l2 x2 r2 h2)) by auto; intuition_in; order. +Qed. + +Lemma subset_spec : forall s1 s2 `{Ok s1, Ok s2}, + (subset s1 s2 = true <-> Subset s1 s2). +Proof. + induction s1 as [|l1 IHl1 x1 r1 IHr1 h1]; simpl; intros. + unfold Subset; intuition_in. + destruct s2 as [|l2 x2 r2 h2]; simpl; intros. + unfold Subset; intuition_in; try discriminate. + assert (H': InT x1 Leaf) by auto; inversion H'. + inv. + elim_compare x1 x2. + + rewrite <-andb_lazy_alt, andb_true_iff, IHl1, IHr1 by auto. + clear IHl1 IHr1. + unfold Subset; intuition_in. + assert (X.eq a x2) by order; intuition_in. + assert (InT a (Node l2 x2 r2 h2)) by auto; intuition_in; order. + assert (InT a (Node l2 x2 r2 h2)) by auto; intuition_in; order. + + rewrite <-andb_lazy_alt, andb_true_iff, IHr1 by auto. + rewrite (@subsetl_spec (subset l1) l1 x1 h1) by auto. + clear IHl1 IHr1. + unfold Subset; intuition_in. + assert (InT a (Node l2 x2 r2 h2)) by auto; intuition_in; order. + assert (InT a (Node l2 x2 r2 h2)) by auto; intuition_in; order. + + rewrite <-andb_lazy_alt, andb_true_iff, IHl1 by auto. + rewrite (@subsetr_spec (subset r1) r1 x1 h1) by auto. + clear IHl1 IHr1. + unfold Subset; intuition_in. + assert (InT a (Node l2 x2 r2 h2)) by auto; intuition_in; order. + assert (InT a (Node l2 x2 r2 h2)) by auto; intuition_in; order. +Qed. + + +(** * Comparison *) + +(** ** Relations [eq] and [lt] over trees *) + +Module L := MakeListOrdering X. + +Definition eq := Equal. +Instance eq_equiv : Equivalence eq. +Proof. firstorder. Qed. + +Lemma eq_Leq : forall s s', eq s s' <-> L.eq (elements s) (elements s'). +Proof. + unfold eq, Equal, L.eq; intros. + setoid_rewrite elements_spec1; firstorder. +Qed. + +Definition lt (s1 s2 : t) : Prop := + exists s1', exists s2', Ok s1' /\ Ok s2' /\ eq s1 s1' /\ eq s2 s2' + /\ L.lt (elements s1') (elements s2'). + +Instance lt_strorder : StrictOrder lt. +Proof. + split. + intros s (s1 & s2 & B1 & B2 & E1 & E2 & L). + assert (eqlistA X.eq (elements s1) (elements s2)). + apply SortA_equivlistA_eqlistA with (ltA:=X.lt); auto with *. + rewrite <- eq_Leq. transitivity s; auto. symmetry; auto. + rewrite H in L. + apply (StrictOrder_Irreflexive (elements s2)); auto. + intros s1 s2 s3 (s1' & s2' & B1 & B2 & E1 & E2 & L12) + (s2'' & s3' & B2' & B3 & E2' & E3 & L23). + exists s1', s3'; do 4 (split; trivial). + assert (eqlistA X.eq (elements s2') (elements s2'')). + apply SortA_equivlistA_eqlistA with (ltA:=X.lt); auto with *. + rewrite <- eq_Leq. transitivity s2; auto. symmetry; auto. + transitivity (elements s2'); auto. + rewrite H; auto. +Qed. + +Instance lt_compat : Proper (eq==>eq==>iff) lt. +Proof. + intros s1 s2 E12 s3 s4 E34. split. + intros (s1' & s3' & B1 & B3 & E1 & E3 & LT). + exists s1', s3'; do 2 (split; trivial). + split. transitivity s1; auto. symmetry; auto. + split; auto. transitivity s3; auto. symmetry; auto. + intros (s1' & s3' & B1 & B3 & E1 & E3 & LT). + exists s1', s3'; do 2 (split; trivial). + split. transitivity s2; auto. + split; auto. transitivity s4; auto. +Qed. + + +(** * Proof of the comparison algorithm *) + +(** [flatten_e e] returns the list of elements of [e] i.e. the list + of elements actually compared *) + +Fixpoint flatten_e (e : enumeration) : list elt := match e with + | End => nil + | More x t r => x :: elements t ++ flatten_e r + end. + +Lemma flatten_e_elements : + forall l x r h e, + elements l ++ flatten_e (More x r e) = elements (Node l x r h) ++ flatten_e e. +Proof. + intros; simpl; apply elements_node. +Qed. + +Lemma cons_1 : forall s e, + flatten_e (cons s e) = elements s ++ flatten_e e. +Proof. + induction s; simpl; auto; intros. + rewrite IHs1; apply flatten_e_elements. +Qed. + +(** Correctness of this comparison *) + +Definition Cmp c x y := CompSpec L.eq L.lt x y c. + +Local Hint Unfold Cmp flip. + +Lemma compare_end_Cmp : + forall e2, Cmp (compare_end e2) nil (flatten_e e2). +Proof. + destruct e2; simpl; constructor; auto. reflexivity. +Qed. + +Lemma compare_more_Cmp : forall x1 cont x2 r2 e2 l, + Cmp (cont (cons r2 e2)) l (elements r2 ++ flatten_e e2) -> + 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. +Qed. + +Lemma compare_cont_Cmp : forall s1 cont e2 l, + (forall e, Cmp (cont e) l (flatten_e e)) -> + Cmp (compare_cont s1 cont e2) (elements s1 ++ l) (flatten_e e2). +Proof. + induction s1 as [|l1 Hl1 x1 r1 Hr1 h1]; simpl; intros; auto. + rewrite <- elements_node; simpl. + apply Hl1; auto. clear e2. intros [|x2 r2 e2]. + simpl; auto. + apply compare_more_Cmp. + rewrite <- cons_1; auto. +Qed. + +Lemma compare_Cmp : forall s1 s2, + Cmp (compare s1 s2) (elements s1) (elements s2). +Proof. + intros; unfold compare. + rewrite (app_nil_end (elements s1)). + replace (elements s2) with (flatten_e (cons s2 End)) by + (rewrite cons_1; simpl; rewrite <- app_nil_end; auto). + apply compare_cont_Cmp; auto. + intros. + apply compare_end_Cmp; auto. +Qed. + +Lemma compare_spec : forall s1 s2 `{Ok s1, Ok s2}, + CompSpec eq lt s1 s2 (compare s1 s2). +Proof. + intros. + destruct (compare_Cmp s1 s2); constructor. + rewrite eq_Leq; auto. + intros; exists s1, s2; repeat split; auto. + intros; exists s2, s1; repeat split; auto. +Qed. + + +(** * Equality test *) + +Lemma equal_spec : forall s1 s2 `{Ok s1, Ok s2}, + equal s1 s2 = true <-> eq s1 s2. +Proof. +unfold equal; intros s1 s2 B1 B2. +destruct (@compare_spec s1 s2 B1 B2) as [H|H|H]; + split; intros H'; auto; try discriminate. +rewrite H' in H. elim (StrictOrder_Irreflexive s2); auto. +rewrite H' in H. elim (StrictOrder_Irreflexive s2); auto. +Qed. + +End MakeRaw. + + + +(** * Encapsulation + + Now, in order to really provide a functor implementing [S], we + need to encapsulate everything into a type of binary search trees. + They also happen to be well-balanced, but this has no influence + on the correctness of operations, so we won't state this here, + see [MSetFullAVL] if you need more than just the MSet interface. +*) + +Module IntMake (I:Int)(X: OrderedType) <: S with Module E := X. + Module Raw := MakeRaw I X. + Include Raw2Sets X Raw. +End IntMake. + +(* For concrete use inside Coq, we propose an instantiation of [Int] by [Z]. *) + +Module Make (X: OrderedType) <: S with Module E := X + :=IntMake(Z_as_Int)(X). diff --git a/theories/MSets/MSetDecide.v b/theories/MSets/MSetDecide.v new file mode 100644 index 00000000..07c9955a --- /dev/null +++ b/theories/MSets/MSetDecide.v @@ -0,0 +1,880 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) +(* \VV/ *************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(***********************************************************************) + +(* $Id$ *) + +(**************************************************************) +(* MSetDecide.v *) +(* *) +(* Author: Aaron Bohannon *) +(**************************************************************) + +(** This file implements a decision procedure for a certain + class of propositions involving finite sets. *) + +Require Import Decidable DecidableTypeEx MSetFacts. + +(** First, a version for Weak Sets in functorial presentation *) + +Module WDecideOn (E : DecidableType)(Import M : WSetsOn E). + Module F := MSetFacts.WFactsOn E M. + +(** * Overview + This functor defines the tactic [fsetdec], which will + solve any valid goal of the form +<< + forall s1 ... sn, + forall x1 ... xm, + P1 -> ... -> Pk -> P +>> + where [P]'s are defined by the grammar: +<< + +P ::= +| Q +| Empty F +| Subset F F' +| Equal F F' + +Q ::= +| E.eq X X' +| In X F +| Q /\ Q' +| Q \/ Q' +| Q -> Q' +| Q <-> Q' +| ~ Q +| True +| False + +F ::= +| S +| empty +| singleton X +| add X F +| remove X F +| union F F' +| inter F F' +| diff F F' + +X ::= x1 | ... | xm +S ::= s1 | ... | sn + +>> + +The tactic will also work on some goals that vary slightly from +the above form: +- The variables and hypotheses may be mixed in any order and may + have already been introduced into the context. Moreover, + there may be additional, unrelated hypotheses mixed in (these + will be ignored). +- A conjunction of hypotheses will be handled as easily as + separate hypotheses, i.e., [P1 /\ P2 -> P] can be solved iff + [P1 -> P2 -> P] can be solved. +- [fsetdec] should solve any goal if the MSet-related hypotheses + are contradictory. +- [fsetdec] will first perform any necessary zeta and beta + reductions and will invoke [subst] to eliminate any Coq + equalities between finite sets or their elements. +- If [E.eq] is convertible with Coq's equality, it will not + matter which one is used in the hypotheses or conclusion. +- The tactic can solve goals where the finite sets or set + elements are expressed by Coq terms that are more complicated + than variables. However, non-local definitions are not + expanded, and Coq equalities between non-variable terms are + not used. For example, this goal will be solved: +<< + forall (f : t -> t), + forall (g : elt -> elt), + forall (s1 s2 : t), + forall (x1 x2 : elt), + Equal s1 (f s2) -> + E.eq x1 (g (g x2)) -> + In x1 s1 -> + In (g (g x2)) (f s2) +>> + This one will not be solved: +<< + forall (f : t -> t), + forall (g : elt -> elt), + forall (s1 s2 : t), + forall (x1 x2 : elt), + Equal s1 (f s2) -> + E.eq x1 (g x2) -> + In x1 s1 -> + g x2 = g (g x2) -> + In (g (g x2)) (f s2) +>> +*) + + (** * Facts and Tactics for Propositional Logic + These lemmas and tactics are in a module so that they do + not affect the namespace if you import the enclosing + module [Decide]. *) + Module MSetLogicalFacts. + Require Export Decidable. + Require Export Setoid. + + (** ** Lemmas and Tactics About Decidable Propositions *) + + (** ** Propositional Equivalences Involving Negation + These are all written with the unfolded form of + negation, since I am not sure if setoid rewriting will + always perform conversion. *) + + (** ** Tactics for Negations *) + + Tactic Notation "fold" "any" "not" := + repeat ( + match goal with + | H: context [?P -> False] |- _ => + fold (~ P) in H + | |- context [?P -> False] => + fold (~ P) + end). + + (** [push not using db] will pushes all negations to the + leaves of propositions in the goal, using the lemmas in + [db] to assist in checking the decidability of the + propositions involved. If [using db] is omitted, then + [core] will be used. Additional versions are provided + to manipulate the hypotheses or the hypotheses and goal + together. + + XXX: This tactic and the similar subsequent ones should + have been defined using [autorewrite]. However, dealing + with multiples rewrite sites and side-conditions is + done more cleverly with the following explicit + analysis of goals. *) + + Ltac or_not_l_iff P Q tac := + (rewrite (or_not_l_iff_1 P Q) by tac) || + (rewrite (or_not_l_iff_2 P Q) by tac). + + Ltac or_not_r_iff P Q tac := + (rewrite (or_not_r_iff_1 P Q) by tac) || + (rewrite (or_not_r_iff_2 P Q) by tac). + + Ltac or_not_l_iff_in P Q H tac := + (rewrite (or_not_l_iff_1 P Q) in H by tac) || + (rewrite (or_not_l_iff_2 P Q) in H by tac). + + Ltac or_not_r_iff_in P Q H tac := + (rewrite (or_not_r_iff_1 P Q) in H by tac) || + (rewrite (or_not_r_iff_2 P Q) in H by tac). + + Tactic Notation "push" "not" "using" ident(db) := + let dec := solve_decidable using db in + unfold not, iff; + repeat ( + match goal with + | |- context [True -> False] => rewrite not_true_iff + | |- context [False -> False] => rewrite not_false_iff + | |- context [(?P -> False) -> False] => rewrite (not_not_iff P) by dec + | |- context [(?P -> False) -> (?Q -> False)] => + rewrite (contrapositive P Q) by dec + | |- context [(?P -> False) \/ ?Q] => or_not_l_iff P Q dec + | |- context [?P \/ (?Q -> False)] => or_not_r_iff P Q dec + | |- context [(?P -> False) -> ?Q] => rewrite (imp_not_l P Q) by dec + | |- context [?P \/ ?Q -> False] => rewrite (not_or_iff P Q) + | |- context [?P /\ ?Q -> False] => rewrite (not_and_iff P Q) + | |- context [(?P -> ?Q) -> False] => rewrite (not_imp_iff P Q) by dec + end); + fold any not. + + Tactic Notation "push" "not" := + push not using core. + + Tactic Notation + "push" "not" "in" "*" "|-" "using" ident(db) := + let dec := solve_decidable using db in + unfold not, iff in * |-; + repeat ( + match goal with + | H: context [True -> False] |- _ => rewrite not_true_iff in H + | H: context [False -> False] |- _ => rewrite not_false_iff in H + | H: context [(?P -> False) -> False] |- _ => + rewrite (not_not_iff P) in H by dec + | H: context [(?P -> False) -> (?Q -> False)] |- _ => + rewrite (contrapositive P Q) in H by dec + | H: context [(?P -> False) \/ ?Q] |- _ => or_not_l_iff_in P Q H dec + | H: context [?P \/ (?Q -> False)] |- _ => or_not_r_iff_in P Q H dec + | H: context [(?P -> False) -> ?Q] |- _ => + rewrite (imp_not_l P Q) in H by dec + | H: context [?P \/ ?Q -> False] |- _ => rewrite (not_or_iff P Q) in H + | H: context [?P /\ ?Q -> False] |- _ => rewrite (not_and_iff P Q) in H + | H: context [(?P -> ?Q) -> False] |- _ => + rewrite (not_imp_iff P Q) in H by dec + end); + fold any not. + + Tactic Notation "push" "not" "in" "*" "|-" := + push not in * |- using core. + + Tactic Notation "push" "not" "in" "*" "using" ident(db) := + push not using db; push not in * |- using db. + Tactic Notation "push" "not" "in" "*" := + push not in * using core. + + (** A simple test case to see how this works. *) + Lemma test_push : forall P Q R : Prop, + decidable P -> + decidable Q -> + (~ True) -> + (~ False) -> + (~ ~ P) -> + (~ (P /\ Q) -> ~ R) -> + ((P /\ Q) \/ ~ R) -> + (~ (P /\ Q) \/ R) -> + (R \/ ~ (P /\ Q)) -> + (~ R \/ (P /\ Q)) -> + (~ P -> R) -> + (~ ((R -> P) \/ (Q -> R))) -> + (~ (P /\ R)) -> + (~ (P -> R)) -> + True. + Proof. + intros. push not in *. + (* note that ~(R->P) remains (since R isnt decidable) *) + tauto. + Qed. + + (** [pull not using db] will pull as many negations as + possible toward the top of the propositions in the goal, + using the lemmas in [db] to assist in checking the + decidability of the propositions involved. If [using + db] is omitted, then [core] will be used. Additional + versions are provided to manipulate the hypotheses or + the hypotheses and goal together. *) + + Tactic Notation "pull" "not" "using" ident(db) := + let dec := solve_decidable using db in + unfold not, iff; + repeat ( + match goal with + | |- context [True -> False] => rewrite not_true_iff + | |- context [False -> False] => rewrite not_false_iff + | |- context [(?P -> False) -> False] => rewrite (not_not_iff P) by dec + | |- context [(?P -> False) -> (?Q -> False)] => + rewrite (contrapositive P Q) by dec + | |- context [(?P -> False) \/ ?Q] => or_not_l_iff P Q dec + | |- context [?P \/ (?Q -> False)] => or_not_r_iff P Q dec + | |- context [(?P -> False) -> ?Q] => rewrite (imp_not_l P Q) by dec + | |- context [(?P -> False) /\ (?Q -> False)] => + rewrite <- (not_or_iff P Q) + | |- context [?P -> ?Q -> False] => rewrite <- (not_and_iff P Q) + | |- context [?P /\ (?Q -> False)] => rewrite <- (not_imp_iff P Q) by dec + | |- context [(?Q -> False) /\ ?P] => + rewrite <- (not_imp_rev_iff P Q) by dec + end); + fold any not. + + Tactic Notation "pull" "not" := + pull not using core. + + Tactic Notation + "pull" "not" "in" "*" "|-" "using" ident(db) := + let dec := solve_decidable using db in + unfold not, iff in * |-; + repeat ( + match goal with + | H: context [True -> False] |- _ => rewrite not_true_iff in H + | H: context [False -> False] |- _ => rewrite not_false_iff in H + | H: context [(?P -> False) -> False] |- _ => + rewrite (not_not_iff P) in H by dec + | H: context [(?P -> False) -> (?Q -> False)] |- _ => + rewrite (contrapositive P Q) in H by dec + | H: context [(?P -> False) \/ ?Q] |- _ => or_not_l_iff_in P Q H dec + | H: context [?P \/ (?Q -> False)] |- _ => or_not_r_iff_in P Q H dec + | H: context [(?P -> False) -> ?Q] |- _ => + rewrite (imp_not_l P Q) in H by dec + | H: context [(?P -> False) /\ (?Q -> False)] |- _ => + rewrite <- (not_or_iff P Q) in H + | H: context [?P -> ?Q -> False] |- _ => + rewrite <- (not_and_iff P Q) in H + | H: context [?P /\ (?Q -> False)] |- _ => + rewrite <- (not_imp_iff P Q) in H by dec + | H: context [(?Q -> False) /\ ?P] |- _ => + rewrite <- (not_imp_rev_iff P Q) in H by dec + end); + fold any not. + + Tactic Notation "pull" "not" "in" "*" "|-" := + pull not in * |- using core. + + Tactic Notation "pull" "not" "in" "*" "using" ident(db) := + pull not using db; pull not in * |- using db. + Tactic Notation "pull" "not" "in" "*" := + pull not in * using core. + + (** A simple test case to see how this works. *) + Lemma test_pull : forall P Q R : Prop, + decidable P -> + decidable Q -> + (~ True) -> + (~ False) -> + (~ ~ P) -> + (~ (P /\ Q) -> ~ R) -> + ((P /\ Q) \/ ~ R) -> + (~ (P /\ Q) \/ R) -> + (R \/ ~ (P /\ Q)) -> + (~ R \/ (P /\ Q)) -> + (~ P -> R) -> + (~ (R -> P) /\ ~ (Q -> R)) -> + (~ P \/ ~ R) -> + (P /\ ~ R) -> + (~ R /\ P) -> + True. + Proof. + intros. pull not in *. tauto. + Qed. + + End MSetLogicalFacts. + Import MSetLogicalFacts. + + (** * Auxiliary Tactics + Again, these lemmas and tactics are in a module so that + they do not affect the namespace if you import the + enclosing module [Decide]. *) + Module MSetDecideAuxiliary. + + (** ** Generic Tactics + We begin by defining a few generic, useful tactics. *) + + (** remove logical hypothesis inter-dependencies (fix #2136). *) + + Ltac no_logical_interdep := + match goal with + | H : ?P |- _ => + match type of P with + | Prop => + match goal with H' : context [ H ] |- _ => clear dependent H' end + | _ => fail + end; no_logical_interdep + | _ => idtac + end. + + (** [if t then t1 else t2] executes [t] and, if it does not + fail, then [t1] will be applied to all subgoals + produced. If [t] fails, then [t2] is executed. *) + Tactic Notation + "if" tactic(t) + "then" tactic(t1) + "else" tactic(t2) := + first [ t; first [ t1 | fail 2 ] | t2 ]. + + (** [prop P holds by t] succeeds (but does not modify the + goal or context) if the proposition [P] can be proved by + [t] in the current context. Otherwise, the tactic + fails. *) + Tactic Notation "prop" constr(P) "holds" "by" tactic(t) := + let H := fresh in + assert P as H by t; + clear H. + + (** This tactic acts just like [assert ... by ...] but will + fail if the context already contains the proposition. *) + Tactic Notation "assert" "new" constr(e) "by" tactic(t) := + match goal with + | H: e |- _ => fail 1 + | _ => assert e by t + end. + + (** [subst++] is similar to [subst] except that + - it never fails (as [subst] does on recursive + equations), + - it substitutes locally defined variable for their + definitions, + - it performs beta reductions everywhere, which may + arise after substituting a locally defined function + for its definition. + *) + Tactic Notation "subst" "++" := + repeat ( + match goal with + | x : _ |- _ => subst x + end); + cbv zeta beta in *. + + (** [decompose records] calls [decompose record H] on every + relevant hypothesis [H]. *) + Tactic Notation "decompose" "records" := + repeat ( + match goal with + | H: _ |- _ => progress (decompose record H); clear H + end). + + (** ** Discarding Irrelevant Hypotheses + We will want to clear the context of any + non-MSet-related hypotheses in order to increase the + speed of the tactic. To do this, we will need to be + able to decide which are relevant. We do this by making + a simple inductive definition classifying the + propositions of interest. *) + + Inductive MSet_elt_Prop : Prop -> Prop := + | eq_Prop : forall (S : Type) (x y : S), + MSet_elt_Prop (x = y) + | eq_elt_prop : forall x y, + MSet_elt_Prop (E.eq x y) + | In_elt_prop : forall x s, + MSet_elt_Prop (In x s) + | True_elt_prop : + MSet_elt_Prop True + | False_elt_prop : + MSet_elt_Prop False + | conj_elt_prop : forall P Q, + MSet_elt_Prop P -> + MSet_elt_Prop Q -> + MSet_elt_Prop (P /\ Q) + | disj_elt_prop : forall P Q, + MSet_elt_Prop P -> + MSet_elt_Prop Q -> + MSet_elt_Prop (P \/ Q) + | impl_elt_prop : forall P Q, + MSet_elt_Prop P -> + MSet_elt_Prop Q -> + MSet_elt_Prop (P -> Q) + | not_elt_prop : forall P, + MSet_elt_Prop P -> + MSet_elt_Prop (~ P). + + Inductive MSet_Prop : Prop -> Prop := + | elt_MSet_Prop : forall P, + MSet_elt_Prop P -> + MSet_Prop P + | Empty_MSet_Prop : forall s, + MSet_Prop (Empty s) + | Subset_MSet_Prop : forall s1 s2, + MSet_Prop (Subset s1 s2) + | Equal_MSet_Prop : forall s1 s2, + MSet_Prop (Equal s1 s2). + + (** Here is the tactic that will throw away hypotheses that + are not useful (for the intended scope of the [fsetdec] + tactic). *) + Hint Constructors MSet_elt_Prop MSet_Prop : MSet_Prop. + Ltac discard_nonMSet := + decompose records; + repeat ( + match goal with + | H : ?P |- _ => + if prop (MSet_Prop P) holds by + (auto 100 with MSet_Prop) + then fail + else clear H + end). + + (** ** Turning Set Operators into Propositional Connectives + The lemmas from [MSetFacts] will be used to break down + set operations into propositional formulas built over + the predicates [In] and [E.eq] applied only to + variables. We are going to use them with [autorewrite]. + *) + + Hint Rewrite + F.empty_iff F.singleton_iff F.add_iff F.remove_iff + F.union_iff F.inter_iff F.diff_iff + : set_simpl. + + (** ** Decidability of MSet Propositions *) + + (** [In] is decidable. *) + Lemma dec_In : forall x s, + decidable (In x s). + Proof. + red; intros; generalize (F.mem_iff s x); case (mem x s); intuition. + Qed. + + (** [E.eq] is decidable. *) + Lemma dec_eq : forall (x y : E.t), + decidable (E.eq x y). + Proof. + red; intros x y; destruct (E.eq_dec x y); auto. + Qed. + + (** The hint database [MSet_decidability] will be given to + the [push_neg] tactic from the module [Negation]. *) + Hint Resolve dec_In dec_eq : MSet_decidability. + + (** ** Normalizing Propositions About Equality + We have to deal with the fact that [E.eq] may be + convertible with Coq's equality. Thus, we will find the + following tactics useful to replace one form with the + other everywhere. *) + + (** The next tactic, [Logic_eq_to_E_eq], mentions the term + [E.t]; thus, we must ensure that [E.t] is used in favor + of any other convertible but syntactically distinct + term. *) + Ltac change_to_E_t := + repeat ( + match goal with + | H : ?T |- _ => + progress (change T with E.t in H); + repeat ( + match goal with + | J : _ |- _ => progress (change T with E.t in J) + | |- _ => progress (change T with E.t) + end ) + | H : forall x : ?T, _ |- _ => + progress (change T with E.t in H); + repeat ( + match goal with + | J : _ |- _ => progress (change T with E.t in J) + | |- _ => progress (change T with E.t) + end ) + end). + + (** These two tactics take us from Coq's built-in equality + to [E.eq] (and vice versa) when possible. *) + + Ltac Logic_eq_to_E_eq := + repeat ( + match goal with + | H: _ |- _ => + progress (change (@Logic.eq E.t) with E.eq in H) + | |- _ => + progress (change (@Logic.eq E.t) with E.eq) + end). + + Ltac E_eq_to_Logic_eq := + repeat ( + match goal with + | H: _ |- _ => + progress (change E.eq with (@Logic.eq E.t) in H) + | |- _ => + progress (change E.eq with (@Logic.eq E.t)) + end). + + (** This tactic works like the built-in tactic [subst], but + at the level of set element equality (which may not be + the convertible with Coq's equality). *) + Ltac substMSet := + repeat ( + match goal with + | H: E.eq ?x ?y |- _ => rewrite H in *; clear H + end). + + (** ** Considering Decidability of Base Propositions + This tactic adds assertions about the decidability of + [E.eq] and [In] to the context. This is necessary for + the completeness of the [fsetdec] tactic. However, in + order to minimize the cost of proof search, we should be + careful to not add more than we need. Once negations + have been pushed to the leaves of the propositions, we + only need to worry about decidability for those base + propositions that appear in a negated form. *) + Ltac assert_decidability := + (** We actually don't want these rules to fire if the + syntactic context in the patterns below is trivially + empty, but we'll just do some clean-up at the + afterward. *) + repeat ( + match goal with + | H: context [~ E.eq ?x ?y] |- _ => + assert new (E.eq x y \/ ~ E.eq x y) by (apply dec_eq) + | H: context [~ In ?x ?s] |- _ => + assert new (In x s \/ ~ In x s) by (apply dec_In) + | |- context [~ E.eq ?x ?y] => + assert new (E.eq x y \/ ~ E.eq x y) by (apply dec_eq) + | |- context [~ In ?x ?s] => + assert new (In x s \/ ~ In x s) by (apply dec_In) + end); + (** Now we eliminate the useless facts we added (because + they would likely be very harmful to performance). *) + repeat ( + match goal with + | _: ~ ?P, H : ?P \/ ~ ?P |- _ => clear H + end). + + (** ** Handling [Empty], [Subset], and [Equal] + This tactic instantiates universally quantified + hypotheses (which arise from the unfolding of [Empty], + [Subset], and [Equal]) for each of the set element + expressions that is involved in some membership or + equality fact. Then it throws away those hypotheses, + which should no longer be needed. *) + Ltac inst_MSet_hypotheses := + repeat ( + match goal with + | H : forall a : E.t, _, + _ : context [ In ?x _ ] |- _ => + let P := type of (H x) in + assert new P by (exact (H x)) + | H : forall a : E.t, _ + |- context [ In ?x _ ] => + let P := type of (H x) in + assert new P by (exact (H x)) + | H : forall a : E.t, _, + _ : context [ E.eq ?x _ ] |- _ => + let P := type of (H x) in + assert new P by (exact (H x)) + | H : forall a : E.t, _ + |- context [ E.eq ?x _ ] => + let P := type of (H x) in + assert new P by (exact (H x)) + | H : forall a : E.t, _, + _ : context [ E.eq _ ?x ] |- _ => + let P := type of (H x) in + assert new P by (exact (H x)) + | H : forall a : E.t, _ + |- context [ E.eq _ ?x ] => + let P := type of (H x) in + assert new P by (exact (H x)) + end); + repeat ( + match goal with + | H : forall a : E.t, _ |- _ => + clear H + end). + + (** ** The Core [fsetdec] Auxiliary Tactics *) + + (** Here is the crux of the proof search. Recursion through + [intuition]! (This will terminate if I correctly + understand the behavior of [intuition].) *) + Ltac fsetdec_rec := + try (match goal with + | H: E.eq ?x ?x -> False |- _ => destruct H + end); + (reflexivity || + contradiction || + (progress substMSet; intuition fsetdec_rec)). + + (** If we add [unfold Empty, Subset, Equal in *; intros;] to + the beginning of this tactic, it will satisfy the same + specification as the [fsetdec] tactic; however, it will + be much slower than necessary without the pre-processing + done by the wrapper tactic [fsetdec]. *) + Ltac fsetdec_body := + inst_MSet_hypotheses; + autorewrite with set_simpl in *; + push not in * using MSet_decidability; + substMSet; + assert_decidability; + auto using (@Equivalence_Reflexive _ _ E.eq_equiv); + (intuition fsetdec_rec) || + fail 1 + "because the goal is beyond the scope of this tactic". + + End MSetDecideAuxiliary. + Import MSetDecideAuxiliary. + + (** * The [fsetdec] Tactic + Here is the top-level tactic (the only one intended for + clients of this library). It's specification is given at + the top of the file. *) + Ltac fsetdec := + (** We first unfold any occurrences of [iff]. *) + unfold iff in *; + (** We fold occurrences of [not] because it is better for + [intros] to leave us with a goal of [~ P] than a goal of + [False]. *) + fold any not; intros; + (** We remove dependencies to logical hypothesis. This way, + later "clear" will work nicely (see bug #2136) *) + no_logical_interdep; + (** Now we decompose conjunctions, which will allow the + [discard_nonMSet] and [assert_decidability] tactics to + do a much better job. *) + decompose records; + discard_nonMSet; + (** We unfold these defined propositions on finite sets. If + our goal was one of them, then have one more item to + introduce now. *) + unfold Empty, Subset, Equal in *; intros; + (** We now want to get rid of all uses of [=] in favor of + [E.eq]. However, the best way to eliminate a [=] is in + the context is with [subst], so we will try that first. + In fact, we may as well convert uses of [E.eq] into [=] + when possible before we do [subst] so that we can even + more mileage out of it. Then we will convert all + remaining uses of [=] back to [E.eq] when possible. We + use [change_to_E_t] to ensure that we have a canonical + name for set elements, so that [Logic_eq_to_E_eq] will + work properly. *) + change_to_E_t; E_eq_to_Logic_eq; subst++; Logic_eq_to_E_eq; + (** The next optimization is to swap a negated goal with a + negated hypothesis when possible. Any swap will improve + performance by eliminating the total number of + negations, but we will get the maximum benefit if we + swap the goal with a hypotheses mentioning the same set + element, so we try that first. If we reach the fourth + branch below, we attempt any swap. However, to maintain + completeness of this tactic, we can only perform such a + swap with a decidable proposition; hence, we first test + whether the hypothesis is an [MSet_elt_Prop], noting + that any [MSet_elt_Prop] is decidable. *) + pull not using MSet_decidability; + unfold not in *; + match goal with + | H: (In ?x ?r) -> False |- (In ?x ?s) -> False => + contradict H; fsetdec_body + | H: (In ?x ?r) -> False |- (E.eq ?x ?y) -> False => + contradict H; fsetdec_body + | H: (In ?x ?r) -> False |- (E.eq ?y ?x) -> False => + contradict H; fsetdec_body + | H: ?P -> False |- ?Q -> False => + if prop (MSet_elt_Prop P) holds by + (auto 100 with MSet_Prop) + then (contradict H; fsetdec_body) + else fsetdec_body + | |- _ => + fsetdec_body + end. + + (** * Examples *) + + Module MSetDecideTestCases. + + Lemma test_eq_trans_1 : forall x y z s, + E.eq x y -> + ~ ~ E.eq z y -> + In x s -> + In z s. + Proof. fsetdec. Qed. + + Lemma test_eq_trans_2 : forall x y z r s, + In x (singleton y) -> + ~ In z r -> + ~ ~ In z (add y r) -> + In x s -> + In z s. + Proof. fsetdec. Qed. + + Lemma test_eq_neq_trans_1 : forall w x y z s, + E.eq x w -> + ~ ~ E.eq x y -> + ~ E.eq y z -> + In w s -> + In w (remove z s). + Proof. fsetdec. Qed. + + Lemma test_eq_neq_trans_2 : forall w x y z r1 r2 s, + In x (singleton w) -> + ~ In x r1 -> + In x (add y r1) -> + In y r2 -> + In y (remove z r2) -> + In w s -> + In w (remove z s). + Proof. fsetdec. Qed. + + Lemma test_In_singleton : forall x, + In x (singleton x). + Proof. fsetdec. Qed. + + Lemma test_add_In : forall x y s, + In x (add y s) -> + ~ E.eq x y -> + In x s. + Proof. fsetdec. Qed. + + Lemma test_Subset_add_remove : forall x s, + s [<=] (add x (remove x s)). + Proof. fsetdec. Qed. + + Lemma test_eq_disjunction : forall w x y z, + In w (add x (add y (singleton z))) -> + E.eq w x \/ E.eq w y \/ E.eq w z. + Proof. fsetdec. Qed. + + Lemma test_not_In_disj : forall x y s1 s2 s3 s4, + ~ In x (union s1 (union s2 (union s3 (add y s4)))) -> + ~ (In x s1 \/ In x s4 \/ E.eq y x). + Proof. fsetdec. Qed. + + Lemma test_not_In_conj : forall x y s1 s2 s3 s4, + ~ In x (union s1 (union s2 (union s3 (add y s4)))) -> + ~ In x s1 /\ ~ In x s4 /\ ~ E.eq y x. + Proof. fsetdec. Qed. + + Lemma test_iff_conj : forall a x s s', + (In a s' <-> E.eq x a \/ In a s) -> + (In a s' <-> In a (add x s)). + Proof. fsetdec. Qed. + + Lemma test_set_ops_1 : forall x q r s, + (singleton x) [<=] s -> + Empty (union q r) -> + Empty (inter (diff s q) (diff s r)) -> + ~ In x s. + Proof. fsetdec. Qed. + + Lemma eq_chain_test : forall x1 x2 x3 x4 s1 s2 s3 s4, + Empty s1 -> + In x2 (add x1 s1) -> + In x3 s2 -> + ~ In x3 (remove x2 s2) -> + ~ In x4 s3 -> + In x4 (add x3 s3) -> + In x1 s4 -> + Subset (add x4 s4) s4. + Proof. fsetdec. Qed. + + Lemma test_too_complex : forall x y z r s, + E.eq x y -> + (In x (singleton y) -> r [<=] s) -> + In z r -> + In z s. + Proof. + (** [fsetdec] is not intended to solve this directly. *) + intros until s; intros Heq H Hr; lapply H; fsetdec. + Qed. + + Lemma function_test_1 : + forall (f : t -> t), + forall (g : elt -> elt), + forall (s1 s2 : t), + forall (x1 x2 : elt), + Equal s1 (f s2) -> + E.eq x1 (g (g x2)) -> + In x1 s1 -> + In (g (g x2)) (f s2). + Proof. fsetdec. Qed. + + Lemma function_test_2 : + forall (f : t -> t), + forall (g : elt -> elt), + forall (s1 s2 : t), + forall (x1 x2 : elt), + Equal s1 (f s2) -> + E.eq x1 (g x2) -> + In x1 s1 -> + g x2 = g (g x2) -> + In (g (g x2)) (f s2). + Proof. + (** [fsetdec] is not intended to solve this directly. *) + intros until 3. intros g_eq. rewrite <- g_eq. fsetdec. + Qed. + + Lemma test_baydemir : + forall (f : t -> t), + forall (s : t), + forall (x y : elt), + In x (add y (f s)) -> + ~ E.eq x y -> + In x (f s). + Proof. + fsetdec. + Qed. + + End MSetDecideTestCases. + +End WDecideOn. + +Require Import MSetInterface. + +(** Now comes variants for self-contained weak sets and for full sets. + For these variants, only one argument is necessary. Thanks to + the subtyping [WS<=S], the [Decide] functor which is meant to be + used on modules [(M:S)] can simply be an alias of [WDecide]. *) + +Module WDecide (M:WSets) := WDecideOn M.E M. +Module Decide := WDecide. diff --git a/theories/MSets/MSetEqProperties.v b/theories/MSets/MSetEqProperties.v new file mode 100644 index 00000000..fe6c3c79 --- /dev/null +++ b/theories/MSets/MSetEqProperties.v @@ -0,0 +1,936 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) +(* \VV/ *************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(***********************************************************************) + +(* $Id$ *) + +(** * Finite sets library *) + +(** This module proves many properties of finite sets that + are consequences of the axiomatization in [FsetInterface] + Contrary to the functor in [FsetProperties] it uses + sets operations instead of predicates over sets, i.e. + [mem x s=true] instead of [In x s], + [equal s s'=true] instead of [Equal s s'], etc. *) + +Require Import MSetProperties Zerob Sumbool Omega DecidableTypeEx. + +Module WEqPropertiesOn (Import E:DecidableType)(M:WSetsOn E). +Module Import MP := WPropertiesOn E M. +Import FM Dec.F. +Import M. + +Definition Add := MP.Add. + +Section BasicProperties. + +(** Some old specifications written with boolean equalities. *) + +Variable s s' s'': t. +Variable x y z : elt. + +Lemma mem_eq: + E.eq x y -> mem x s=mem y s. +Proof. +intro H; rewrite H; auto. +Qed. + +Lemma equal_mem_1: + (forall a, mem a s=mem a s') -> equal s s'=true. +Proof. +intros; apply equal_1; unfold Equal; intros. +do 2 rewrite mem_iff; rewrite H; tauto. +Qed. + +Lemma equal_mem_2: + equal s s'=true -> forall a, mem a s=mem a s'. +Proof. +intros; rewrite (equal_2 H); auto. +Qed. + +Lemma subset_mem_1: + (forall a, mem a s=true->mem a s'=true) -> subset s s'=true. +Proof. +intros; apply subset_1; unfold Subset; intros a. +do 2 rewrite mem_iff; auto. +Qed. + +Lemma subset_mem_2: + subset s s'=true -> forall a, mem a s=true -> mem a s'=true. +Proof. +intros H a; do 2 rewrite <- mem_iff; apply subset_2; auto. +Qed. + +Lemma empty_mem: mem x empty=false. +Proof. +rewrite <- not_mem_iff; auto with set. +Qed. + +Lemma is_empty_equal_empty: is_empty s = equal s empty. +Proof. +apply bool_1; split; intros. +auto with set. +rewrite <- is_empty_iff; auto with set. +Qed. + +Lemma choose_mem_1: choose s=Some x -> mem x s=true. +Proof. +auto with set. +Qed. + +Lemma choose_mem_2: choose s=None -> is_empty s=true. +Proof. +auto with set. +Qed. + +Lemma add_mem_1: mem x (add x s)=true. +Proof. +auto with set relations. +Qed. + +Lemma add_mem_2: ~E.eq x y -> mem y (add x s)=mem y s. +Proof. +apply add_neq_b. +Qed. + +Lemma remove_mem_1: mem x (remove x s)=false. +Proof. +rewrite <- not_mem_iff; auto with set relations. +Qed. + +Lemma remove_mem_2: ~E.eq x y -> mem y (remove x s)=mem y s. +Proof. +apply remove_neq_b. +Qed. + +Lemma singleton_equal_add: + equal (singleton x) (add x empty)=true. +Proof. +rewrite (singleton_equal_add x); auto with set. +Qed. + +Lemma union_mem: + mem x (union s s')=mem x s || mem x s'. +Proof. +apply union_b. +Qed. + +Lemma inter_mem: + mem x (inter s s')=mem x s && mem x s'. +Proof. +apply inter_b. +Qed. + +Lemma diff_mem: + mem x (diff s s')=mem x s && negb (mem x s'). +Proof. +apply diff_b. +Qed. + +(** properties of [mem] *) + +Lemma mem_3 : ~In x s -> mem x s=false. +Proof. +intros; rewrite <- not_mem_iff; auto. +Qed. + +Lemma mem_4 : mem x s=false -> ~In x s. +Proof. +intros; rewrite not_mem_iff; auto. +Qed. + +(** Properties of [equal] *) + +Lemma equal_refl: equal s s=true. +Proof. +auto with set. +Qed. + +Lemma equal_sym: equal s s'=equal s' s. +Proof. +intros; apply bool_1; do 2 rewrite <- equal_iff; intuition. +Qed. + +Lemma equal_trans: + equal s s'=true -> equal s' s''=true -> equal s s''=true. +Proof. +intros; rewrite (equal_2 H); auto. +Qed. + +Lemma equal_equal: + equal s s'=true -> equal s s''=equal s' s''. +Proof. +intros; rewrite (equal_2 H); auto. +Qed. + +Lemma equal_cardinal: + equal s s'=true -> cardinal s=cardinal s'. +Proof. +auto with set. +Qed. + +(* Properties of [subset] *) + +Lemma subset_refl: subset s s=true. +Proof. +auto with set. +Qed. + +Lemma subset_antisym: + subset s s'=true -> subset s' s=true -> equal s s'=true. +Proof. +auto with set. +Qed. + +Lemma subset_trans: + subset s s'=true -> subset s' s''=true -> subset s s''=true. +Proof. +do 3 rewrite <- subset_iff; intros. +apply subset_trans with s'; auto. +Qed. + +Lemma subset_equal: + equal s s'=true -> subset s s'=true. +Proof. +auto with set. +Qed. + +(** Properties of [choose] *) + +Lemma choose_mem_3: + is_empty s=false -> {x:elt|choose s=Some x /\ mem x s=true}. +Proof. +intros. +generalize (@choose_1 s) (@choose_2 s). +destruct (choose s);intros. +exists e;auto with set. +generalize (H1 (refl_equal None)); clear H1. +intros; rewrite (is_empty_1 H1) in H; discriminate. +Qed. + +Lemma choose_mem_4: choose empty=None. +Proof. +generalize (@choose_1 empty). +case (@choose empty);intros;auto. +elim (@empty_1 e); auto. +Qed. + +(** Properties of [add] *) + +Lemma add_mem_3: + mem y s=true -> mem y (add x s)=true. +Proof. +auto with set. +Qed. + +Lemma add_equal: + mem x s=true -> equal (add x s) s=true. +Proof. +auto with set. +Qed. + +(** Properties of [remove] *) + +Lemma remove_mem_3: + mem y (remove x s)=true -> mem y s=true. +Proof. +rewrite remove_b; intros H;destruct (andb_prop _ _ H); auto. +Qed. + +Lemma remove_equal: + mem x s=false -> equal (remove x s) s=true. +Proof. +intros; apply equal_1; apply remove_equal. +rewrite not_mem_iff; auto. +Qed. + +Lemma add_remove: + mem x s=true -> equal (add x (remove x s)) s=true. +Proof. +intros; apply equal_1; apply add_remove; auto with set. +Qed. + +Lemma remove_add: + mem x s=false -> equal (remove x (add x s)) s=true. +Proof. +intros; apply equal_1; apply remove_add; auto. +rewrite not_mem_iff; auto. +Qed. + +(** Properties of [is_empty] *) + +Lemma is_empty_cardinal: is_empty s = zerob (cardinal s). +Proof. +intros; apply bool_1; split; intros. +rewrite MP.cardinal_1; simpl; auto with set. +assert (cardinal s = 0) by (apply zerob_true_elim; auto). +auto with set. +Qed. + +(** Properties of [singleton] *) + +Lemma singleton_mem_1: mem x (singleton x)=true. +Proof. +auto with set relations. +Qed. + +Lemma singleton_mem_2: ~E.eq x y -> mem y (singleton x)=false. +Proof. +intros; rewrite singleton_b. +unfold eqb; destruct (E.eq_dec x y); intuition. +Qed. + +Lemma singleton_mem_3: mem y (singleton x)=true -> E.eq x y. +Proof. +intros; apply singleton_1; auto with set. +Qed. + +(** Properties of [union] *) + +Lemma union_sym: + equal (union s s') (union s' s)=true. +Proof. +auto with set. +Qed. + +Lemma union_subset_equal: + subset s s'=true -> equal (union s s') s'=true. +Proof. +auto with set. +Qed. + +Lemma union_equal_1: + equal s s'=true-> equal (union s s'') (union s' s'')=true. +Proof. +auto with set. +Qed. + +Lemma union_equal_2: + equal s' s''=true-> equal (union s s') (union s s'')=true. +Proof. +auto with set. +Qed. + +Lemma union_assoc: + equal (union (union s s') s'') (union s (union s' s''))=true. +Proof. +auto with set. +Qed. + +Lemma add_union_singleton: + equal (add x s) (union (singleton x) s)=true. +Proof. +auto with set. +Qed. + +Lemma union_add: + equal (union (add x s) s') (add x (union s s'))=true. +Proof. +auto with set. +Qed. + +(* caracterisation of [union] via [subset] *) + +Lemma union_subset_1: subset s (union s s')=true. +Proof. +auto with set. +Qed. + +Lemma union_subset_2: subset s' (union s s')=true. +Proof. +auto with set. +Qed. + +Lemma union_subset_3: + subset s s''=true -> subset s' s''=true -> + subset (union s s') s''=true. +Proof. +intros; apply subset_1; apply union_subset_3; auto with set. +Qed. + +(** Properties of [inter] *) + +Lemma inter_sym: equal (inter s s') (inter s' s)=true. +Proof. +auto with set. +Qed. + +Lemma inter_subset_equal: + subset s s'=true -> equal (inter s s') s=true. +Proof. +auto with set. +Qed. + +Lemma inter_equal_1: + equal s s'=true -> equal (inter s s'') (inter s' s'')=true. +Proof. +auto with set. +Qed. + +Lemma inter_equal_2: + equal s' s''=true -> equal (inter s s') (inter s s'')=true. +Proof. +auto with set. +Qed. + +Lemma inter_assoc: + equal (inter (inter s s') s'') (inter s (inter s' s''))=true. +Proof. +auto with set. +Qed. + +Lemma union_inter_1: + equal (inter (union s s') s'') (union (inter s s'') (inter s' s''))=true. +Proof. +auto with set. +Qed. + +Lemma union_inter_2: + equal (union (inter s s') s'') (inter (union s s'') (union s' s''))=true. +Proof. +auto with set. +Qed. + +Lemma inter_add_1: mem x s'=true -> + equal (inter (add x s) s') (add x (inter s s'))=true. +Proof. +auto with set. +Qed. + +Lemma inter_add_2: mem x s'=false -> + equal (inter (add x s) s') (inter s s')=true. +Proof. +intros; apply equal_1; apply inter_add_2. +rewrite not_mem_iff; auto. +Qed. + +(* caracterisation of [union] via [subset] *) + +Lemma inter_subset_1: subset (inter s s') s=true. +Proof. +auto with set. +Qed. + +Lemma inter_subset_2: subset (inter s s') s'=true. +Proof. +auto with set. +Qed. + +Lemma inter_subset_3: + subset s'' s=true -> subset s'' s'=true -> + subset s'' (inter s s')=true. +Proof. +intros; apply subset_1; apply inter_subset_3; auto with set. +Qed. + +(** Properties of [diff] *) + +Lemma diff_subset: subset (diff s s') s=true. +Proof. +auto with set. +Qed. + +Lemma diff_subset_equal: + subset s s'=true -> equal (diff s s') empty=true. +Proof. +auto with set. +Qed. + +Lemma remove_inter_singleton: + equal (remove x s) (diff s (singleton x))=true. +Proof. +auto with set. +Qed. + +Lemma diff_inter_empty: + equal (inter (diff s s') (inter s s')) empty=true. +Proof. +auto with set. +Qed. + +Lemma diff_inter_all: + equal (union (diff s s') (inter s s')) s=true. +Proof. +auto with set. +Qed. + +End BasicProperties. + +Hint Immediate empty_mem is_empty_equal_empty add_mem_1 + remove_mem_1 singleton_equal_add union_mem inter_mem + diff_mem equal_sym add_remove remove_add : set. +Hint Resolve equal_mem_1 subset_mem_1 choose_mem_1 + choose_mem_2 add_mem_2 remove_mem_2 equal_refl equal_equal + subset_refl subset_equal subset_antisym + add_mem_3 add_equal remove_mem_3 remove_equal : set. + + +(** General recursion principle *) + +Lemma set_rec: forall (P:t->Type), + (forall s s', equal s s'=true -> P s -> P s') -> + (forall s x, mem x s=false -> P s -> P (add x s)) -> + P empty -> forall s, P s. +Proof. +intros. +apply set_induction; auto; intros. +apply X with empty; auto with set. +apply X with (add x s0); auto with set. +apply equal_1; intro a; rewrite add_iff; rewrite (H0 a); tauto. +apply X0; auto with set; apply mem_3; auto. +Qed. + +(** Properties of [fold] *) + +Lemma exclusive_set : forall s s' x, + ~(In x s/\In x s') <-> mem x s && mem x s'=false. +Proof. +intros; do 2 rewrite mem_iff. +destruct (mem x s); destruct (mem x s'); intuition. +Qed. + +Section Fold. +Variables (A:Type)(eqA:A->A->Prop)(st:Equivalence eqA). +Variables (f:elt->A->A)(Comp:Proper (E.eq==>eqA==>eqA) f)(Ass:transpose eqA f). +Variables (i:A). +Variables (s s':t)(x:elt). + +Lemma fold_empty: (fold f empty i) = i. +Proof. +apply fold_empty; auto. +Qed. + +Lemma fold_equal: + equal s s'=true -> eqA (fold f s i) (fold f s' i). +Proof. +intros; apply fold_equal with (eqA:=eqA); auto with set. +Qed. + +Lemma fold_add: + mem x s=false -> eqA (fold f (add x s) i) (f x (fold f s i)). +Proof. +intros; apply fold_add with (eqA:=eqA); auto. +rewrite not_mem_iff; auto. +Qed. + +Lemma add_fold: + mem x s=true -> eqA (fold f (add x s) i) (fold f s i). +Proof. +intros; apply add_fold with (eqA:=eqA); auto with set. +Qed. + +Lemma remove_fold_1: + mem x s=true -> eqA (f x (fold f (remove x s) i)) (fold f s i). +Proof. +intros; apply remove_fold_1 with (eqA:=eqA); auto with set. +Qed. + +Lemma remove_fold_2: + mem x s=false -> eqA (fold f (remove x s) i) (fold f s i). +Proof. +intros; apply remove_fold_2 with (eqA:=eqA); auto. +rewrite not_mem_iff; auto. +Qed. + +Lemma fold_union: + (forall x, mem x s && mem x s'=false) -> + eqA (fold f (union s s') i) (fold f s (fold f s' i)). +Proof. +intros; apply fold_union with (eqA:=eqA); auto. +intros; rewrite exclusive_set; auto. +Qed. + +End Fold. + +(** Properties of [cardinal] *) + +Lemma add_cardinal_1: + forall s x, mem x s=true -> cardinal (add x s)=cardinal s. +Proof. +auto with set. +Qed. + +Lemma add_cardinal_2: + forall s x, mem x s=false -> cardinal (add x s)=S (cardinal s). +Proof. +intros; apply add_cardinal_2; auto. +rewrite not_mem_iff; auto. +Qed. + +Lemma remove_cardinal_1: + forall s x, mem x s=true -> S (cardinal (remove x s))=cardinal s. +Proof. +intros; apply remove_cardinal_1; auto with set. +Qed. + +Lemma remove_cardinal_2: + forall s x, mem x s=false -> cardinal (remove x s)=cardinal s. +Proof. +intros; apply Equal_cardinal; apply equal_2; auto with set. +Qed. + +Lemma union_cardinal: + forall s s', (forall x, mem x s && mem x s'=false) -> + cardinal (union s s')=cardinal s+cardinal s'. +Proof. +intros; apply union_cardinal; auto; intros. +rewrite exclusive_set; auto. +Qed. + +Lemma subset_cardinal: + forall s s', subset s s'=true -> cardinal s<=cardinal s'. +Proof. +intros; apply subset_cardinal; auto with set. +Qed. + +Section Bool. + +(** Properties of [filter] *) + +Variable f:elt->bool. +Variable Comp: Proper (E.eq==>Logic.eq) f. + +Let Comp' : Proper (E.eq==>Logic.eq) (fun x =>negb (f x)). +Proof. +repeat red; intros; f_equal; auto. +Qed. + +Lemma filter_mem: forall s x, mem x (filter f s)=mem x s && f x. +Proof. +intros; apply filter_b; auto. +Qed. + +Lemma for_all_filter: + forall s, for_all f s=is_empty (filter (fun x => negb (f x)) s). +Proof. +intros; apply bool_1; split; intros. +apply is_empty_1. +unfold Empty; intros. +rewrite filter_iff; auto. +red; destruct 1. +rewrite <- (@for_all_iff s f) in H; auto. +rewrite (H a H0) in H1; discriminate. +apply for_all_1; auto; red; intros. +revert H; rewrite <- is_empty_iff. +unfold Empty; intro H; generalize (H x); clear H. +rewrite filter_iff; auto. +destruct (f x); auto. +Qed. + +Lemma exists_filter : + forall s, exists_ f s=negb (is_empty (filter f s)). +Proof. +intros; apply bool_1; split; intros. +destruct (exists_2 Comp H) as (a,(Ha1,Ha2)). +apply bool_6. +red; intros; apply (@is_empty_2 _ H0 a); auto with set. +generalize (@choose_1 (filter f s)) (@choose_2 (filter f s)). +destruct (choose (filter f s)). +intros H0 _; apply exists_1; auto. +exists e; generalize (H0 e); rewrite filter_iff; auto. +intros _ H0. +rewrite (is_empty_1 (H0 (refl_equal None))) in H; auto; discriminate. +Qed. + +Lemma partition_filter_1: + forall s, equal (fst (partition f s)) (filter f s)=true. +Proof. +auto with set. +Qed. + +Lemma partition_filter_2: + forall s, equal (snd (partition f s)) (filter (fun x => negb (f x)) s)=true. +Proof. +auto with set. +Qed. + +Lemma filter_add_1 : forall s x, f x = true -> + filter f (add x s) [=] add x (filter f s). +Proof. +red; intros; set_iff; do 2 (rewrite filter_iff; auto); set_iff. +intuition. +rewrite <- H; apply Comp; auto with relations. +Qed. + +Lemma filter_add_2 : forall s x, f x = false -> + filter f (add x s) [=] filter f s. +Proof. +red; intros; do 2 (rewrite filter_iff; auto); set_iff. +intuition. +assert (f x = f a) by (apply Comp; auto). +rewrite H in H1; rewrite H2 in H1; discriminate. +Qed. + +Lemma add_filter_1 : forall s s' x, + f x=true -> (Add x s s') -> (Add x (filter f s) (filter f s')). +Proof. +unfold Add, MP.Add; intros. +repeat rewrite filter_iff; auto. +rewrite H0; clear H0. +intuition. +setoid_replace y with x; auto with relations. +Qed. + +Lemma add_filter_2 : forall s s' x, + f x=false -> (Add x s s') -> filter f s [=] filter f s'. +Proof. +unfold Add, MP.Add, Equal; intros. +repeat rewrite filter_iff; auto. +rewrite H0; clear H0. +intuition. +setoid_replace x with a in H; auto. congruence. +Qed. + +Lemma union_filter: forall f g, + Proper (E.eq==>Logic.eq) f -> Proper (E.eq==>Logic.eq) g -> + forall s, union (filter f s) (filter g s) [=] filter (fun x=>orb (f x) (g x)) s. +Proof. +clear Comp' Comp f. +intros. +assert (Proper (E.eq==>Logic.eq) (fun x => orb (f x) (g x))). + repeat red; intros. + rewrite (H x y H1); rewrite (H0 x y H1); auto. +unfold Equal; intros; set_iff; repeat rewrite filter_iff; auto. +assert (f a || g a = true <-> f a = true \/ g a = true). + split; auto with bool. + intro H3; destruct (orb_prop _ _ H3); auto. +tauto. +Qed. + +Lemma filter_union: forall s s', filter f (union s s') [=] union (filter f s) (filter f s'). +Proof. +unfold Equal; intros; set_iff; repeat rewrite filter_iff; auto; set_iff; tauto. +Qed. + +(** Properties of [for_all] *) + +Lemma for_all_mem_1: forall s, + (forall x, (mem x s)=true->(f x)=true) -> (for_all f s)=true. +Proof. +intros. +rewrite for_all_filter; auto. +rewrite is_empty_equal_empty. +apply equal_mem_1;intros. +rewrite filter_b; auto. +rewrite empty_mem. +generalize (H a); case (mem a s);intros;auto. +rewrite H0;auto. +Qed. + +Lemma for_all_mem_2: forall s, + (for_all f s)=true -> forall x,(mem x s)=true -> (f x)=true. +Proof. +intros. +rewrite for_all_filter in H; auto. +rewrite is_empty_equal_empty in H. +generalize (equal_mem_2 _ _ H x). +rewrite filter_b; auto. +rewrite empty_mem. +rewrite H0; simpl;intros. +rewrite <- negb_false_iff; auto. +Qed. + +Lemma for_all_mem_3: + forall s x,(mem x s)=true -> (f x)=false -> (for_all f s)=false. +Proof. +intros. +apply (bool_eq_ind (for_all f s));intros;auto. +rewrite for_all_filter in H1; auto. +rewrite is_empty_equal_empty in H1. +generalize (equal_mem_2 _ _ H1 x). +rewrite filter_b; auto. +rewrite empty_mem. +rewrite H. +rewrite H0. +simpl;auto. +Qed. + +Lemma for_all_mem_4: + forall s, for_all f s=false -> {x:elt | mem x s=true /\ f x=false}. +Proof. +intros. +rewrite for_all_filter in H; auto. +destruct (choose_mem_3 _ H) as (x,(H0,H1));intros. +exists x. +rewrite filter_b in H1; auto. +elim (andb_prop _ _ H1). +split;auto. +rewrite <- negb_true_iff; auto. +Qed. + +(** Properties of [exists] *) + +Lemma for_all_exists: + forall s, exists_ f s = negb (for_all (fun x =>negb (f x)) s). +Proof. +intros. +rewrite for_all_b; auto. +rewrite exists_b; auto. +induction (elements s); simpl; auto. +destruct (f a); simpl; auto. +Qed. + +End Bool. +Section Bool'. + +Variable f:elt->bool. +Variable Comp: Proper (E.eq==>Logic.eq) f. + +Let Comp' : Proper (E.eq==>Logic.eq) (fun x => negb (f x)). +Proof. +repeat red; intros; f_equal; auto. +Qed. + +Lemma exists_mem_1: + forall s, (forall x, mem x s=true->f x=false) -> exists_ f s=false. +Proof. +intros. +rewrite for_all_exists; auto. +rewrite for_all_mem_1;auto with bool. +intros;generalize (H x H0);intros. +rewrite negb_true_iff; auto. +Qed. + +Lemma exists_mem_2: + forall s, exists_ f s=false -> forall x, mem x s=true -> f x=false. +Proof. +intros. +rewrite for_all_exists in H; auto. +rewrite negb_false_iff in H. +rewrite <- negb_true_iff. +apply for_all_mem_2 with (2:=H); auto. +Qed. + +Lemma exists_mem_3: + forall s x, mem x s=true -> f x=true -> exists_ f s=true. +Proof. +intros. +rewrite for_all_exists; auto. +rewrite negb_true_iff. +apply for_all_mem_3 with x;auto. +rewrite negb_false_iff; auto. +Qed. + +Lemma exists_mem_4: + forall s, exists_ f s=true -> {x:elt | (mem x s)=true /\ (f x)=true}. +Proof. +intros. +rewrite for_all_exists in H; auto. +rewrite negb_true_iff in H. +elim (@for_all_mem_4 (fun x =>negb (f x)) Comp' s);intros;auto. +elim p;intros. +exists x;split;auto. +rewrite <-negb_false_iff; auto. +Qed. + +End Bool'. + +Section Sum. + +(** Adding a valuation function on all elements of a set. *) + +Definition sum (f:elt -> nat)(s:t) := fold (fun x => plus (f x)) s 0. +Notation compat_opL := (Proper (E.eq==>Logic.eq==>Logic.eq)). +Notation transposeL := (transpose Logic.eq). + +Lemma sum_plus : + forall f g, + Proper (E.eq==>Logic.eq) f -> Proper (E.eq==>Logic.eq) g -> + forall s, sum (fun x =>f x+g x) s = sum f s + sum g s. +Proof. +unfold sum. +intros f g Hf Hg. +assert (fc : compat_opL (fun x:elt =>plus (f x))) by + (repeat red; intros; rewrite Hf; auto). +assert (ft : transposeL (fun x:elt =>plus (f x))) by (red; intros; omega). +assert (gc : compat_opL (fun x:elt => plus (g x))) by + (repeat red; intros; rewrite Hg; auto). +assert (gt : transposeL (fun x:elt =>plus (g x))) by (red; intros; omega). +assert (fgc : compat_opL (fun x:elt =>plus ((f x)+(g x)))) by + (repeat red; intros; rewrite Hf,Hg; auto). +assert (fgt : transposeL (fun x:elt=>plus ((f x)+(g x)))) by (red; intros; omega). +intros s;pattern s; apply set_rec. +intros. +rewrite <- (fold_equal _ _ _ _ fc ft 0 _ _ H). +rewrite <- (fold_equal _ _ _ _ gc gt 0 _ _ H). +rewrite <- (fold_equal _ _ _ _ fgc fgt 0 _ _ H); auto. +intros; do 3 (rewrite fold_add; auto with *). +do 3 rewrite fold_empty;auto. +Qed. + +Lemma sum_filter : forall f : elt -> bool, Proper (E.eq==>Logic.eq) f -> + forall s, (sum (fun x => if f x then 1 else 0) s) = (cardinal (filter f s)). +Proof. +unfold sum; intros f Hf. +assert (st : Equivalence (@Logic.eq nat)) by (split; congruence). +assert (cc : compat_opL (fun x => plus (if f x then 1 else 0))) by + (repeat red; intros; rewrite Hf; auto). +assert (ct : transposeL (fun x => plus (if f x then 1 else 0))) by + (red; intros; omega). +intros s;pattern s; apply set_rec. +intros. +change elt with E.t. +rewrite <- (fold_equal _ _ st _ cc ct 0 _ _ H). +apply equal_2 in H; rewrite <- H, <-H0; auto. +intros; rewrite (fold_add _ _ st _ cc ct); auto. +generalize (@add_filter_1 f Hf s0 (add x s0) x) (@add_filter_2 f Hf s0 (add x s0) x) . +assert (~ In x (filter f s0)). + intro H1; rewrite (mem_1 (filter_1 Hf H1)) in H; discriminate H. +case (f x); simpl; intros. +rewrite (MP.cardinal_2 H1 (H2 (refl_equal true) (MP.Add_add s0 x))); auto. +rewrite <- (MP.Equal_cardinal (H3 (refl_equal false) (MP.Add_add s0 x))); auto. +intros; rewrite fold_empty;auto. +rewrite MP.cardinal_1; auto. +unfold Empty; intros. +rewrite filter_iff; auto; set_iff; tauto. +Qed. + +Lemma fold_compat : + forall (A:Type)(eqA:A->A->Prop)(st:Equivalence eqA) + (f g:elt->A->A), + Proper (E.eq==>eqA==>eqA) f -> transpose eqA f -> + Proper (E.eq==>eqA==>eqA) g -> transpose eqA g -> + forall (i:A)(s:t),(forall x:elt, (In x s) -> forall y, (eqA (f x y) (g x y))) -> + (eqA (fold f s i) (fold g s i)). +Proof. +intros A eqA st f g fc ft gc gt i. +intro s; pattern s; apply set_rec; intros. +transitivity (fold f s0 i). +apply fold_equal with (eqA:=eqA); auto. +rewrite equal_sym; auto. +transitivity (fold g s0 i). +apply H0; intros; apply H1; auto with set. +elim (equal_2 H x); auto with set; intros. +apply fold_equal with (eqA:=eqA); auto with set. +transitivity (f x (fold f s0 i)). +apply fold_add with (eqA:=eqA); auto with set. +transitivity (g x (fold f s0 i)); auto with set relations. +transitivity (g x (fold g s0 i)); auto with set relations. +apply gc; auto with set relations. +symmetry; apply fold_add with (eqA:=eqA); auto. +do 2 rewrite fold_empty; reflexivity. +Qed. + +Lemma sum_compat : + forall f g, Proper (E.eq==>Logic.eq) f -> Proper (E.eq==>Logic.eq) g -> + forall s, (forall x, In x s -> f x=g x) -> sum f s=sum g s. +intros. +unfold sum; apply (@fold_compat _ (@Logic.eq nat)); + repeat red; auto with *. +Qed. + +End Sum. + +End WEqPropertiesOn. + +(** Now comes variants for self-contained weak sets and for full sets. + For these variants, only one argument is necessary. Thanks to + the subtyping [WS<=S], the [EqProperties] functor which is meant to be + used on modules [(M:S)] can simply be an alias of [WEqProperties]. *) + +Module WEqProperties (M:WSets) := WEqPropertiesOn M.E M. +Module EqProperties := WEqProperties. diff --git a/theories/MSets/MSetFacts.v b/theories/MSets/MSetFacts.v new file mode 100644 index 00000000..6d38b696 --- /dev/null +++ b/theories/MSets/MSetFacts.v @@ -0,0 +1,528 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) +(* \VV/ *************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(***********************************************************************) + +(* $Id$ *) + +(** * Finite sets library *) + +(** This functor derives additional facts from [MSetInterface.S]. These + facts are mainly the specifications of [MSetInterface.S] written using + different styles: equivalence and boolean equalities. + Moreover, we prove that [E.Eq] and [Equal] are setoid equalities. +*) + +Require Import DecidableTypeEx. +Require Export MSetInterface. +Set Implicit Arguments. +Unset Strict Implicit. + +(** First, a functor for Weak Sets in functorial version. *) + +Module WFactsOn (Import E : DecidableType)(Import M : WSetsOn E). + +Notation eq_dec := E.eq_dec. +Definition eqb x y := if eq_dec x y then true else false. + +(** * Specifications written using implications : + this used to be the default interface. *) + +Section ImplSpec. +Variable s s' : t. +Variable x y : elt. + +Lemma In_1 : E.eq x y -> In x s -> In y s. +Proof. intros E; rewrite E; auto. Qed. + +Lemma mem_1 : In x s -> mem x s = true. +Proof. intros; apply <- mem_spec; auto. Qed. +Lemma mem_2 : mem x s = true -> In x s. +Proof. intros; apply -> mem_spec; auto. Qed. + +Lemma equal_1 : Equal s s' -> equal s s' = true. +Proof. intros; apply <- equal_spec; auto. Qed. +Lemma equal_2 : equal s s' = true -> Equal s s'. +Proof. intros; apply -> equal_spec; auto. Qed. + +Lemma subset_1 : Subset s s' -> subset s s' = true. +Proof. intros; apply <- subset_spec; auto. Qed. +Lemma subset_2 : subset s s' = true -> Subset s s'. +Proof. intros; apply -> subset_spec; auto. Qed. + +Lemma is_empty_1 : Empty s -> is_empty s = true. +Proof. intros; apply <- is_empty_spec; auto. Qed. +Lemma is_empty_2 : is_empty s = true -> Empty s. +Proof. intros; apply -> is_empty_spec; auto. Qed. + +Lemma add_1 : E.eq x y -> In y (add x s). +Proof. intros; apply <- add_spec. auto with relations. Qed. +Lemma add_2 : In y s -> In y (add x s). +Proof. intros; apply <- add_spec; auto. Qed. +Lemma add_3 : ~ E.eq x y -> In y (add x s) -> In y s. +Proof. rewrite add_spec. intros H [H'|H']; auto. elim H; auto with relations. Qed. + +Lemma remove_1 : E.eq x y -> ~ In y (remove x s). +Proof. intros; rewrite remove_spec; intuition. Qed. +Lemma remove_2 : ~ E.eq x y -> In y s -> In y (remove x s). +Proof. intros; apply <- remove_spec; auto with relations. Qed. +Lemma remove_3 : In y (remove x s) -> In y s. +Proof. rewrite remove_spec; intuition. Qed. + +Lemma singleton_1 : In y (singleton x) -> E.eq x y. +Proof. rewrite singleton_spec; auto with relations. Qed. +Lemma singleton_2 : E.eq x y -> In y (singleton x). +Proof. rewrite singleton_spec; auto with relations. Qed. + +Lemma union_1 : In x (union s s') -> In x s \/ In x s'. +Proof. rewrite union_spec; auto. Qed. +Lemma union_2 : In x s -> In x (union s s'). +Proof. rewrite union_spec; auto. Qed. +Lemma union_3 : In x s' -> In x (union s s'). +Proof. rewrite union_spec; auto. Qed. + +Lemma inter_1 : In x (inter s s') -> In x s. +Proof. rewrite inter_spec; intuition. Qed. +Lemma inter_2 : In x (inter s s') -> In x s'. +Proof. rewrite inter_spec; intuition. Qed. +Lemma inter_3 : In x s -> In x s' -> In x (inter s s'). +Proof. rewrite inter_spec; intuition. Qed. + +Lemma diff_1 : In x (diff s s') -> In x s. +Proof. rewrite diff_spec; intuition. Qed. +Lemma diff_2 : In x (diff s s') -> ~ In x s'. +Proof. rewrite diff_spec; intuition. Qed. +Lemma diff_3 : In x s -> ~ In x s' -> In x (diff s s'). +Proof. rewrite diff_spec; auto. Qed. + +Variable f : elt -> bool. +Notation compatb := (Proper (E.eq==>Logic.eq)) (only parsing). + +Lemma filter_1 : compatb f -> In x (filter f s) -> In x s. +Proof. intros P; rewrite filter_spec; intuition. Qed. +Lemma filter_2 : compatb f -> In x (filter f s) -> f x = true. +Proof. intros P; rewrite filter_spec; intuition. Qed. +Lemma filter_3 : compatb f -> In x s -> f x = true -> In x (filter f s). +Proof. intros P; rewrite filter_spec; intuition. Qed. + +Lemma for_all_1 : compatb f -> + For_all (fun x => f x = true) s -> for_all f s = true. +Proof. intros; apply <- for_all_spec; auto. Qed. +Lemma for_all_2 : compatb f -> + for_all f s = true -> For_all (fun x => f x = true) s. +Proof. intros; apply -> for_all_spec; auto. Qed. + +Lemma exists_1 : compatb f -> + Exists (fun x => f x = true) s -> exists_ f s = true. +Proof. intros; apply <- exists_spec; auto. Qed. + +Lemma exists_2 : compatb f -> + exists_ f s = true -> Exists (fun x => f x = true) s. +Proof. intros; apply -> exists_spec; auto. Qed. + +Lemma elements_1 : In x s -> InA E.eq x (elements s). +Proof. intros; apply <- elements_spec1; auto. Qed. +Lemma elements_2 : InA E.eq x (elements s) -> In x s. +Proof. intros; apply -> elements_spec1; auto. Qed. + +End ImplSpec. + +Notation empty_1 := empty_spec (only parsing). +Notation fold_1 := fold_spec (only parsing). +Notation cardinal_1 := cardinal_spec (only parsing). +Notation partition_1 := partition_spec1 (only parsing). +Notation partition_2 := partition_spec2 (only parsing). +Notation choose_1 := choose_spec1 (only parsing). +Notation choose_2 := choose_spec2 (only parsing). +Notation elements_3w := elements_spec2w (only parsing). + +Hint Resolve mem_1 equal_1 subset_1 empty_1 + is_empty_1 choose_1 choose_2 add_1 add_2 remove_1 + remove_2 singleton_2 union_1 union_2 union_3 + inter_3 diff_3 fold_1 filter_3 for_all_1 exists_1 + partition_1 partition_2 elements_1 elements_3w + : set. +Hint Immediate In_1 mem_2 equal_2 subset_2 is_empty_2 add_3 + remove_3 singleton_1 inter_1 inter_2 diff_1 diff_2 + filter_1 filter_2 for_all_2 exists_2 elements_2 + : set. + + +(** * Specifications written using equivalences : + this is now provided by the default interface. *) + +Section IffSpec. +Variable s s' s'' : t. +Variable x y z : elt. + +Lemma In_eq_iff : E.eq x y -> (In x s <-> In y s). +Proof. +intros E; rewrite E; intuition. +Qed. + +Lemma mem_iff : In x s <-> mem x s = true. +Proof. apply iff_sym, mem_spec. Qed. + +Lemma not_mem_iff : ~In x s <-> mem x s = false. +Proof. +rewrite <-mem_spec; destruct (mem x s); intuition. +Qed. + +Lemma equal_iff : s[=]s' <-> equal s s' = true. +Proof. apply iff_sym, equal_spec. Qed. + +Lemma subset_iff : s[<=]s' <-> subset s s' = true. +Proof. apply iff_sym, subset_spec. Qed. + +Lemma empty_iff : In x empty <-> False. +Proof. intuition; apply (empty_spec H). Qed. + +Lemma is_empty_iff : Empty s <-> is_empty s = true. +Proof. apply iff_sym, is_empty_spec. Qed. + +Lemma singleton_iff : In y (singleton x) <-> E.eq x y. +Proof. rewrite singleton_spec; intuition. Qed. + +Lemma add_iff : In y (add x s) <-> E.eq x y \/ In y s. +Proof. rewrite add_spec; intuition. Qed. + +Lemma add_neq_iff : ~ E.eq x y -> (In y (add x s) <-> In y s). +Proof. rewrite add_spec; intuition. elim H; auto with relations. Qed. + +Lemma remove_iff : In y (remove x s) <-> In y s /\ ~E.eq x y. +Proof. rewrite remove_spec; intuition. Qed. + +Lemma remove_neq_iff : ~ E.eq x y -> (In y (remove x s) <-> In y s). +Proof. rewrite remove_spec; intuition. Qed. + +Variable f : elt -> bool. + +Lemma for_all_iff : Proper (E.eq==>Logic.eq) f -> + (For_all (fun x => f x = true) s <-> for_all f s = true). +Proof. intros; apply iff_sym, for_all_spec; auto. Qed. + +Lemma exists_iff : Proper (E.eq==>Logic.eq) f -> + (Exists (fun x => f x = true) s <-> exists_ f s = true). +Proof. intros; apply iff_sym, exists_spec; auto. Qed. + +Lemma elements_iff : In x s <-> InA E.eq x (elements s). +Proof. apply iff_sym, elements_spec1. Qed. + +End IffSpec. + +Notation union_iff := union_spec (only parsing). +Notation inter_iff := inter_spec (only parsing). +Notation diff_iff := diff_spec (only parsing). +Notation filter_iff := filter_spec (only parsing). + +(** Useful tactic for simplifying expressions like [In y (add x (union s s'))] *) + +Ltac set_iff := + repeat (progress ( + rewrite add_iff || rewrite remove_iff || rewrite singleton_iff + || rewrite union_iff || rewrite inter_iff || rewrite diff_iff + || rewrite empty_iff)). + +(** * Specifications written using boolean predicates *) + +Section BoolSpec. +Variable s s' s'' : t. +Variable x y z : elt. + +Lemma mem_b : E.eq x y -> mem x s = mem y s. +Proof. +intros. +generalize (mem_iff s x) (mem_iff s y)(In_eq_iff s H). +destruct (mem x s); destruct (mem y s); intuition. +Qed. + +Lemma empty_b : mem y empty = false. +Proof. +generalize (empty_iff y)(mem_iff empty y). +destruct (mem y empty); intuition. +Qed. + +Lemma add_b : mem y (add x s) = eqb x y || mem y s. +Proof. +generalize (mem_iff (add x s) y)(mem_iff s y)(add_iff s x y); unfold eqb. +destruct (eq_dec x y); destruct (mem y s); destruct (mem y (add x s)); intuition. +Qed. + +Lemma add_neq_b : ~ E.eq x y -> mem y (add x s) = mem y s. +Proof. +intros; generalize (mem_iff (add x s) y)(mem_iff s y)(add_neq_iff s H). +destruct (mem y s); destruct (mem y (add x s)); intuition. +Qed. + +Lemma remove_b : mem y (remove x s) = mem y s && negb (eqb x y). +Proof. +generalize (mem_iff (remove x s) y)(mem_iff s y)(remove_iff s x y); unfold eqb. +destruct (eq_dec x y); destruct (mem y s); destruct (mem y (remove x s)); simpl; intuition. +Qed. + +Lemma remove_neq_b : ~ E.eq x y -> mem y (remove x s) = mem y s. +Proof. +intros; generalize (mem_iff (remove x s) y)(mem_iff s y)(remove_neq_iff s H). +destruct (mem y s); destruct (mem y (remove x s)); intuition. +Qed. + +Lemma singleton_b : mem y (singleton x) = eqb x y. +Proof. +generalize (mem_iff (singleton x) y)(singleton_iff x y); unfold eqb. +destruct (eq_dec x y); destruct (mem y (singleton x)); intuition. +Qed. + +Lemma union_b : mem x (union s s') = mem x s || mem x s'. +Proof. +generalize (mem_iff (union s s') x)(mem_iff s x)(mem_iff s' x)(union_iff s s' x). +destruct (mem x s); destruct (mem x s'); destruct (mem x (union s s')); intuition. +Qed. + +Lemma inter_b : mem x (inter s s') = mem x s && mem x s'. +Proof. +generalize (mem_iff (inter s s') x)(mem_iff s x)(mem_iff s' x)(inter_iff s s' x). +destruct (mem x s); destruct (mem x s'); destruct (mem x (inter s s')); intuition. +Qed. + +Lemma diff_b : mem x (diff s s') = mem x s && negb (mem x s'). +Proof. +generalize (mem_iff (diff s s') x)(mem_iff s x)(mem_iff s' x)(diff_iff s s' x). +destruct (mem x s); destruct (mem x s'); destruct (mem x (diff s s')); simpl; intuition. +Qed. + +Lemma elements_b : mem x s = existsb (eqb x) (elements s). +Proof. +generalize (mem_iff s x)(elements_iff s x)(existsb_exists (eqb x) (elements s)). +rewrite InA_alt. +destruct (mem x s); destruct (existsb (eqb x) (elements s)); auto; intros. +symmetry. +rewrite H1. +destruct H0 as (H0,_). +destruct H0 as (a,(Ha1,Ha2)); [ intuition |]. +exists a; intuition. +unfold eqb; destruct (eq_dec x a); auto. +rewrite <- H. +rewrite H0. +destruct H1 as (H1,_). +destruct H1 as (a,(Ha1,Ha2)); [intuition|]. +exists a; intuition. +unfold eqb in *; destruct (eq_dec x a); auto; discriminate. +Qed. + +Variable f : elt->bool. + +Lemma filter_b : Proper (E.eq==>Logic.eq) f -> mem x (filter f s) = mem x s && f x. +Proof. +intros. +generalize (mem_iff (filter f s) x)(mem_iff s x)(filter_iff s x H). +destruct (mem x s); destruct (mem x (filter f s)); destruct (f x); simpl; intuition. +Qed. + +Lemma for_all_b : Proper (E.eq==>Logic.eq) f -> + for_all f s = forallb f (elements s). +Proof. +intros. +generalize (forallb_forall f (elements s))(for_all_iff s H)(elements_iff s). +unfold For_all. +destruct (forallb f (elements s)); destruct (for_all f s); auto; intros. +rewrite <- H1; intros. +destruct H0 as (H0,_). +rewrite (H2 x0) in H3. +rewrite (InA_alt E.eq x0 (elements s)) in H3. +destruct H3 as (a,(Ha1,Ha2)). +rewrite (H _ _ Ha1). +apply H0; auto. +symmetry. +rewrite H0; intros. +destruct H1 as (_,H1). +apply H1; auto. +rewrite H2. +rewrite InA_alt. exists x0; split; auto with relations. +Qed. + +Lemma exists_b : Proper (E.eq==>Logic.eq) f -> + exists_ f s = existsb f (elements s). +Proof. +intros. +generalize (existsb_exists f (elements s))(exists_iff s H)(elements_iff s). +unfold Exists. +destruct (existsb f (elements s)); destruct (exists_ f s); auto; intros. +rewrite <- H1; intros. +destruct H0 as (H0,_). +destruct H0 as (a,(Ha1,Ha2)); auto. +exists a; split; auto. +rewrite H2; rewrite InA_alt; exists a; auto with relations. +symmetry. +rewrite H0. +destruct H1 as (_,H1). +destruct H1 as (a,(Ha1,Ha2)); auto. +rewrite (H2 a) in Ha1. +rewrite (InA_alt E.eq a (elements s)) in Ha1. +destruct Ha1 as (b,(Hb1,Hb2)). +exists b; auto. +rewrite <- (H _ _ Hb1); auto. +Qed. + +End BoolSpec. + +(** * Declarations of morphisms with respects to [E.eq] and [Equal] *) + +Instance In_m : Proper (E.eq==>Equal==>iff) In. +Proof. +unfold Equal; intros x y H s s' H0. +rewrite (In_eq_iff s H); auto. +Qed. + +Instance Empty_m : Proper (Equal==>iff) Empty. +Proof. +repeat red; unfold Empty; intros s s' E. +setoid_rewrite E; auto. +Qed. + +Instance is_empty_m : Proper (Equal==>Logic.eq) is_empty. +Proof. +intros s s' H. +generalize (is_empty_iff s). rewrite H at 1. rewrite is_empty_iff. +destruct (is_empty s); destruct (is_empty s'); intuition. +Qed. + +Instance mem_m : Proper (E.eq==>Equal==>Logic.eq) mem. +Proof. +intros x x' Hx s s' Hs. +generalize (mem_iff s x). rewrite Hs, Hx at 1; rewrite mem_iff. +destruct (mem x s), (mem x' s'); intuition. +Qed. + +Instance singleton_m : Proper (E.eq==>Equal) singleton. +Proof. +intros x y H a. rewrite !singleton_iff, H; intuition. +Qed. + +Instance add_m : Proper (E.eq==>Equal==>Equal) add. +Proof. +intros x x' Hx s s' Hs a. rewrite !add_iff, Hx, Hs; intuition. +Qed. + +Instance remove_m : Proper (E.eq==>Equal==>Equal) remove. +Proof. +intros x x' Hx s s' Hs a. rewrite !remove_iff, Hx, Hs; intuition. +Qed. + +Instance union_m : Proper (Equal==>Equal==>Equal) union. +Proof. +intros s1 s1' Hs1 s2 s2' Hs2 a. rewrite !union_iff, Hs1, Hs2; intuition. +Qed. + +Instance inter_m : Proper (Equal==>Equal==>Equal) inter. +Proof. +intros s1 s1' Hs1 s2 s2' Hs2 a. rewrite !inter_iff, Hs1, Hs2; intuition. +Qed. + +Instance diff_m : Proper (Equal==>Equal==>Equal) diff. +Proof. +intros s1 s1' Hs1 s2 s2' Hs2 a. rewrite !diff_iff, Hs1, Hs2; intuition. +Qed. + +Instance Subset_m : Proper (Equal==>Equal==>iff) Subset. +Proof. +unfold Equal, Subset; firstorder. +Qed. + +Instance subset_m : Proper (Equal==>Equal==>Logic.eq) subset. +Proof. +intros s1 s1' Hs1 s2 s2' Hs2. +generalize (subset_iff s1 s2). rewrite Hs1, Hs2 at 1. rewrite subset_iff. +destruct (subset s1 s2); destruct (subset s1' s2'); intuition. +Qed. + +Instance equal_m : Proper (Equal==>Equal==>Logic.eq) equal. +Proof. +intros s1 s1' Hs1 s2 s2' Hs2. +generalize (equal_iff s1 s2). rewrite Hs1,Hs2 at 1. rewrite equal_iff. +destruct (equal s1 s2); destruct (equal s1' s2'); intuition. +Qed. + +Instance SubsetSetoid : PreOrder Subset. (* reflexive + transitive *) +Proof. firstorder. Qed. + +Definition Subset_refl := @PreOrder_Reflexive _ _ SubsetSetoid. +Definition Subset_trans := @PreOrder_Transitive _ _ SubsetSetoid. + +Instance In_s_m : Morphisms.Proper (E.eq ==> Subset ++> impl) In | 1. +Proof. + simpl_relation. eauto with set. +Qed. + +Instance Empty_s_m : Proper (Subset-->impl) Empty. +Proof. firstorder. Qed. + +Instance add_s_m : Proper (E.eq==>Subset++>Subset) add. +Proof. +intros x x' Hx s s' Hs a. rewrite !add_iff, Hx; intuition. +Qed. + +Instance remove_s_m : Proper (E.eq==>Subset++>Subset) remove. +Proof. +intros x x' Hx s s' Hs a. rewrite !remove_iff, Hx; intuition. +Qed. + +Instance union_s_m : Proper (Subset++>Subset++>Subset) union. +Proof. +intros s1 s1' Hs1 s2 s2' Hs2 a. rewrite !union_iff, Hs1, Hs2; intuition. +Qed. + +Instance inter_s_m : Proper (Subset++>Subset++>Subset) inter. +Proof. +intros s1 s1' Hs1 s2 s2' Hs2 a. rewrite !inter_iff, Hs1, Hs2; intuition. +Qed. + +Instance diff_s_m : Proper (Subset++>Subset-->Subset) diff. +Proof. +intros s1 s1' Hs1 s2 s2' Hs2 a. rewrite !diff_iff, Hs1, Hs2; intuition. +Qed. + + +(* [fold], [filter], [for_all], [exists_] and [partition] requires + some knowledge on [f] in order to be known as morphisms. *) + +Generalizable Variables f. + +Instance filter_equal : forall `(Proper _ (E.eq==>Logic.eq) f), + Proper (Equal==>Equal) (filter f). +Proof. +intros f Hf s s' Hs a. rewrite !filter_iff, Hs by auto; intuition. +Qed. + +Instance filter_subset : forall `(Proper _ (E.eq==>Logic.eq) f), + Proper (Subset==>Subset) (filter f). +Proof. +intros f Hf s s' Hs a. rewrite !filter_iff, Hs by auto; intuition. +Qed. + +Lemma filter_ext : forall f f', Proper (E.eq==>Logic.eq) f -> (forall x, f x = f' x) -> + forall s s', s[=]s' -> filter f s [=] filter f' s'. +Proof. +intros f f' Hf Hff' s s' Hss' x. rewrite 2 filter_iff; auto. +rewrite Hff', Hss'; intuition. +red; red; intros; rewrite <- 2 Hff'; auto. +Qed. + +(* For [elements], [min_elt], [max_elt] and [choose], we would need setoid + structures on [list elt] and [option elt]. *) + +(* Later: +Add Morphism cardinal ; cardinal_m. +*) + +End WFactsOn. + +(** Now comes variants for self-contained weak sets and for full sets. + For these variants, only one argument is necessary. Thanks to + the subtyping [WS<=S], the [Facts] functor which is meant to be + used on modules [(M:S)] can simply be an alias of [WFacts]. *) + +Module WFacts (M:WSets) := WFactsOn M.E M. +Module Facts := WFacts. diff --git a/theories/MSets/MSetInterface.v b/theories/MSets/MSetInterface.v new file mode 100644 index 00000000..194cb904 --- /dev/null +++ b/theories/MSets/MSetInterface.v @@ -0,0 +1,732 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) +(* \VV/ *************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(***********************************************************************) + +(* $Id$ *) + +(** * Finite set library *) + +(** Set interfaces, inspired by the one of Ocaml. When compared with + Ocaml, the main differences are: + - the lack of [iter] function, useless since Coq is purely functional + - the use of [option] types instead of [Not_found] exceptions + - the use of [nat] instead of [int] for the [cardinal] function + + Several variants of the set interfaces are available: + - [WSetsOn] : functorial signature for weak sets + - [WSets] : self-contained version of [WSets] + - [SetsOn] : functorial signature for ordered sets + - [Sets] : self-contained version of [Sets] + - [WRawSets] : a signature for weak sets that may be ill-formed + - [RawSets] : same for ordered sets + + If unsure, [S = Sets] is probably what you're looking for: most other + signatures are subsets of it, while [Sets] can be obtained from + [RawSets] via the use of a subset type (see (W)Raw2Sets below). +*) + +Require Export Bool SetoidList RelationClasses Morphisms + RelationPairs Equalities Orders OrdersFacts. +Set Implicit Arguments. +Unset Strict Implicit. + +Module Type TypElt. + Parameters t elt : Type. +End TypElt. + +Module Type HasWOps (Import T:TypElt). + + Parameter empty : t. + (** The empty set. *) + + Parameter is_empty : t -> bool. + (** Test whether a set is empty or not. *) + + Parameter mem : elt -> t -> bool. + (** [mem x s] tests whether [x] belongs to the set [s]. *) + + Parameter add : elt -> t -> t. + (** [add x s] returns a set containing all elements of [s], + plus [x]. If [x] was already in [s], [s] is returned unchanged. *) + + Parameter singleton : elt -> t. + (** [singleton x] returns the one-element set containing only [x]. *) + + Parameter remove : elt -> t -> t. + (** [remove x s] returns a set containing all elements of [s], + except [x]. If [x] was not in [s], [s] is returned unchanged. *) + + Parameter union : t -> t -> t. + (** Set union. *) + + Parameter inter : t -> t -> t. + (** Set intersection. *) + + Parameter diff : t -> t -> t. + (** Set difference. *) + + Parameter equal : t -> t -> bool. + (** [equal s1 s2] tests whether the sets [s1] and [s2] are + equal, that is, contain equal elements. *) + + Parameter subset : t -> t -> bool. + (** [subset s1 s2] tests whether the set [s1] is a subset of + the set [s2]. *) + + Parameter fold : forall A : Type, (elt -> A -> A) -> t -> A -> A. + (** [fold f s a] computes [(f xN ... (f x2 (f x1 a))...)], + where [x1 ... xN] are the elements of [s]. + The order in which elements of [s] are presented to [f] is + unspecified. *) + + Parameter for_all : (elt -> bool) -> t -> bool. + (** [for_all p s] checks if all elements of the set + satisfy the predicate [p]. *) + + Parameter exists_ : (elt -> bool) -> t -> bool. + (** [exists p s] checks if at least one element of + the set satisfies the predicate [p]. *) + + Parameter filter : (elt -> bool) -> t -> t. + (** [filter p s] returns the set of all elements in [s] + that satisfy predicate [p]. *) + + Parameter partition : (elt -> bool) -> t -> t * t. + (** [partition p s] returns a pair of sets [(s1, s2)], where + [s1] is the set of all the elements of [s] that satisfy the + predicate [p], and [s2] is the set of all the elements of + [s] that do not satisfy [p]. *) + + Parameter cardinal : t -> nat. + (** Return the number of elements of a set. *) + + Parameter elements : t -> list elt. + (** Return the list of all elements of the given set, in any order. *) + + Parameter choose : t -> option elt. + (** Return one element of the given set, or [None] if + the set is empty. Which element is chosen is unspecified. + Equal sets could return different elements. *) + +End HasWOps. + +Module Type WOps (E : DecidableType). + Definition elt := E.t. + Parameter t : Type. (** the abstract type of sets *) + Include HasWOps. +End WOps. + + +(** ** Functorial signature for weak sets + + Weak sets are sets without ordering on base elements, only + a decidable equality. *) + +Module Type WSetsOn (E : DecidableType). + (** First, we ask for all the functions *) + Include WOps E. + + (** Logical predicates *) + Parameter In : elt -> t -> Prop. + Declare Instance In_compat : Proper (E.eq==>eq==>iff) In. + + Definition Equal s s' := forall a : elt, In a s <-> In a s'. + Definition Subset s s' := forall a : elt, In a s -> In a s'. + Definition Empty s := forall a : elt, ~ In a s. + Definition For_all (P : elt -> Prop) s := forall x, In x s -> P x. + Definition Exists (P : elt -> Prop) s := exists x, In x s /\ P x. + + Notation "s [=] t" := (Equal s t) (at level 70, no associativity). + Notation "s [<=] t" := (Subset s t) (at level 70, no associativity). + + Definition eq : t -> t -> Prop := Equal. + Include IsEq. (** [eq] is obviously an equivalence, for subtyping only *) + Include HasEqDec. + + (** Specifications of set operators *) + + Section Spec. + Variable s s': t. + Variable x y : elt. + Variable f : elt -> bool. + Notation compatb := (Proper (E.eq==>Logic.eq)) (only parsing). + + Parameter mem_spec : mem x s = true <-> In x s. + Parameter equal_spec : equal s s' = true <-> s[=]s'. + Parameter subset_spec : subset s s' = true <-> s[<=]s'. + Parameter empty_spec : Empty empty. + Parameter is_empty_spec : is_empty s = true <-> Empty s. + Parameter add_spec : In y (add x s) <-> E.eq y x \/ In y s. + Parameter remove_spec : In y (remove x s) <-> In y s /\ ~E.eq y x. + Parameter singleton_spec : In y (singleton x) <-> E.eq y x. + Parameter union_spec : In x (union s s') <-> In x s \/ In x s'. + Parameter inter_spec : In x (inter s s') <-> In x s /\ In x s'. + Parameter diff_spec : In x (diff s s') <-> In x s /\ ~In x s'. + Parameter fold_spec : forall (A : Type) (i : A) (f : elt -> A -> A), + fold f s i = fold_left (flip f) (elements s) i. + Parameter cardinal_spec : cardinal s = length (elements s). + Parameter filter_spec : compatb f -> + (In x (filter f s) <-> In x s /\ f x = true). + Parameter for_all_spec : compatb f -> + (for_all f s = true <-> For_all (fun x => f x = true) s). + Parameter exists_spec : compatb f -> + (exists_ f s = true <-> Exists (fun x => f x = true) s). + Parameter partition_spec1 : compatb f -> + fst (partition f s) [=] filter f s. + Parameter partition_spec2 : compatb f -> + snd (partition f s) [=] filter (fun x => negb (f x)) s. + Parameter elements_spec1 : InA E.eq x (elements s) <-> In x s. + (** When compared with ordered sets, here comes the only + property that is really weaker: *) + Parameter elements_spec2w : NoDupA E.eq (elements s). + Parameter choose_spec1 : choose s = Some x -> In x s. + Parameter choose_spec2 : choose s = None -> Empty s. + + End Spec. + +End WSetsOn. + +(** ** Static signature for weak sets + + Similar to the functorial signature [WSetsOn], except that the + module [E] of base elements is incorporated in the signature. *) + +Module Type WSets. + Declare Module E : DecidableType. + Include WSetsOn E. +End WSets. + +(** ** Functorial signature for sets on ordered elements + + Based on [WSetsOn], plus ordering on sets and [min_elt] and [max_elt] + and some stronger specifications for other functions. *) + +Module Type HasOrdOps (Import T:TypElt). + + Parameter compare : t -> t -> comparison. + (** Total ordering between sets. Can be used as the ordering function + for doing sets of sets. *) + + Parameter min_elt : t -> option elt. + (** Return the smallest element of the given set + (with respect to the [E.compare] ordering), + or [None] if the set is empty. *) + + Parameter max_elt : t -> option elt. + (** Same as [min_elt], but returns the largest element of the + given set. *) + +End HasOrdOps. + +Module Type Ops (E : OrderedType) := WOps E <+ HasOrdOps. + + +Module Type SetsOn (E : OrderedType). + Include WSetsOn E <+ HasOrdOps <+ HasLt <+ IsStrOrder. + + Section Spec. + Variable s s': t. + Variable x y : elt. + + Parameter compare_spec : CompSpec eq lt s s' (compare s s'). + + (** Additional specification of [elements] *) + Parameter elements_spec2 : sort E.lt (elements s). + + (** Remark: since [fold] is specified via [elements], this stronger + specification of [elements] has an indirect impact on [fold], + which can now be proved to receive elements in increasing order. + *) + + Parameter min_elt_spec1 : min_elt s = Some x -> In x s. + Parameter min_elt_spec2 : min_elt s = Some x -> In y s -> ~ E.lt y x. + Parameter min_elt_spec3 : min_elt s = None -> Empty s. + + Parameter max_elt_spec1 : max_elt s = Some x -> In x s. + Parameter max_elt_spec2 : max_elt s = Some x -> In y s -> ~ E.lt x y. + Parameter max_elt_spec3 : max_elt s = None -> Empty s. + + (** Additional specification of [choose] *) + Parameter choose_spec3 : choose s = Some x -> choose s' = Some y -> + Equal s s' -> E.eq x y. + + End Spec. + +End SetsOn. + + +(** ** Static signature for sets on ordered elements + + Similar to the functorial signature [SetsOn], except that the + module [E] of base elements is incorporated in the signature. *) + +Module Type Sets. + Declare Module E : OrderedType. + Include SetsOn E. +End Sets. + +Module Type S := Sets. + + +(** ** Some subtyping tests +<< +WSetsOn ---> WSets + | | + | | + V V +SetsOn ---> Sets + +Module S_WS (M : Sets) <: WSets := M. +Module Sfun_WSfun (E:OrderedType)(M : SetsOn E) <: WSetsOn E := M. +Module S_Sfun (M : Sets) <: SetsOn M.E := M. +Module WS_WSfun (M : WSets) <: WSetsOn M.E := M. +>> +*) + + + +(** ** Signatures for set representations with ill-formed values. + + Motivation: + + For many implementation of finite sets (AVL trees, sorted + lists, lists without duplicates), we use the same two-layer + approach: + + - A first module deals with the datatype (eg. list or tree) without + any restriction on the values we consider. In this module (named + "Raw" in the past), some results are stated under the assumption + that some invariant (e.g. sortedness) holds for the input sets. We + also prove that this invariant is preserved by set operators. + + - A second module implements the exact Sets interface by + using a subtype, for instance [{ l : list A | sorted l }]. + This module is a mere wrapper around the first Raw module. + + With the interfaces below, we give some respectability to + the "Raw" modules. This allows the interested users to directly + access them via the interfaces. Even better, we can build once + and for all a functor doing the transition between Raw and usual Sets. + + Description: + + The type [t] of sets may contain ill-formed values on which our + set operators may give wrong answers. In particular, [mem] + may not see a element in a ill-formed set (think for instance of a + unsorted list being given to an optimized [mem] that stops + its search as soon as a strictly larger element is encountered). + + Unlike optimized operators, the [In] predicate is supposed to + always be correct, even on ill-formed sets. Same for [Equal] and + other logical predicates. + + A predicate parameter [Ok] is used to discriminate between + well-formed and ill-formed values. Some lemmas hold only on sets + validating [Ok]. This predicate [Ok] is required to be + preserved by set operators. Moreover, a boolean function [isok] + should exist for identifying (at least some of) the well-formed sets. + +*) + + +Module Type WRawSets (E : DecidableType). + (** First, we ask for all the functions *) + Include WOps E. + + (** Is a set well-formed or ill-formed ? *) + + Parameter IsOk : t -> Prop. + Class Ok (s:t) : Prop := ok : IsOk s. + + (** In order to be able to validate (at least some) particular sets as + well-formed, we ask for a boolean function for (semi-)deciding + predicate [Ok]. If [Ok] isn't decidable, [isok] may be the + always-false function. *) + Parameter isok : t -> bool. + Declare Instance isok_Ok s `(isok s = true) : Ok s | 10. + + (** Logical predicates *) + Parameter In : elt -> t -> Prop. + Declare Instance In_compat : Proper (E.eq==>eq==>iff) In. + + Definition Equal s s' := forall a : elt, In a s <-> In a s'. + Definition Subset s s' := forall a : elt, In a s -> In a s'. + Definition Empty s := forall a : elt, ~ In a s. + Definition For_all (P : elt -> Prop) s := forall x, In x s -> P x. + Definition Exists (P : elt -> Prop) s := exists x, In x s /\ P x. + + Notation "s [=] t" := (Equal s t) (at level 70, no associativity). + Notation "s [<=] t" := (Subset s t) (at level 70, no associativity). + + Definition eq : t -> t -> Prop := Equal. + Declare Instance eq_equiv : Equivalence eq. + + (** First, all operations are compatible with the well-formed predicate. *) + + Declare Instance empty_ok : Ok empty. + Declare Instance add_ok s x `(Ok s) : Ok (add x s). + Declare Instance remove_ok s x `(Ok s) : Ok (remove x s). + Declare Instance singleton_ok x : Ok (singleton x). + Declare Instance union_ok s s' `(Ok s, Ok s') : Ok (union s s'). + Declare Instance inter_ok s s' `(Ok s, Ok s') : Ok (inter s s'). + Declare Instance diff_ok s s' `(Ok s, Ok s') : Ok (diff s s'). + Declare Instance filter_ok s f `(Ok s) : Ok (filter f s). + Declare Instance partition_ok1 s f `(Ok s) : Ok (fst (partition f s)). + Declare Instance partition_ok2 s f `(Ok s) : Ok (snd (partition f s)). + + (** Now, the specifications, with constraints on the input sets. *) + + Section Spec. + Variable s s': t. + Variable x y : elt. + Variable f : elt -> bool. + Notation compatb := (Proper (E.eq==>Logic.eq)) (only parsing). + + Parameter mem_spec : forall `{Ok s}, mem x s = true <-> In x s. + Parameter equal_spec : forall `{Ok s, Ok s'}, + equal s s' = true <-> s[=]s'. + Parameter subset_spec : forall `{Ok s, Ok s'}, + subset s s' = true <-> s[<=]s'. + Parameter empty_spec : Empty empty. + Parameter is_empty_spec : is_empty s = true <-> Empty s. + Parameter add_spec : forall `{Ok s}, + In y (add x s) <-> E.eq y x \/ In y s. + Parameter remove_spec : forall `{Ok s}, + In y (remove x s) <-> In y s /\ ~E.eq y x. + Parameter singleton_spec : In y (singleton x) <-> E.eq y x. + Parameter union_spec : forall `{Ok s, Ok s'}, + In x (union s s') <-> In x s \/ In x s'. + Parameter inter_spec : forall `{Ok s, Ok s'}, + In x (inter s s') <-> In x s /\ In x s'. + Parameter diff_spec : forall `{Ok s, Ok s'}, + In x (diff s s') <-> In x s /\ ~In x s'. + Parameter fold_spec : forall (A : Type) (i : A) (f : elt -> A -> A), + fold f s i = fold_left (flip f) (elements s) i. + Parameter cardinal_spec : forall `{Ok s}, + cardinal s = length (elements s). + Parameter filter_spec : compatb f -> + (In x (filter f s) <-> In x s /\ f x = true). + Parameter for_all_spec : compatb f -> + (for_all f s = true <-> For_all (fun x => f x = true) s). + Parameter exists_spec : compatb f -> + (exists_ f s = true <-> Exists (fun x => f x = true) s). + Parameter partition_spec1 : compatb f -> + fst (partition f s) [=] filter f s. + Parameter partition_spec2 : compatb f -> + snd (partition f s) [=] filter (fun x => negb (f x)) s. + Parameter elements_spec1 : InA E.eq x (elements s) <-> In x s. + Parameter elements_spec2w : forall `{Ok s}, NoDupA E.eq (elements s). + Parameter choose_spec1 : choose s = Some x -> In x s. + Parameter choose_spec2 : choose s = None -> Empty s. + + End Spec. + +End WRawSets. + +(** From weak raw sets to weak usual sets *) + +Module WRaw2SetsOn (E:DecidableType)(M:WRawSets E) <: WSetsOn E. + + (** We avoid creating induction principles for the Record *) + Local Unset Elimination Schemes. + Local Unset Case Analysis Schemes. + + Definition elt := E.t. + + Record t_ := Mkt {this :> M.t; is_ok : M.Ok this}. + Definition t := t_. + Implicit Arguments Mkt [ [is_ok] ]. + Hint Resolve is_ok : typeclass_instances. + + Definition In (x : elt)(s : t) := M.In x s.(this). + Definition Equal (s s' : t) := forall a : elt, In a s <-> In a s'. + Definition Subset (s s' : t) := forall a : elt, In a s -> In a s'. + Definition Empty (s : t) := forall a : elt, ~ In a s. + Definition For_all (P : elt -> Prop)(s : t) := forall x, In x s -> P x. + Definition Exists (P : elt -> Prop)(s : t) := exists x, In x s /\ P x. + + Definition mem (x : elt)(s : t) := M.mem x s. + Definition add (x : elt)(s : t) : t := Mkt (M.add x s). + Definition remove (x : elt)(s : t) : t := Mkt (M.remove x s). + Definition singleton (x : elt) : t := Mkt (M.singleton x). + Definition union (s s' : t) : t := Mkt (M.union s s'). + Definition inter (s s' : t) : t := Mkt (M.inter s s'). + Definition diff (s s' : t) : t := Mkt (M.diff s s'). + Definition equal (s s' : t) := M.equal s s'. + Definition subset (s s' : t) := M.subset s s'. + Definition empty : t := Mkt M.empty. + Definition is_empty (s : t) := M.is_empty s. + Definition elements (s : t) : list elt := M.elements s. + Definition choose (s : t) : option elt := M.choose s. + Definition fold (A : Type)(f : elt -> A -> A)(s : t) : A -> A := M.fold f s. + Definition cardinal (s : t) := M.cardinal s. + Definition filter (f : elt -> bool)(s : t) : t := Mkt (M.filter f s). + Definition for_all (f : elt -> bool)(s : t) := M.for_all f s. + Definition exists_ (f : elt -> bool)(s : t) := M.exists_ f s. + Definition partition (f : elt -> bool)(s : t) : t * t := + let p := M.partition f s in (Mkt (fst p), Mkt (snd p)). + + Instance In_compat : Proper (E.eq==>eq==>iff) In. + Proof. repeat red. intros; apply M.In_compat; congruence. Qed. + + Definition eq : t -> t -> Prop := Equal. + + Instance eq_equiv : Equivalence eq. + Proof. firstorder. Qed. + + Definition eq_dec : forall (s s':t), { eq s s' }+{ ~eq s s' }. + Proof. + intros (s,Hs) (s',Hs'). + change ({M.Equal s s'}+{~M.Equal s s'}). + destruct (M.equal s s') as [ ]_eqn:H; [left|right]; + rewrite <- M.equal_spec; congruence. + Defined. + + + Section Spec. + Variable s s' : t. + Variable x y : elt. + Variable f : elt -> bool. + Notation compatb := (Proper (E.eq==>Logic.eq)) (only parsing). + + Lemma mem_spec : mem x s = true <-> In x s. + Proof. exact (@M.mem_spec _ _ _). Qed. + Lemma equal_spec : equal s s' = true <-> Equal s s'. + Proof. exact (@M.equal_spec _ _ _ _). Qed. + Lemma subset_spec : subset s s' = true <-> Subset s s'. + Proof. exact (@M.subset_spec _ _ _ _). Qed. + Lemma empty_spec : Empty empty. + Proof. exact M.empty_spec. Qed. + Lemma is_empty_spec : is_empty s = true <-> Empty s. + Proof. exact (@M.is_empty_spec _). Qed. + Lemma add_spec : In y (add x s) <-> E.eq y x \/ In y s. + Proof. exact (@M.add_spec _ _ _ _). Qed. + Lemma remove_spec : In y (remove x s) <-> In y s /\ ~E.eq y x. + Proof. exact (@M.remove_spec _ _ _ _). Qed. + Lemma singleton_spec : In y (singleton x) <-> E.eq y x. + Proof. exact (@M.singleton_spec _ _). Qed. + Lemma union_spec : In x (union s s') <-> In x s \/ In x s'. + Proof. exact (@M.union_spec _ _ _ _ _). Qed. + Lemma inter_spec : In x (inter s s') <-> In x s /\ In x s'. + Proof. exact (@M.inter_spec _ _ _ _ _). Qed. + Lemma diff_spec : In x (diff s s') <-> In x s /\ ~In x s'. + Proof. exact (@M.diff_spec _ _ _ _ _). Qed. + Lemma fold_spec : forall (A : Type) (i : A) (f : elt -> A -> A), + fold f s i = fold_left (fun a e => f e a) (elements s) i. + Proof. exact (@M.fold_spec _). Qed. + Lemma cardinal_spec : cardinal s = length (elements s). + Proof. exact (@M.cardinal_spec s _). Qed. + Lemma filter_spec : compatb f -> + (In x (filter f s) <-> In x s /\ f x = true). + Proof. exact (@M.filter_spec _ _ _). Qed. + Lemma for_all_spec : compatb f -> + (for_all f s = true <-> For_all (fun x => f x = true) s). + Proof. exact (@M.for_all_spec _ _). Qed. + Lemma exists_spec : compatb f -> + (exists_ f s = true <-> Exists (fun x => f x = true) s). + Proof. exact (@M.exists_spec _ _). Qed. + Lemma partition_spec1 : compatb f -> Equal (fst (partition f s)) (filter f s). + Proof. exact (@M.partition_spec1 _ _). Qed. + Lemma partition_spec2 : compatb f -> + Equal (snd (partition f s)) (filter (fun x => negb (f x)) s). + Proof. exact (@M.partition_spec2 _ _). Qed. + Lemma elements_spec1 : InA E.eq x (elements s) <-> In x s. + Proof. exact (@M.elements_spec1 _ _). Qed. + Lemma elements_spec2w : NoDupA E.eq (elements s). + Proof. exact (@M.elements_spec2w _ _). Qed. + Lemma choose_spec1 : choose s = Some x -> In x s. + Proof. exact (@M.choose_spec1 _ _). Qed. + Lemma choose_spec2 : choose s = None -> Empty s. + Proof. exact (@M.choose_spec2 _). Qed. + + End Spec. + +End WRaw2SetsOn. + +Module WRaw2Sets (D:DecidableType)(M:WRawSets D) <: WSets with Module E := D. + Module E := D. + Include WRaw2SetsOn D M. +End WRaw2Sets. + +(** Same approach for ordered sets *) + +Module Type RawSets (E : OrderedType). + Include WRawSets E <+ HasOrdOps <+ HasLt <+ IsStrOrder. + + Section Spec. + Variable s s': t. + Variable x y : elt. + + (** Specification of [compare] *) + Parameter compare_spec : forall `{Ok s, Ok s'}, CompSpec eq lt s s' (compare s s'). + + (** Additional specification of [elements] *) + Parameter elements_spec2 : forall `{Ok s}, sort E.lt (elements s). + + (** Specification of [min_elt] *) + Parameter min_elt_spec1 : min_elt s = Some x -> In x s. + Parameter min_elt_spec2 : forall `{Ok s}, min_elt s = Some x -> In y s -> ~ E.lt y x. + Parameter min_elt_spec3 : min_elt s = None -> Empty s. + + (** Specification of [max_elt] *) + Parameter max_elt_spec1 : max_elt s = Some x -> In x s. + Parameter max_elt_spec2 : forall `{Ok s}, max_elt s = Some x -> In y s -> ~ E.lt x y. + Parameter max_elt_spec3 : max_elt s = None -> Empty s. + + (** Additional specification of [choose] *) + Parameter choose_spec3 : forall `{Ok s, Ok s'}, + choose s = Some x -> choose s' = Some y -> Equal s s' -> E.eq x y. + + End Spec. + +End RawSets. + +(** From Raw to usual sets *) + +Module Raw2SetsOn (O:OrderedType)(M:RawSets O) <: SetsOn O. + Include WRaw2SetsOn O M. + + Definition compare (s s':t) := M.compare s s'. + Definition min_elt (s:t) : option elt := M.min_elt s. + Definition max_elt (s:t) : option elt := M.max_elt s. + Definition lt (s s':t) := M.lt s s'. + + (** Specification of [lt] *) + Instance lt_strorder : StrictOrder lt. + Proof. constructor ; unfold lt; red. + unfold complement. red. intros. apply (irreflexivity H). + intros. transitivity y; auto. + Qed. + + Instance lt_compat : Proper (eq==>eq==>iff) lt. + Proof. + repeat red. unfold eq, lt. + intros (s1,p1) (s2,p2) E (s1',p1') (s2',p2') E'; simpl. + change (M.eq s1 s2) in E. + change (M.eq s1' s2') in E'. + rewrite E,E'; intuition. + Qed. + + Section Spec. + Variable s s' s'' : t. + Variable x y : elt. + + Lemma compare_spec : CompSpec eq lt s s' (compare s s'). + Proof. unfold compare; destruct (@M.compare_spec s s' _ _); auto. Qed. + + (** Additional specification of [elements] *) + Lemma elements_spec2 : sort O.lt (elements s). + Proof. exact (@M.elements_spec2 _ _). Qed. + + (** Specification of [min_elt] *) + Lemma min_elt_spec1 : min_elt s = Some x -> In x s. + Proof. exact (@M.min_elt_spec1 _ _). Qed. + Lemma min_elt_spec2 : min_elt s = Some x -> In y s -> ~ O.lt y x. + Proof. exact (@M.min_elt_spec2 _ _ _ _). Qed. + Lemma min_elt_spec3 : min_elt s = None -> Empty s. + Proof. exact (@M.min_elt_spec3 _). Qed. + + (** Specification of [max_elt] *) + Lemma max_elt_spec1 : max_elt s = Some x -> In x s. + Proof. exact (@M.max_elt_spec1 _ _). Qed. + Lemma max_elt_spec2 : max_elt s = Some x -> In y s -> ~ O.lt x y. + Proof. exact (@M.max_elt_spec2 _ _ _ _). Qed. + Lemma max_elt_spec3 : max_elt s = None -> Empty s. + Proof. exact (@M.max_elt_spec3 _). Qed. + + (** Additional specification of [choose] *) + Lemma choose_spec3 : + choose s = Some x -> choose s' = Some y -> Equal s s' -> O.eq x y. + Proof. exact (@M.choose_spec3 _ _ _ _ _ _). Qed. + + End Spec. + +End Raw2SetsOn. + +Module Raw2Sets (O:OrderedType)(M:RawSets O) <: Sets with Module E := O. + Module E := O. + Include Raw2SetsOn O M. +End Raw2Sets. + + +(** We provide an ordering for sets-as-sorted-lists *) + +Module MakeListOrdering (O:OrderedType). + Module MO:=OrderedTypeFacts O. + + Local Notation t := (list O.t). + Local Notation In := (InA O.eq). + + Definition eq s s' := forall x, In x s <-> In x s'. + + Instance eq_equiv : Equivalence eq. + + Inductive lt_list : t -> t -> Prop := + | lt_nil : forall x s, lt_list nil (x :: s) + | lt_cons_lt : forall x y s s', + O.lt x y -> lt_list (x :: s) (y :: s') + | lt_cons_eq : forall x y s s', + O.eq x y -> lt_list s s' -> lt_list (x :: s) (y :: s'). + Hint Constructors lt_list. + + Definition lt := lt_list. + Hint Unfold lt. + + Instance lt_strorder : StrictOrder lt. + Proof. + split. + (* irreflexive *) + assert (forall s s', s=s' -> ~lt s s'). + red; induction 2. + discriminate. + inversion H; subst. + apply (StrictOrder_Irreflexive y); auto. + inversion H; subst; auto. + intros s Hs; exact (H s s (eq_refl s) Hs). + (* transitive *) + intros s s' s'' H; generalize s''; clear s''; elim H. + intros x l s'' H'; inversion_clear H'; auto. + intros x x' l l' E s'' H'; inversion_clear H'; auto. + constructor 2. transitivity x'; auto. + constructor 2. rewrite <- H0; auto. + intros. + inversion_clear H3. + constructor 2. rewrite H0; auto. + constructor 3; auto. transitivity y; auto. unfold lt in *; auto. + Qed. + + Instance lt_compat' : + Proper (eqlistA O.eq==>eqlistA O.eq==>iff) lt. + Proof. + apply proper_sym_impl_iff_2; auto with *. + intros s1 s1' E1 s2 s2' E2 H. + revert s1' E1 s2' E2. + induction H; intros; inversion_clear E1; inversion_clear E2. + constructor 1. + constructor 2. MO.order. + constructor 3. MO.order. unfold lt in *; auto. + Qed. + + Lemma eq_cons : + forall l1 l2 x y, + O.eq x y -> eq l1 l2 -> eq (x :: l1) (y :: l2). + Proof. + unfold eq; intros l1 l2 x y Exy E12 z. + split; inversion_clear 1. + left; MO.order. right; rewrite <- E12; auto. + left; MO.order. right; rewrite E12; auto. + Qed. + Hint Resolve eq_cons. + + Lemma cons_CompSpec : forall c x1 x2 l1 l2, O.eq x1 x2 -> + CompSpec eq lt l1 l2 c -> CompSpec eq lt (x1::l1) (x2::l2) c. + Proof. + destruct c; simpl; inversion_clear 2; auto with relations. + Qed. + Hint Resolve cons_CompSpec. + +End MakeListOrdering. diff --git a/theories/MSets/MSetList.v b/theories/MSets/MSetList.v new file mode 100644 index 00000000..48af7e6a --- /dev/null +++ b/theories/MSets/MSetList.v @@ -0,0 +1,899 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) +(* \VV/ *************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(***********************************************************************) + +(* $Id$ *) + +(** * Finite sets library *) + +(** This file proposes an implementation of the non-dependant + interface [MSetInterface.S] using strictly ordered list. *) + +Require Export MSetInterface OrdersFacts OrdersLists. +Set Implicit Arguments. +Unset Strict Implicit. + +(** * Functions over lists + + First, we provide sets as lists which are not necessarily sorted. + The specs are proved under the additional condition of being sorted. + And the functions returning sets are proved to preserve this invariant. *) + +Module Ops (X:OrderedType) <: WOps X. + + Definition elt := X.t. + Definition t := list elt. + + Definition empty : t := nil. + + Definition is_empty (l : t) := if l then true else false. + + (** ** The set operations. *) + + Fixpoint mem x s := + match s with + | nil => false + | y :: l => + match X.compare x y with + | Lt => false + | Eq => true + | Gt => mem x l + end + end. + + Fixpoint add x s := + match s with + | nil => x :: nil + | y :: l => + match X.compare x y with + | Lt => x :: s + | Eq => s + | Gt => y :: add x l + end + end. + + Definition singleton (x : elt) := x :: nil. + + Fixpoint remove x s := + match s with + | nil => nil + | y :: l => + match X.compare x y with + | Lt => s + | Eq => l + | Gt => y :: remove x l + end + end. + + Fixpoint union (s : t) : t -> t := + match s with + | nil => fun s' => s' + | x :: l => + (fix union_aux (s' : t) : t := + match s' with + | nil => s + | x' :: l' => + match X.compare x x' with + | Lt => x :: union l s' + | Eq => x :: union l l' + | Gt => x' :: union_aux l' + end + end) + end. + + Fixpoint inter (s : t) : t -> t := + match s with + | nil => fun _ => nil + | x :: l => + (fix inter_aux (s' : t) : t := + match s' with + | nil => nil + | x' :: l' => + match X.compare x x' with + | Lt => inter l s' + | Eq => x :: inter l l' + | Gt => inter_aux l' + end + end) + end. + + Fixpoint diff (s : t) : t -> t := + match s with + | nil => fun _ => nil + | x :: l => + (fix diff_aux (s' : t) : t := + match s' with + | nil => s + | x' :: l' => + match X.compare x x' with + | Lt => x :: diff l s' + | Eq => diff l l' + | Gt => diff_aux l' + end + end) + end. + + Fixpoint equal (s : t) : t -> bool := + fun s' : t => + match s, s' with + | nil, nil => true + | x :: l, x' :: l' => + match X.compare x x' with + | Eq => equal l l' + | _ => false + end + | _, _ => false + end. + + Fixpoint subset s s' := + match s, s' with + | nil, _ => true + | x :: l, x' :: l' => + match X.compare x x' with + | Lt => false + | Eq => subset l l' + | Gt => subset s l' + end + | _, _ => false + end. + + Definition fold (B : Type) (f : elt -> B -> B) (s : t) (i : B) : B := + fold_left (flip f) s i. + + Fixpoint filter (f : elt -> bool) (s : t) : t := + match s with + | nil => nil + | x :: l => if f x then x :: filter f l else filter f l + end. + + Fixpoint for_all (f : elt -> bool) (s : t) : bool := + match s with + | nil => true + | x :: l => if f x then for_all f l else false + end. + + Fixpoint exists_ (f : elt -> bool) (s : t) : bool := + match s with + | nil => false + | x :: l => if f x then true else exists_ f l + end. + + Fixpoint partition (f : elt -> bool) (s : t) : t * t := + match s with + | nil => (nil, nil) + | x :: l => + let (s1, s2) := partition f l in + if f x then (x :: s1, s2) else (s1, x :: s2) + end. + + Definition cardinal (s : t) : nat := length s. + + Definition elements (x : t) : list elt := x. + + Definition min_elt (s : t) : option elt := + match s with + | nil => None + | x :: _ => Some x + end. + + Fixpoint max_elt (s : t) : option elt := + match s with + | nil => None + | x :: nil => Some x + | _ :: l => max_elt l + end. + + Definition choose := min_elt. + + Fixpoint compare s s' := + match s, s' with + | nil, nil => Eq + | nil, _ => Lt + | _, nil => Gt + | x::s, x'::s' => + match X.compare x x' with + | Eq => compare s s' + | Lt => Lt + | Gt => Gt + end + end. + +End Ops. + +Module MakeRaw (X: OrderedType) <: RawSets X. + Module Import MX := OrderedTypeFacts X. + Module Import ML := OrderedTypeLists X. + + Include Ops X. + + (** ** Proofs of set operation specifications. *) + + Section ForNotations. + + Definition inf x l := + match l with + | nil => true + | y::_ => match X.compare x y with Lt => true | _ => false end + end. + + Fixpoint isok l := + match l with + | nil => true + | x::l => inf x l && isok l + end. + + Notation Sort l := (isok l = true). + Notation Inf := (lelistA X.lt). + Notation In := (InA X.eq). + + (* TODO: modify proofs in order to avoid these hints *) + Hint Resolve (@Equivalence_Reflexive _ _ X.eq_equiv). + Hint Immediate (@Equivalence_Symmetric _ _ X.eq_equiv). + Hint Resolve (@Equivalence_Transitive _ _ X.eq_equiv). + + Definition IsOk s := Sort s. + + Class Ok (s:t) : Prop := ok : Sort s. + + Hint Resolve @ok. + Hint Unfold Ok. + + Instance Sort_Ok s `(Hs : Sort s) : Ok s := { ok := Hs }. + + Lemma inf_iff : forall x l, Inf x l <-> inf x l = true. + Proof. + intros x l; split; intro H. + (* -> *) + destruct H; simpl in *. + reflexivity. + rewrite <- compare_lt_iff in H; rewrite H; reflexivity. + (* <- *) + destruct l as [|y ys]; simpl in *. + constructor; fail. + revert H; case_eq (X.compare x y); try discriminate; []. + intros Ha _. + rewrite compare_lt_iff in Ha. + constructor; assumption. + Qed. + + Lemma isok_iff : forall l, sort X.lt l <-> Ok l. + Proof. + intro l; split; intro H. + (* -> *) + elim H. + constructor; fail. + intros y ys Ha Hb Hc. + change (inf y ys && isok ys = true). + rewrite inf_iff in Hc. + rewrite andb_true_iff; tauto. + (* <- *) + induction l as [|x xs]. + constructor. + change (inf x xs && isok xs = true) in H. + rewrite andb_true_iff, <- inf_iff in H. + destruct H; constructor; tauto. + Qed. + + Hint Extern 1 (Ok _) => rewrite <- isok_iff. + + Ltac inv_ok := match goal with + | H:sort X.lt (_ :: _) |- _ => inversion_clear H; inv_ok + | H:sort X.lt nil |- _ => clear H; inv_ok + | H:sort X.lt ?l |- _ => change (Ok l) in H; inv_ok + | H:Ok _ |- _ => rewrite <- isok_iff in H; inv_ok + | |- Ok _ => rewrite <- isok_iff + | _ => idtac + end. + + Ltac inv := invlist InA; inv_ok; invlist lelistA. + Ltac constructors := repeat constructor. + + Ltac sort_inf_in := match goal with + | H:Inf ?x ?l, H':In ?y ?l |- _ => + cut (X.lt x y); [ intro | apply Sort_Inf_In with l; auto] + | _ => fail + end. + + Global Instance isok_Ok s `(isok s = true) : Ok s | 10. + Proof. + intros. assumption. + Qed. + + Definition Equal s s' := forall a : elt, In a s <-> In a s'. + Definition Subset s s' := forall a : elt, In a s -> In a s'. + Definition Empty s := forall a : elt, ~ In a s. + Definition For_all (P : elt -> Prop) s := forall x, In x s -> P x. + Definition Exists (P : elt -> Prop) (s : t) := exists x, In x s /\ P x. + + Lemma mem_spec : + forall (s : t) (x : elt) (Hs : Ok s), mem x s = true <-> In x s. + Proof. + induction s; intros x Hs; inv; simpl. + intuition. discriminate. inv. + elim_compare x a; rewrite InA_cons; intuition; try order. + discriminate. + sort_inf_in. order. + rewrite <- IHs; auto. + rewrite IHs; auto. + Qed. + + Lemma add_inf : + forall (s : t) (x a : elt), Inf a s -> X.lt a x -> Inf a (add x s). + Proof. + simple induction s; simpl. + intuition. + intros; elim_compare x a; inv; intuition. + Qed. + Hint Resolve add_inf. + + Global Instance add_ok s x : forall `(Ok s), Ok (add x s). + Proof. + repeat rewrite <- isok_iff; revert s x. + simple induction s; simpl. + intuition. + intros; elim_compare x a; inv; auto. + Qed. + + Lemma add_spec : + forall (s : t) (x y : elt) (Hs : Ok s), + In y (add x s) <-> X.eq y x \/ In y s. + Proof. + induction s; simpl; intros. + intuition. inv; auto. + elim_compare x a; inv; rewrite !InA_cons, ?IHs; intuition. + left; order. + Qed. + + Lemma remove_inf : + forall (s : t) (x a : elt) (Hs : Ok s), Inf a s -> Inf a (remove x s). + Proof. + induction s; simpl. + intuition. + intros; elim_compare x a; inv; auto. + apply Inf_lt with a; auto. + Qed. + Hint Resolve remove_inf. + + Global Instance remove_ok s x : forall `(Ok s), Ok (remove x s). + Proof. + repeat rewrite <- isok_iff; revert s x. + induction s; simpl. + intuition. + intros; elim_compare x a; inv; auto. + Qed. + + Lemma remove_spec : + forall (s : t) (x y : elt) (Hs : Ok s), + In y (remove x s) <-> In y s /\ ~X.eq y x. + Proof. + induction s; simpl; intros. + intuition; inv; auto. + elim_compare x a; inv; rewrite !InA_cons, ?IHs; intuition; + try sort_inf_in; try order. + Qed. + + Global Instance singleton_ok x : Ok (singleton x). + Proof. + unfold singleton; simpl; auto. + Qed. + + Lemma singleton_spec : forall x y : elt, In y (singleton x) <-> X.eq y x. + Proof. + unfold singleton; simpl; split; intros; inv; auto. + Qed. + + Ltac induction2 := + simple induction s; + [ simpl; auto; try solve [ intros; inv ] + | intros x l Hrec; simple induction s'; + [ simpl; auto; try solve [ intros; inv ] + | intros x' l' Hrec'; simpl; elim_compare x x'; intros; inv; auto ]]. + + Lemma union_inf : + forall (s s' : t) (a : elt) (Hs : Ok s) (Hs' : Ok s'), + Inf a s -> Inf a s' -> Inf a (union s s'). + Proof. + induction2. + Qed. + Hint Resolve union_inf. + + Global Instance union_ok s s' : forall `(Ok s, Ok s'), Ok (union s s'). + Proof. + repeat rewrite <- isok_iff; revert s s'. + induction2; constructors; try apply @ok; auto. + apply Inf_eq with x'; auto; apply union_inf; auto; apply Inf_eq with x; auto. + change (Inf x' (union (x :: l) l')); auto. + Qed. + + Lemma union_spec : + forall (s s' : t) (x : elt) (Hs : Ok s) (Hs' : Ok s'), + In x (union s s') <-> In x s \/ In x s'. + Proof. + induction2; try rewrite ?InA_cons, ?Hrec, ?Hrec'; intuition; inv; auto. + left; order. + Qed. + + Lemma inter_inf : + forall (s s' : t) (a : elt) (Hs : Ok s) (Hs' : Ok s'), + Inf a s -> Inf a s' -> Inf a (inter s s'). + Proof. + induction2. + apply Inf_lt with x; auto. + apply Hrec'; auto. + apply Inf_lt with x'; auto. + Qed. + Hint Resolve inter_inf. + + Global Instance inter_ok s s' : forall `(Ok s, Ok s'), Ok (inter s s'). + Proof. + repeat rewrite <- isok_iff; revert s s'. + induction2. + constructors; auto. + apply Inf_eq with x'; auto; apply inter_inf; auto; apply Inf_eq with x; auto. + Qed. + + Lemma inter_spec : + forall (s s' : t) (x : elt) (Hs : Ok s) (Hs' : Ok s'), + In x (inter s s') <-> In x s /\ In x s'. + Proof. + induction2; try rewrite ?InA_cons, ?Hrec, ?Hrec'; intuition; inv; auto; + try sort_inf_in; try order. + left; order. + Qed. + + Lemma diff_inf : + forall (s s' : t) (Hs : Ok s) (Hs' : Ok s') (a : elt), + Inf a s -> Inf a s' -> Inf a (diff s s'). + Proof. + intros s s'; repeat rewrite <- isok_iff; revert s s'. + induction2. + apply Hrec; trivial. + apply Inf_lt with x; auto. + apply Inf_lt with x'; auto. + apply Hrec'; auto. + apply Inf_lt with x'; auto. + Qed. + Hint Resolve diff_inf. + + Global Instance diff_ok s s' : forall `(Ok s, Ok s'), Ok (diff s s'). + Proof. + repeat rewrite <- isok_iff; revert s s'. + induction2. + Qed. + + Lemma diff_spec : + forall (s s' : t) (x : elt) (Hs : Ok s) (Hs' : Ok s'), + In x (diff s s') <-> In x s /\ ~In x s'. + Proof. + induction2; try rewrite ?InA_cons, ?Hrec, ?Hrec'; intuition; inv; auto; + try sort_inf_in; try order. + right; intuition; inv; auto. + Qed. + + Lemma equal_spec : + forall (s s' : t) (Hs : Ok s) (Hs' : Ok s'), + equal s s' = true <-> Equal s s'. + Proof. + induction s as [ | x s IH]; intros [ | x' s'] Hs Hs'; simpl. + intuition. + split; intros H. discriminate. assert (In x' nil) by (rewrite H; auto). inv. + split; intros H. discriminate. assert (In x nil) by (rewrite <-H; auto). inv. + inv. + elim_compare x x' as C; try discriminate. + (* x=x' *) + rewrite IH; auto. + split; intros E y; specialize (E y). + rewrite !InA_cons, E, C; intuition. + rewrite !InA_cons, C in E. intuition; try sort_inf_in; order. + (* x<x' *) + split; intros E. discriminate. + assert (In x (x'::s')) by (rewrite <- E; auto). + inv; try sort_inf_in; order. + (* x>x' *) + split; intros E. discriminate. + assert (In x' (x::s)) by (rewrite E; auto). + inv; try sort_inf_in; order. + Qed. + + Lemma subset_spec : + forall (s s' : t) (Hs : Ok s) (Hs' : Ok s'), + subset s s' = true <-> Subset s s'. + Proof. + intros s s'; revert s. + induction s' as [ | x' s' IH]; intros [ | x s] Hs Hs'; simpl; auto. + split; try red; intros; auto. + split; intros H. discriminate. assert (In x nil) by (apply H; auto). inv. + split; try red; intros; auto. inv. + inv. elim_compare x x' as C. + (* x=x' *) + rewrite IH; auto. + split; intros S y; specialize (S y). + rewrite !InA_cons, C. intuition. + rewrite !InA_cons, C in S. intuition; try sort_inf_in; order. + (* x<x' *) + split; intros S. discriminate. + assert (In x (x'::s')) by (apply S; auto). + inv; try sort_inf_in; order. + (* x>x' *) + rewrite IH; auto. + split; intros S y; specialize (S y). + rewrite !InA_cons. intuition. + rewrite !InA_cons in S. rewrite !InA_cons. intuition; try sort_inf_in; order. + Qed. + + Global Instance empty_ok : Ok empty. + Proof. + constructors. + Qed. + + Lemma empty_spec : Empty empty. + Proof. + unfold Empty, empty; intuition; inv. + Qed. + + Lemma is_empty_spec : forall s : t, is_empty s = true <-> Empty s. + Proof. + intros [ | x s]; simpl. + split; auto. intros _ x H. inv. + split. discriminate. intros H. elim (H x); auto. + Qed. + + Lemma elements_spec1 : forall (s : t) (x : elt), In x (elements s) <-> In x s. + Proof. + intuition. + Qed. + + Lemma elements_spec2 : forall (s : t) (Hs : Ok s), sort X.lt (elements s). + Proof. + intro s; repeat rewrite <- isok_iff; auto. + Qed. + + Lemma elements_spec2w : forall (s : t) (Hs : Ok s), NoDupA X.eq (elements s). + Proof. + intro s; repeat rewrite <- isok_iff; auto. + Qed. + + Lemma min_elt_spec1 : forall (s : t) (x : elt), min_elt s = Some x -> In x s. + Proof. + destruct s; simpl; inversion 1; auto. + Qed. + + Lemma min_elt_spec2 : + forall (s : t) (x y : elt) (Hs : Ok s), + min_elt s = Some x -> In y s -> ~ X.lt y x. + Proof. + induction s as [ | x s IH]; simpl; inversion 2; subst. + intros; inv; try sort_inf_in; order. + Qed. + + Lemma min_elt_spec3 : forall s : t, min_elt s = None -> Empty s. + Proof. + destruct s; simpl; red; intuition. inv. discriminate. + Qed. + + Lemma max_elt_spec1 : forall (s : t) (x : elt), max_elt s = Some x -> In x s. + Proof. + induction s as [ | x s IH]. inversion 1. + destruct s as [ | y s]. simpl. inversion 1; subst; auto. + right; apply IH; auto. + Qed. + + Lemma max_elt_spec2 : + forall (s : t) (x y : elt) (Hs : Ok s), + max_elt s = Some x -> In y s -> ~ X.lt x y. + Proof. + induction s as [ | a s IH]. inversion 2. + destruct s as [ | b s]. inversion 2; subst. intros; inv; order. + intros. inv; auto. + assert (~X.lt x b) by (apply IH; auto). + assert (X.lt a b) by auto. + order. + Qed. + + Lemma max_elt_spec3 : forall s : t, max_elt s = None -> Empty s. + Proof. + induction s as [ | a s IH]. red; intuition; inv. + destruct s as [ | b s]. inversion 1. + intros; elim IH with b; auto. + Qed. + + Definition choose_spec1 : + forall (s : t) (x : elt), choose s = Some x -> In x s := min_elt_spec1. + + Definition choose_spec2 : + forall s : t, choose s = None -> Empty s := min_elt_spec3. + + Lemma choose_spec3: forall s s' x x', Ok s -> Ok s' -> + choose s = Some x -> choose s' = Some x' -> Equal s s' -> X.eq x x'. + Proof. + unfold choose; intros s s' x x' Hs Hs' Hx Hx' H. + assert (~X.lt x x'). + apply min_elt_spec2 with s'; auto. + rewrite <-H; auto using min_elt_spec1. + assert (~X.lt x' x). + apply min_elt_spec2 with s; auto. + rewrite H; auto using min_elt_spec1. + order. + Qed. + + Lemma fold_spec : + forall (s : t) (A : Type) (i : A) (f : elt -> A -> A), + fold f s i = fold_left (flip f) (elements s) i. + Proof. + reflexivity. + Qed. + + Lemma cardinal_spec : + forall (s : t) (Hs : Ok s), + cardinal s = length (elements s). + Proof. + auto. + Qed. + + Lemma filter_inf : + forall (s : t) (x : elt) (f : elt -> bool) (Hs : Ok s), + Inf x s -> Inf x (filter f s). + Proof. + simple induction s; simpl. + intuition. + intros x l Hrec a f Hs Ha; inv. + case (f x); auto. + apply Hrec; auto. + apply Inf_lt with x; auto. + Qed. + + Global Instance filter_ok s f : forall `(Ok s), Ok (filter f s). + Proof. + repeat rewrite <- isok_iff; revert s f. + simple induction s; simpl. + auto. + intros x l Hrec f Hs; inv. + case (f x); auto. + constructors; auto. + apply filter_inf; auto. + Qed. + + Lemma filter_spec : + forall (s : t) (x : elt) (f : elt -> bool), + Proper (X.eq==>eq) f -> + (In x (filter f s) <-> In x s /\ f x = true). + Proof. + induction s; simpl; intros. + split; intuition; inv. + destruct (f a) as [ ]_eqn:F; rewrite !InA_cons, ?IHs; intuition. + setoid_replace x with a; auto. + setoid_replace a with x in F; auto; congruence. + Qed. + + Lemma for_all_spec : + forall (s : t) (f : elt -> bool), + Proper (X.eq==>eq) f -> + (for_all f s = true <-> For_all (fun x => f x = true) s). + Proof. + unfold For_all; induction s; simpl; intros. + split; intros; auto. inv. + destruct (f a) as [ ]_eqn:F. + rewrite IHs; auto. firstorder. inv; auto. + setoid_replace x with a; auto. + split; intros H'. discriminate. + rewrite H' in F; auto. + Qed. + + Lemma exists_spec : + forall (s : t) (f : elt -> bool), + Proper (X.eq==>eq) f -> + (exists_ f s = true <-> Exists (fun x => f x = true) s). + Proof. + unfold Exists; induction s; simpl; intros. + firstorder. discriminate. inv. + destruct (f a) as [ ]_eqn:F. + firstorder. + rewrite IHs; auto. + firstorder. + inv. + setoid_replace a with x in F; auto; congruence. + exists x; auto. + Qed. + + Lemma partition_inf1 : + forall (s : t) (f : elt -> bool) (x : elt) (Hs : Ok s), + Inf x s -> Inf x (fst (partition f s)). + Proof. + intros s f x; repeat rewrite <- isok_iff; revert s f x. + simple induction s; simpl. + intuition. + intros x l Hrec f a Hs Ha; inv. + generalize (Hrec f a H). + case (f x); case (partition f l); simpl. + auto. + intros; apply H2; apply Inf_lt with x; auto. + Qed. + + Lemma partition_inf2 : + forall (s : t) (f : elt -> bool) (x : elt) (Hs : Ok s), + Inf x s -> Inf x (snd (partition f s)). + Proof. + intros s f x; repeat rewrite <- isok_iff; revert s f x. + simple induction s; simpl. + intuition. + intros x l Hrec f a Hs Ha; inv. + generalize (Hrec f a H). + case (f x); case (partition f l); simpl. + intros; apply H2; apply Inf_lt with x; auto. + auto. + Qed. + + Global Instance partition_ok1 s f : forall `(Ok s), Ok (fst (partition f s)). + Proof. + repeat rewrite <- isok_iff; revert s f. + simple induction s; simpl. + auto. + intros x l Hrec f Hs; inv. + generalize (Hrec f H); generalize (@partition_inf1 l f x). + case (f x); case (partition f l); simpl; auto. + Qed. + + Global Instance partition_ok2 s f : forall `(Ok s), Ok (snd (partition f s)). + Proof. + repeat rewrite <- isok_iff; revert s f. + simple induction s; simpl. + auto. + intros x l Hrec f Hs; inv. + generalize (Hrec f H); generalize (@partition_inf2 l f x). + case (f x); case (partition f l); simpl; auto. + Qed. + + Lemma partition_spec1 : + forall (s : t) (f : elt -> bool), + Proper (X.eq==>eq) f -> Equal (fst (partition f s)) (filter f s). + Proof. + simple induction s; simpl; auto; unfold Equal. + split; auto. + intros x l Hrec f Hf. + generalize (Hrec f Hf); clear Hrec. + destruct (partition f l) as [s1 s2]; simpl; intros. + case (f x); simpl; auto. + split; inversion_clear 1; auto. + constructor 2; rewrite <- H; auto. + constructor 2; rewrite H; auto. + Qed. + + Lemma partition_spec2 : + forall (s : t) (f : elt -> bool), + Proper (X.eq==>eq) f -> + Equal (snd (partition f s)) (filter (fun x => negb (f x)) s). + Proof. + simple induction s; simpl; auto; unfold Equal. + split; auto. + intros x l Hrec f Hf. + generalize (Hrec f Hf); clear Hrec. + destruct (partition f l) as [s1 s2]; simpl; intros. + case (f x); simpl; auto. + split; inversion_clear 1; auto. + constructor 2; rewrite <- H; auto. + constructor 2; rewrite H; auto. + Qed. + + End ForNotations. + + Definition In := InA X.eq. + Instance In_compat : Proper (X.eq==>eq==> iff) In. + Proof. repeat red; intros; rewrite H, H0; auto. Qed. + + Module L := MakeListOrdering 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'. + + Instance lt_strorder : StrictOrder lt. + Proof. + split. + intros s (s1 & s2 & B1 & B2 & E1 & E2 & L). + repeat rewrite <- isok_iff in *. + assert (eqlistA X.eq s1 s2). + apply SortA_equivlistA_eqlistA with (ltA:=X.lt); auto using @ok with *. + transitivity s; auto. symmetry; auto. + rewrite H in L. + apply (StrictOrder_Irreflexive s2); auto. + intros s1 s2 s3 (s1' & s2' & B1 & B2 & E1 & E2 & L12) + (s2'' & s3' & B2' & B3 & E2' & E3 & L23). + exists s1', s3'. + repeat rewrite <- isok_iff in *. + do 4 (split; trivial). + assert (eqlistA X.eq s2' s2''). + apply SortA_equivlistA_eqlistA with (ltA:=X.lt); auto using @ok with *. + transitivity s2; auto. symmetry; auto. + transitivity s2'; auto. + rewrite H; auto. + Qed. + + Instance lt_compat : Proper (eq==>eq==>iff) lt. + Proof. + intros s1 s2 E12 s3 s4 E34. split. + intros (s1' & s3' & B1 & B3 & E1 & E3 & LT). + exists s1', s3'; do 2 (split; trivial). + split. transitivity s1; auto. symmetry; auto. + split; auto. transitivity s3; auto. symmetry; auto. + intros (s1' & s3' & B1 & B3 & E1 & E3 & LT). + exists s1', s3'; do 2 (split; trivial). + split. transitivity s2; auto. + split; auto. transitivity s4; auto. + Qed. + + Lemma compare_spec_aux : forall s s', CompSpec eq L.lt s s' (compare s s'). + Proof. + induction s as [|x s IH]; intros [|x' s']; simpl; intuition. + elim_compare x x'; auto. + Qed. + + Lemma compare_spec : forall s s', Ok s -> Ok s' -> + CompSpec eq lt s s' (compare s s'). + Proof. + intros s s' Hs Hs'. + destruct (compare_spec_aux s s'); constructor; auto. + exists s, s'; repeat split; auto using @ok. + exists s', s; repeat split; auto using @ok. + Qed. + +End MakeRaw. + +(** * Encapsulation + + Now, in order to really provide a functor implementing [S], we + need to encapsulate everything into a type of strictly ordered lists. *) + +Module Make (X: OrderedType) <: S with Module E := X. + Module Raw := MakeRaw X. + Include Raw2Sets X Raw. +End Make. + +(** For this specific implementation, eq coincides with Leibniz equality *) + +Require Eqdep_dec. + +Module Type OrderedTypeWithLeibniz. + Include OrderedType. + Parameter eq_leibniz : forall x y, eq x y -> x = y. +End OrderedTypeWithLeibniz. + +Module Type SWithLeibniz. + Declare Module E : OrderedTypeWithLeibniz. + Include SetsOn E. + Parameter eq_leibniz : forall x y, eq x y -> x = y. +End SWithLeibniz. + +Module MakeWithLeibniz (X: OrderedTypeWithLeibniz) <: SWithLeibniz with Module E := X. + Module E := X. + Module Raw := MakeRaw X. + Include Raw2SetsOn X Raw. + + Lemma eq_leibniz_list : forall xs ys, eqlistA X.eq xs ys -> xs = ys. + Proof. + induction xs as [|x xs]; intros [|y ys] H; inversion H; [ | ]. + reflexivity. + f_equal. + apply X.eq_leibniz; congruence. + apply IHxs; subst; assumption. + Qed. + + Lemma eq_leibniz : forall s s', eq s s' -> s = s'. + Proof. + intros [xs Hxs] [ys Hys] Heq. + change (equivlistA X.eq xs ys) in Heq. + assert (H : eqlistA X.eq xs ys). + rewrite <- Raw.isok_iff in Hxs, Hys. + apply SortA_equivlistA_eqlistA with X.lt; auto with *. + apply eq_leibniz_list in H. + subst ys. + f_equal. + apply Eqdep_dec.eq_proofs_unicity. + intros x y; destruct (bool_dec x y); tauto. + Qed. + +End MakeWithLeibniz. diff --git a/theories/MSets/MSetPositive.v b/theories/MSets/MSetPositive.v new file mode 100644 index 00000000..e83ac27d --- /dev/null +++ b/theories/MSets/MSetPositive.v @@ -0,0 +1,1149 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) +(* \VV/ *************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(***********************************************************************) + +(** Efficient implementation of [MSetInterface.S] for positive keys, + inspired from the [FMapPositive] module. + + This module was adapted by Alexandre Ren, Damien Pous, and Thomas + Braibant (2010, LIG, CNRS, UMR 5217), from the [FMapPositive] + module of Pierre Letouzey and Jean-Christophe FilliĂ¢tre, which in + turn comes from the [FMap] framework of a work by Xavier Leroy and + Sandrine Blazy (used for building certified compilers). +*) + +Require Import Bool BinPos Orders MSetInterface. + +Set Implicit Arguments. + +Local Open Scope lazy_bool_scope. +Local Open Scope positive_scope. + +Local Unset Elimination Schemes. +Local Unset Case Analysis Schemes. +Local Unset Boolean Equality Schemes. + + +(** 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 FSetPositive and FMapPositive. *) + +Module PositiveOrderedTypeBits <: UsualOrderedType. + Definition t:=positive. + Include HasUsualEq <+ UsualIsEq. + Definition eqb := Peqb. + Definition eqb_eq := Peqb_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. + + Lemma bits_lt_antirefl : forall x : positive, ~ bits_lt x x. + Proof. + induction x; simpl; auto. + Qed. + + Lemma bits_lt_trans : + forall x y z : positive, bits_lt x y -> bits_lt y z -> bits_lt x z. + Proof. + induction x; destruct y,z; simpl; eauto; intuition. + Qed. + + Instance lt_compat : Proper (eq==>eq==>iff) lt. + Proof. + intros x x' Hx y y' Hy. rewrite Hx, Hy; intuition. + Qed. + + Instance lt_strorder : StrictOrder lt. + Proof. + split; [ 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. + + Lemma compare_spec : forall x y, CompSpec eq lt x y (compare x y). + Proof. + unfold eq, lt. + induction x; destruct y; try constructor; simpl; auto. + destruct (IHx y); subst; auto. + destruct (IHx y); subst; auto. + Qed. + +End PositiveOrderedTypeBits. + +Module PositiveSet <: S with Module E:=PositiveOrderedTypeBits. + + Module E:=PositiveOrderedTypeBits. + + Definition elt := positive. + + Inductive tree := + | Leaf : tree + | Node : tree -> bool -> tree -> tree. + + Scheme tree_ind := Induction for tree Sort Prop. + + Definition t := tree. + + Definition empty := Leaf. + + Fixpoint is_empty (m : t) : bool := + match m with + | Leaf => true + | Node l b r => negb b &&& is_empty l &&& is_empty r + end. + + Fixpoint mem (i : positive) (m : t) : bool := + match m with + | Leaf => false + | Node l o r => + match i with + | 1 => o + | i~0 => mem i l + | i~1 => mem i r + end + end. + + Fixpoint add (i : positive) (m : t) : t := + match m with + | Leaf => + match i with + | 1 => Node Leaf true Leaf + | i~0 => Node (add i Leaf) false Leaf + | i~1 => Node Leaf false (add i Leaf) + end + | Node l o r => + match i with + | 1 => Node l true r + | i~0 => Node (add i l) o r + | i~1 => Node l o (add i r) + end + end. + + Definition singleton i := add i empty. + + (** helper function to avoid creating empty trees that are not leaves *) + + Definition node l (b: bool) r := + if b then Node l b r else + match l,r with + | Leaf,Leaf => Leaf + | _,_ => Node l false r end. + + Fixpoint remove (i : positive) (m : t) : t := + match m with + | Leaf => Leaf + | Node l o r => + match i with + | 1 => node l false r + | i~0 => node (remove i l) o r + | i~1 => node l o (remove i r) + end + end. + + Fixpoint union (m m': t) := + match m with + | Leaf => m' + | Node l o r => + match m' with + | Leaf => m + | Node l' o' r' => Node (union l l') (o||o') (union r r') + end + end. + + Fixpoint inter (m m': t) := + match m with + | Leaf => Leaf + | Node l o r => + match m' with + | Leaf => Leaf + | Node l' o' r' => node (inter l l') (o&&o') (inter r r') + end + end. + + Fixpoint diff (m m': t) := + match m with + | Leaf => Leaf + | Node l o r => + match m' with + | Leaf => m + | Node l' o' r' => node (diff l l') (o&&negb o') (diff r r') + end + end. + + Fixpoint equal (m m': t): bool := + match m with + | Leaf => is_empty m' + | Node l o r => + match m' with + | Leaf => is_empty m + | Node l' o' r' => eqb o o' &&& equal l l' &&& equal r r' + end + end. + + Fixpoint subset (m m': t): bool := + match m with + | Leaf => true + | Node l o r => + match m' with + | Leaf => is_empty m + | Node l' o' r' => (negb o ||| o') &&& subset l l' &&& subset r r' + end + end. + + (** reverses [y] and concatenate it with [x] *) + + Fixpoint rev_append y x := + match y with + | 1 => x + | y~1 => rev_append y x~1 + | y~0 => rev_append y x~0 + end. + Infix "@" := rev_append (at level 60). + Definition rev x := x@1. + + Section Fold. + + Variables B : Type. + Variable f : positive -> B -> B. + + (** the additional argument, [i], records the current path, in + reverse order (this should be more efficient: we reverse this argument + only at present nodes only, rather than at each node of the tree). + we also use this convention in all functions below + *) + + Fixpoint xfold (m : t) (v : B) (i : positive) := + match m with + | Leaf => v + | Node l true r => + xfold r (f (rev i) (xfold l v i~0)) i~1 + | Node l false r => + xfold r (xfold l v i~0) i~1 + end. + Definition fold m i := xfold m i 1. + + End Fold. + + Section Quantifiers. + + Variable f : positive -> bool. + + Fixpoint xforall (m : t) (i : positive) := + match m with + | Leaf => true + | Node l o r => + (negb o ||| f (rev i)) &&& xforall r i~1 &&& xforall l i~0 + end. + Definition for_all m := xforall m 1. + + Fixpoint xexists (m : t) (i : positive) := + match m with + | Leaf => false + | Node l o r => (o &&& f (rev i)) ||| xexists r i~1 ||| xexists l i~0 + end. + Definition exists_ m := xexists m 1. + + Fixpoint xfilter (m : t) (i : positive) := + match m with + | Leaf => Leaf + | Node l o r => node (xfilter l i~0) (o &&& f (rev i)) (xfilter r i~1) + end. + Definition filter m := xfilter m 1. + + Fixpoint xpartition (m : t) (i : positive) := + match m with + | Leaf => (Leaf,Leaf) + | Node l o r => + let (lt,lf) := xpartition l i~0 in + let (rt,rf) := xpartition r i~1 in + if o then + let fi := f (rev i) in + (node lt fi rt, node lf (negb fi) rf) + else + (node lt false rt, node lf false rf) + end. + Definition partition m := xpartition m 1. + + End Quantifiers. + + (** uses [a] to accumulate values rather than doing a lot of concatenations *) + + Fixpoint xelements (m : t) (i : positive) (a: list positive) := + match m with + | Leaf => a + | Node l false r => xelements l i~0 (xelements r i~1 a) + | Node l true r => xelements l i~0 (rev i :: xelements r i~1 a) + end. + + Definition elements (m : t) := xelements m 1 nil. + + Fixpoint cardinal (m : t) : nat := + match m with + | Leaf => O + | Node l false r => (cardinal l + cardinal r)%nat + | Node l true r => S (cardinal l + cardinal r) + end. + + Definition omap (f: elt -> elt) x := + match x with + | None => None + | Some i => Some (f i) + end. + + (** would it be more efficient to use a path like in the above functions ? *) + + Fixpoint choose (m: t) := + match m with + | Leaf => None + | Node l o r => if o then Some 1 else + match choose l with + | None => omap xI (choose r) + | Some i => Some i~0 + end + end. + + Fixpoint min_elt (m: t) := + match m with + | Leaf => None + | Node l o r => + match min_elt l with + | None => if o then Some 1 else omap xI (min_elt r) + | Some i => Some i~0 + end + end. + + Fixpoint max_elt (m: t) := + match m with + | Leaf => None + | Node l o r => + match max_elt r with + | None => if o then Some 1 else omap xO (max_elt l) + | Some i => Some i~1 + end + end. + + (** lexicographic product, defined using a notation to keep things lazy *) + + Notation lex u v := match u with Eq => v | Lt => Lt | Gt => Gt end. + + Definition compare_bool a b := + match a,b with + | false, true => Lt + | true, false => Gt + | _,_ => Eq + end. + + Fixpoint compare (m m': t): comparison := + match m,m' with + | Leaf,_ => if is_empty m' then Eq else Lt + | _,Leaf => if is_empty m then Eq else Gt + | Node l o r,Node l' o' r' => + lex (compare_bool o o') (lex (compare l l') (compare r r')) + end. + + + Definition In i t := mem i t = true. + Definition Equal s s' := forall a : elt, In a s <-> In a s'. + Definition Subset s s' := forall a : elt, In a s -> In a s'. + Definition Empty s := forall a : elt, ~ In a s. + Definition For_all (P : elt -> Prop) s := forall x, In x s -> P x. + Definition Exists (P : elt -> Prop) s := exists x, In x s /\ P x. + + Notation "s [=] t" := (Equal s t) (at level 70, no associativity). + Notation "s [<=] t" := (Subset s t) (at level 70, no associativity). + + Definition eq := Equal. + Definition lt m m' := compare m m' = Lt. + + (** Specification of [In] *) + + Instance In_compat : Proper (E.eq==>Logic.eq==>iff) In. + Proof. + intros s s' Hs x x' Hx. rewrite Hs, Hx; intuition. + Qed. + + (** Specification of [eq] *) + + Local Instance eq_equiv : Equivalence eq. + Proof. firstorder. Qed. + + (** Specification of [mem] *) + + Lemma mem_spec: forall s x, mem x s = true <-> In x s. + Proof. unfold In. intuition. Qed. + + (** Additional lemmas for mem *) + + Lemma mem_Leaf: forall x, mem x Leaf = false. + Proof. destruct x; trivial. Qed. + + (** Specification of [empty] *) + + Lemma empty_spec : Empty empty. + Proof. unfold Empty, In. intro. rewrite mem_Leaf. discriminate. Qed. + + (** Specification of node *) + + Lemma mem_node: forall x l o r, mem x (node l o r) = mem x (Node l o r). + Proof. + intros x l o r. + case o; trivial. + destruct l; trivial. + destruct r; trivial. + symmetry. destruct x. + apply mem_Leaf. + apply mem_Leaf. + reflexivity. + Qed. + Local Opaque node. + + (** Specification of [is_empty] *) + + Lemma is_empty_spec: forall s, is_empty s = true <-> Empty s. + Proof. + unfold Empty, In. + induction s as [|l IHl o r IHr]; simpl. + setoid_rewrite mem_Leaf. firstorder. + rewrite <- 2andb_lazy_alt, 2andb_true_iff, IHl, IHr. clear IHl IHr. + destruct o; simpl; split. + intuition discriminate. + intro H. elim (H 1). reflexivity. + intros H [a|a|]; apply H || intro; discriminate. + intro H. split. split. reflexivity. + intro a. apply (H a~0). + intro a. apply (H a~1). + Qed. + + (** Specification of [subset] *) + + Lemma subset_Leaf_s: forall s, Leaf [<=] s. + Proof. intros s i Hi. apply empty_spec in Hi. elim Hi. Qed. + + Lemma subset_spec: forall s s', subset s s' = true <-> s [<=] s'. + Proof. + induction s as [|l IHl o r IHr]; intros [|l' o' r']; simpl. + split; intros. apply subset_Leaf_s. reflexivity. + + split; intros. apply subset_Leaf_s. reflexivity. + + rewrite <- 2andb_lazy_alt, 2andb_true_iff, 2is_empty_spec. + destruct o; simpl. + split. + intuition discriminate. + intro H. elim (@empty_spec 1). apply H. reflexivity. + split; intro H. + destruct H as [[_ Hl] Hr]. + intros [i|i|] Hi. + elim (Hr i Hi). + elim (Hl i Hi). + discriminate. + split. split. reflexivity. + unfold Empty. intros a H1. apply (@empty_spec (a~0)), H. assumption. + unfold Empty. intros a H1. apply (@empty_spec (a~1)), H. assumption. + + rewrite <- 2andb_lazy_alt, 2andb_true_iff, IHl, IHr. clear. + destruct o; simpl. + split; intro H. + destruct H as [[Ho' Hl] Hr]. rewrite Ho'. + intros i Hi. destruct i. + apply (Hr i). assumption. + apply (Hl i). assumption. + assumption. + split. split. + destruct o'; trivial. + specialize (H 1). unfold In in H. simpl in H. apply H. reflexivity. + intros i Hi. apply (H i~0). apply Hi. + intros i Hi. apply (H i~1). apply Hi. + split; intros. + intros i Hi. destruct i; destruct H as [[H Hl] Hr]. + apply (Hr i). assumption. + apply (Hl i). assumption. + discriminate Hi. + split. split. reflexivity. + intros i Hi. apply (H i~0). apply Hi. + intros i Hi. apply (H i~1). apply Hi. + Qed. + + (** Specification of [equal] (via subset) *) + + Lemma equal_subset: forall s s', equal s s' = subset s s' && subset s' s. + Proof. + induction s as [|l IHl o r IHr]; intros [|l' o' r']; simpl; trivial. + destruct o. reflexivity. rewrite andb_comm. reflexivity. + rewrite <- 6andb_lazy_alt. rewrite eq_iff_eq_true. + rewrite 7andb_true_iff, eqb_true_iff. + rewrite IHl, IHr, 2andb_true_iff. clear IHl IHr. intuition subst. + destruct o'; reflexivity. + destruct o'; reflexivity. + destruct o; auto. destruct o'; trivial. + Qed. + + Lemma equal_spec: forall s s', equal s s' = true <-> Equal s s'. + Proof. + intros. rewrite equal_subset. rewrite andb_true_iff. + rewrite 2subset_spec. unfold Equal, Subset. firstorder. + Qed. + + Lemma eq_dec : forall s s', { eq s s' } + { ~ eq s s' }. + Proof. + unfold eq. + intros. case_eq (equal s s'); intro H. + left. apply equal_spec, H. + right. rewrite <- equal_spec. congruence. + Defined. + + (** (Specified) definition of [compare] *) + + Lemma lex_Opp: forall u v u' v', u = CompOpp u' -> v = CompOpp v' -> + lex u v = CompOpp (lex u' v'). + Proof. intros ? ? u' ? -> ->. case u'; reflexivity. Qed. + + Lemma compare_bool_inv: forall b b', + compare_bool b b' = CompOpp (compare_bool b' b). + Proof. intros [|] [|]; reflexivity. Qed. + + Lemma compare_inv: forall s s', compare s s' = CompOpp (compare s' s). + Proof. + induction s as [|l IHl o r IHr]; destruct s' as [|l' o' r']; trivial. + unfold compare. case is_empty; reflexivity. + unfold compare. case is_empty; reflexivity. + simpl. rewrite compare_bool_inv. + case compare_bool; simpl; trivial; apply lex_Opp; auto. + Qed. + + Lemma lex_Eq: forall u v, lex u v = Eq <-> u=Eq /\ v=Eq. + Proof. intros u v; destruct u; intuition discriminate. Qed. + + Lemma compare_bool_Eq: forall b1 b2, + compare_bool b1 b2 = Eq <-> eqb b1 b2 = true. + Proof. intros [|] [|]; intuition discriminate. Qed. + + Lemma compare_equal: forall s s', compare s s' = Eq <-> equal s s' = true. + Proof. + induction s as [|l IHl o r IHr]; destruct s' as [|l' o' r']. + simpl. tauto. + unfold compare, equal. case is_empty; intuition discriminate. + unfold compare, equal. case is_empty; intuition discriminate. + simpl. rewrite <- 2andb_lazy_alt, 2andb_true_iff. + rewrite <- IHl, <- IHr, <- compare_bool_Eq. clear IHl IHr. + rewrite and_assoc. rewrite <- 2lex_Eq. reflexivity. + Qed. + + + Lemma compare_gt: forall s s', compare s s' = Gt -> lt s' s. + Proof. + unfold lt. intros s s'. rewrite compare_inv. + case compare; trivial; intros; discriminate. + Qed. + + Lemma compare_eq: forall s s', compare s s' = Eq -> eq s s'. + Proof. + unfold eq. intros s s'. rewrite compare_equal, equal_spec. trivial. + Qed. + + Lemma compare_spec : forall s s' : t, CompSpec eq lt s s' (compare s s'). + Proof. + intros. case_eq (compare s s'); intro H; constructor. + apply compare_eq, H. + assumption. + apply compare_gt, H. + Qed. + + Section lt_spec. + + Inductive ct: comparison -> comparison -> comparison -> Prop := + | ct_xxx: forall x, ct x x x + | ct_xex: forall x, ct x Eq x + | ct_exx: forall x, ct Eq x x + | ct_glx: forall x, ct Gt Lt x + | ct_lgx: forall x, ct Lt Gt x. + + Lemma ct_cxe: forall x, ct (CompOpp x) x Eq. + Proof. destruct x; constructor. Qed. + + Lemma ct_xce: forall x, ct x (CompOpp x) Eq. + Proof. destruct x; constructor. Qed. + + Lemma ct_lxl: forall x, ct Lt x Lt. + Proof. destruct x; constructor. Qed. + + Lemma ct_gxg: forall x, ct Gt x Gt. + Proof. destruct x; constructor. Qed. + + Lemma ct_xll: forall x, ct x Lt Lt. + Proof. destruct x; constructor. Qed. + + Lemma ct_xgg: forall x, ct x Gt Gt. + Proof. destruct x; constructor. Qed. + + Local Hint Constructors ct: ct. + Local Hint Resolve ct_cxe ct_xce ct_lxl ct_xll ct_gxg ct_xgg: ct. + Ltac ct := trivial with ct. + + Lemma ct_lex: forall u v w u' v' w', + ct u v w -> ct u' v' w' -> ct (lex u u') (lex v v') (lex w w'). + Proof. + intros u v w u' v' w' H H'. + inversion_clear H; inversion_clear H'; ct; destruct w; ct; destruct w'; ct. + Qed. + + Lemma ct_compare_bool: + forall a b c, ct (compare_bool a b) (compare_bool b c) (compare_bool a c). + Proof. + intros [|] [|] [|]; constructor. + Qed. + + Lemma compare_x_Leaf: forall s, + compare s Leaf = if is_empty s then Eq else Gt. + Proof. + intros. rewrite compare_inv. simpl. case (is_empty s); reflexivity. + Qed. + + Lemma compare_empty_x: forall a, is_empty a = true -> + forall b, compare a b = if is_empty b then Eq else Lt. + Proof. + induction a as [|l IHl o r IHr]; trivial. + destruct o. intro; discriminate. + simpl is_empty. rewrite <- andb_lazy_alt, andb_true_iff. + intros [Hl Hr]. + destruct b as [|l' [|] r']; simpl compare; trivial. + rewrite Hl, Hr. trivial. + rewrite (IHl Hl), (IHr Hr). simpl. + case (is_empty l'); case (is_empty r'); trivial. + Qed. + + Lemma compare_x_empty: forall a, is_empty a = true -> + forall b, compare b a = if is_empty b then Eq else Gt. + Proof. + setoid_rewrite <- compare_x_Leaf. + intros. rewrite 2(compare_inv b), (compare_empty_x _ H). reflexivity. + Qed. + + Lemma ct_compare: + forall a b c, ct (compare a b) (compare b c) (compare a c). + Proof. + induction a as [|l IHl o r IHr]; intros s' s''. + destruct s' as [|l' o' r']; destruct s'' as [|l'' o'' r'']; ct. + rewrite compare_inv. ct. + unfold compare at 1. case_eq (is_empty (Node l' o' r')); intro H'. + rewrite (compare_empty_x _ H'). ct. + unfold compare at 2. case_eq (is_empty (Node l'' o'' r'')); intro H''. + rewrite (compare_x_empty _ H''), H'. ct. + ct. + + destruct s' as [|l' o' r']; destruct s'' as [|l'' o'' r'']. + ct. + unfold compare at 2. rewrite compare_x_Leaf. + case_eq (is_empty (Node l o r)); intro H. + rewrite (compare_empty_x _ H). ct. + case_eq (is_empty (Node l'' o'' r'')); intro H''. + rewrite (compare_x_empty _ H''), H. ct. + ct. + + rewrite 2 compare_x_Leaf. + case_eq (is_empty (Node l o r)); intro H. + rewrite compare_inv, (compare_x_empty _ H). ct. + case_eq (is_empty (Node l' o' r')); intro H'. + rewrite (compare_x_empty _ H'), H. ct. + ct. + + simpl compare. apply ct_lex. apply ct_compare_bool. + apply ct_lex; trivial. + Qed. + + End lt_spec. + + Instance lt_strorder : StrictOrder lt. + Proof. + unfold lt. split. + intros x H. + assert (compare x x = Eq). + apply compare_equal, equal_spec. reflexivity. + congruence. + intros a b c. assert (H := ct_compare a b c). + inversion_clear H; trivial; intros; discriminate. + Qed. + + Local Instance compare_compat_1 : Proper (eq==>Logic.eq==>Logic.eq) compare. + Proof. + intros x x' Hx y y' Hy. subst y'. + unfold eq in *. rewrite <- equal_spec, <- compare_equal in *. + assert (C:=ct_compare x x' y). rewrite Hx in C. inversion C; auto. + Qed. + + Instance compare_compat : Proper (eq==>eq==>Logic.eq) compare. + Proof. + intros x x' Hx y y' Hy. rewrite Hx. + rewrite compare_inv, Hy, <- compare_inv. reflexivity. + Qed. + + Local Instance lt_compat : Proper (eq==>eq==>iff) lt. + Proof. + intros x x' Hx y y' Hy. unfold lt. rewrite Hx, Hy. intuition. + Qed. + + (** Specification of [add] *) + + Lemma add_spec: forall s x y, In y (add x s) <-> y=x \/ In y s. + Proof. + unfold In. intros s x y; revert x y s. + induction x; intros [y|y|] [|l o r]; simpl mem; + try (rewrite IHx; clear IHx); rewrite ?mem_Leaf; intuition congruence. + Qed. + + (** Specification of [remove] *) + + Lemma remove_spec: forall s x y, In y (remove x s) <-> In y s /\ y<>x. + Proof. + unfold In. intros s x y; revert x y s. + induction x; intros [y|y|] [|l o r]; simpl remove; rewrite ?mem_node; + simpl mem; try (rewrite IHx; clear IHx); rewrite ?mem_Leaf; + intuition congruence. + Qed. + + (** Specification of [singleton] *) + + Lemma singleton_spec : forall x y, In y (singleton x) <-> y=x. + Proof. + unfold singleton. intros x y. rewrite add_spec. intuition. + unfold In in *. rewrite mem_Leaf in *. discriminate. + Qed. + + (** Specification of [union] *) + + Lemma union_spec: forall s s' x, In x (union s s') <-> In x s \/ In x s'. + Proof. + unfold In. intros s s' x; revert x s s'. + induction x; destruct s; destruct s'; simpl union; simpl mem; + try (rewrite IHx; clear IHx); try intuition congruence. + apply orb_true_iff. + Qed. + + (** Specification of [inter] *) + + Lemma inter_spec: forall s s' x, In x (inter s s') <-> In x s /\ In x s'. + Proof. + unfold In. intros s s' x; revert x s s'. + induction x; destruct s; destruct s'; simpl inter; rewrite ?mem_node; + simpl mem; try (rewrite IHx; clear IHx); try intuition congruence. + apply andb_true_iff. + Qed. + + (** Specification of [diff] *) + + Lemma diff_spec: forall s s' x, In x (diff s s') <-> In x s /\ ~ In x s'. + Proof. + unfold In. intros s s' x; revert x s s'. + induction x; destruct s; destruct s' as [|l' o' r']; simpl diff; + rewrite ?mem_node; simpl mem; + try (rewrite IHx; clear IHx); try intuition congruence. + rewrite andb_true_iff. destruct o'; intuition discriminate. + Qed. + + (** Specification of [fold] *) + + Lemma fold_spec: forall s (A : Type) (i : A) (f : elt -> A -> A), + fold f s i = fold_left (fun a e => f e a) (elements s) i. + Proof. + unfold fold, elements. intros s A i f. revert s i. + set (f' := fun a e => f e a). + assert (H: forall s i j acc, + fold_left f' acc (xfold f s i j) = + fold_left f' (xelements s j acc) i). + + induction s as [|l IHl o r IHr]; intros; trivial. + destruct o; simpl xelements; simpl xfold. + rewrite IHr, <- IHl. reflexivity. + rewrite IHr. apply IHl. + + intros. exact (H s i 1 nil). + Qed. + + (** Specification of [cardinal] *) + + Lemma cardinal_spec: forall s, cardinal s = length (elements s). + Proof. + unfold elements. + assert (H: forall s j acc, + (cardinal s + length acc)%nat = length (xelements s j acc)). + + induction s as [|l IHl b r IHr]; intros j acc; simpl; trivial. destruct b. + rewrite <- IHl. simpl. rewrite <- IHr. + rewrite <- plus_n_Sm, Plus.plus_assoc. reflexivity. + rewrite <- IHl, <- IHr. rewrite Plus.plus_assoc. reflexivity. + + intros. rewrite <- H. simpl. rewrite Plus.plus_comm. reflexivity. + Qed. + + (** Specification of [filter] *) + + Lemma xfilter_spec: forall f s x i, + In x (xfilter f s i) <-> In x s /\ f (i@x) = true. + Proof. + intro f. unfold In. + induction s as [|l IHl o r IHr]; intros x i; simpl xfilter. + rewrite mem_Leaf. intuition discriminate. + rewrite mem_node. destruct x; simpl. + rewrite IHr. reflexivity. + rewrite IHl. reflexivity. + rewrite <- andb_lazy_alt. apply andb_true_iff. + Qed. + + Lemma filter_spec: forall s x f, compat_bool E.eq f -> + (In x (filter f s) <-> In x s /\ f x = true). + Proof. intros. apply xfilter_spec. Qed. + + (** Specification of [for_all] *) + + Lemma xforall_spec: forall f s i, + xforall f s i = true <-> For_all (fun x => f (i@x) = true) s. + Proof. + unfold For_all, In. intro f. + induction s as [|l IHl o r IHr]; intros i; simpl. + setoid_rewrite mem_Leaf. intuition discriminate. + rewrite <- 2andb_lazy_alt, <- orb_lazy_alt, 2 andb_true_iff. + rewrite IHl, IHr. clear IHl IHr. + split. + intros [[Hi Hr] Hl] x. destruct x; simpl; intro H. + apply Hr, H. + apply Hl, H. + rewrite H in Hi. assumption. + intro H; intuition. + specialize (H 1). destruct o. apply H. reflexivity. reflexivity. + apply H. assumption. + apply H. assumption. + Qed. + + Lemma for_all_spec: forall s f, compat_bool E.eq f -> + (for_all f s = true <-> For_all (fun x => f x = true) s). + Proof. intros. apply xforall_spec. Qed. + + (** Specification of [exists] *) + + Lemma xexists_spec: forall f s i, + xexists f s i = true <-> Exists (fun x => f (i@x) = true) s. + Proof. + unfold Exists, In. intro f. + induction s as [|l IHl o r IHr]; intros i; simpl. + setoid_rewrite mem_Leaf. firstorder. + rewrite <- 2orb_lazy_alt, 2orb_true_iff, <- andb_lazy_alt, andb_true_iff. + rewrite IHl, IHr. clear IHl IHr. + split. + intros [[Hi|[x Hr]]|[x Hl]]. + exists 1. exact Hi. + exists x~1. exact Hr. + exists x~0. exact Hl. + intros [[x|x|] H]; eauto. + Qed. + + Lemma exists_spec : forall s f, compat_bool E.eq f -> + (exists_ f s = true <-> Exists (fun x => f x = true) s). + Proof. intros. apply xexists_spec. Qed. + + + (** Specification of [partition] *) + + Lemma partition_filter : forall s f, + partition f s = (filter f s, filter (fun x => negb (f x)) s). + Proof. + unfold partition, filter. intros s f. generalize 1 as j. + induction s as [|l IHl o r IHr]; intro j. + reflexivity. + destruct o; simpl; rewrite IHl, IHr; reflexivity. + Qed. + + Lemma partition_spec1 : forall s f, compat_bool E.eq f -> + Equal (fst (partition f s)) (filter f s). + Proof. intros. rewrite partition_filter. reflexivity. Qed. + + Lemma partition_spec2 : forall s f, compat_bool E.eq f -> + Equal (snd (partition f s)) (filter (fun x => negb (f x)) s). + Proof. intros. rewrite partition_filter. reflexivity. Qed. + + + (** Specification of [elements] *) + + Notation InL := (InA E.eq). + + Lemma xelements_spec: forall s j acc y, + InL y (xelements s j acc) + <-> + InL y acc \/ exists x, y=(j@x) /\ mem x s = true. + Proof. + induction s as [|l IHl o r IHr]; simpl. + intros. split; intro H. + left. assumption. + destruct H as [H|[x [Hx Hx']]]. assumption. elim (empty_spec Hx'). + + intros j acc y. case o. + rewrite IHl. rewrite InA_cons. rewrite IHr. clear IHl IHr. split. + intros [[H|[H|[x [-> H]]]]|[x [-> H]]]; eauto. + right. exists x~1. auto. + right. exists x~0. auto. + intros [H|[x [-> H]]]. + eauto. + destruct x. + left. right. right. exists x; auto. + right. exists x; auto. + left. left. reflexivity. + + rewrite IHl, IHr. clear IHl IHr. split. + intros [[H|[x [-> H]]]|[x [-> H]]]. + eauto. + right. exists x~1. auto. + right. exists x~0. auto. + intros [H|[x [-> H]]]. + eauto. + destruct x. + left. right. exists x; auto. + right. exists x; auto. + discriminate. + Qed. + + Lemma elements_spec1: forall s x, InL x (elements s) <-> In x s. + Proof. + unfold elements. intros. rewrite xelements_spec. + split; [ intros [A|(y & B & C)] | intros IN ]. + inversion A. simpl in *. congruence. + right. exists x. auto. + Qed. + + Lemma lt_rev_append: forall j x y, E.lt x y -> E.lt (j@x) (j@y). + Proof. induction j; intros; simpl; auto. Qed. + + Lemma elements_spec2: forall s, sort E.lt (elements s). + Proof. + unfold elements. + assert (H: forall s j acc, + sort E.lt acc -> + (forall x y, In x s -> InL y acc -> E.lt (j@x) y) -> + sort E.lt (xelements s j acc)). + + induction s as [|l IHl o r IHr]; simpl; trivial. + intros j acc Hacc Hsacc. destruct o. + apply IHl. constructor. + apply IHr. apply Hacc. + intros x y Hx Hy. apply Hsacc; assumption. + case_eq (xelements r j~1 acc). constructor. + intros z q H. constructor. + assert (H': InL z (xelements r j~1 acc)). + rewrite H. constructor. reflexivity. + clear H q. rewrite xelements_spec in H'. destruct H' as [Hy|[x [-> Hx]]]. + apply (Hsacc 1 z); trivial. reflexivity. + simpl. apply lt_rev_append. exact I. + intros x y Hx Hy. inversion_clear Hy. + rewrite H. simpl. apply lt_rev_append. exact I. + rewrite xelements_spec in H. destruct H as [Hy|[z [-> Hy]]]. + apply Hsacc; assumption. + simpl. apply lt_rev_append. exact I. + + apply IHl. apply IHr. apply Hacc. + intros x y Hx Hy. apply Hsacc; assumption. + intros x y Hx Hy. rewrite xelements_spec in Hy. + destruct Hy as [Hy|[z [-> Hy]]]. + apply Hsacc; assumption. + simpl. apply lt_rev_append. exact I. + + intros. apply H. constructor. + intros x y _ H'. inversion H'. + Qed. + + Lemma elements_spec2w: forall s, NoDupA E.eq (elements s). + Proof. + intro. apply SortA_NoDupA with E.lt; auto with *. + apply E.eq_equiv. + apply elements_spec2. + Qed. + + + (** Specification of [choose] *) + + Lemma choose_spec1: forall s x, choose s = Some x -> In x s. + Proof. + induction s as [| l IHl o r IHr]; simpl. + intros. discriminate. + destruct o. + intros x H. injection H; intros; subst. reflexivity. + revert IHl. case choose. + intros p Hp x H. injection H; intros; subst; clear H. apply Hp. + reflexivity. + intros _ x. revert IHr. case choose. + intros p Hp H. injection H; intros; subst; clear H. apply Hp. + reflexivity. + intros. discriminate. + Qed. + + Lemma choose_spec2: forall s, choose s = None -> Empty s. + Proof. + unfold Empty, In. intros s H. + induction s as [|l IHl o r IHr]. + intro. apply empty_spec. + destruct o. + discriminate. + simpl in H. destruct (choose l). + discriminate. + destruct (choose r). + discriminate. + intros [a|a|]. + apply IHr. reflexivity. + apply IHl. reflexivity. + discriminate. + Qed. + + Lemma choose_empty: forall s, is_empty s = true -> choose s = None. + Proof. + intros s Hs. case_eq (choose s); trivial. + intros p Hp. apply choose_spec1 in Hp. apply is_empty_spec in Hs. + elim (Hs _ Hp). + Qed. + + Lemma choose_spec3': forall s s', Equal s s' -> choose s = choose s'. + Proof. + setoid_rewrite <- equal_spec. + induction s as [|l IHl o r IHr]. + intros. symmetry. apply choose_empty. assumption. + + destruct s' as [|l' o' r']. + generalize (Node l o r) as s. simpl. intros. apply choose_empty. + rewrite equal_spec in H. symmetry in H. rewrite <- equal_spec in H. + assumption. + + simpl. rewrite <- 2andb_lazy_alt, 2andb_true_iff, eqb_true_iff. + intros [[<- Hl] Hr]. rewrite (IHl _ Hl), (IHr _ Hr). reflexivity. + Qed. + + Lemma choose_spec3: forall s s' x y, + choose s = Some x -> choose s' = Some y -> Equal s s' -> E.eq x y. + Proof. intros s s' x y Hx Hy H. apply choose_spec3' in H. congruence. Qed. + + + (** Specification of [min_elt] *) + + Lemma min_elt_spec1: forall s x, min_elt s = Some x -> In x s. + Proof. + unfold In. + induction s as [| l IHl o r IHr]; simpl. + intros. discriminate. + intros x. destruct (min_elt l); intros. + injection H. intros <-. apply IHl. reflexivity. + destruct o; simpl. + injection H. intros <-. reflexivity. + destruct (min_elt r); simpl in *. + injection H. intros <-. apply IHr. reflexivity. + discriminate. + Qed. + + Lemma min_elt_spec3: forall s, min_elt s = None -> Empty s. + Proof. + unfold Empty, In. intros s H. + induction s as [|l IHl o r IHr]. + intro. apply empty_spec. + intros [a|a|]. + apply IHr. revert H. clear. simpl. destruct (min_elt r); trivial. + case min_elt; intros; try discriminate. destruct o; discriminate. + apply IHl. revert H. clear. simpl. destruct (min_elt l); trivial. + intro; discriminate. + revert H. clear. simpl. case min_elt; intros; try discriminate. + destruct o; discriminate. + Qed. + + Lemma min_elt_spec2: forall s x y, min_elt s = Some x -> In y s -> ~ E.lt y x. + Proof. + unfold In. + induction s as [|l IHl o r IHr]; intros x y H H'. + discriminate. + simpl in H. case_eq (min_elt l). + intros p Hp. rewrite Hp in H. injection H; intros <-. + destruct y as [z|z|]; simpl; intro; trivial. apply (IHl p z); trivial. + intro Hp; rewrite Hp in H. apply min_elt_spec3 in Hp. + destruct o. + injection H. intros <- Hl. clear H. + destruct y as [z|z|]; simpl; trivial. elim (Hp _ H'). + + destruct (min_elt r). + injection H. intros <-. clear H. + destruct y as [z|z|]. + apply (IHr p z); trivial. + elim (Hp _ H'). + discriminate. + discriminate. + Qed. + + + (** Specification of [max_elt] *) + + Lemma max_elt_spec1: forall s x, max_elt s = Some x -> In x s. + Proof. + unfold In. + induction s as [| l IHl o r IHr]; simpl. + intros. discriminate. + intros x. destruct (max_elt r); intros. + injection H. intros <-. apply IHr. reflexivity. + destruct o; simpl. + injection H. intros <-. reflexivity. + destruct (max_elt l); simpl in *. + injection H. intros <-. apply IHl. reflexivity. + discriminate. + Qed. + + Lemma max_elt_spec3: forall s, max_elt s = None -> Empty s. + Proof. + unfold Empty, In. intros s H. + induction s as [|l IHl o r IHr]. + intro. apply empty_spec. + intros [a|a|]. + apply IHr. revert H. clear. simpl. destruct (max_elt r); trivial. + intro; discriminate. + apply IHl. revert H. clear. simpl. destruct (max_elt l); trivial. + case max_elt; intros; try discriminate. destruct o; discriminate. + revert H. clear. simpl. case max_elt; intros; try discriminate. + destruct o; discriminate. + Qed. + + Lemma max_elt_spec2: forall s x y, max_elt s = Some x -> In y s -> ~ E.lt x y. + Proof. + unfold In. + induction s as [|l IHl o r IHr]; intros x y H H'. + discriminate. + simpl in H. case_eq (max_elt r). + intros p Hp. rewrite Hp in H. injection H; intros <-. + destruct y as [z|z|]; simpl; intro; trivial. apply (IHr p z); trivial. + intro Hp; rewrite Hp in H. apply max_elt_spec3 in Hp. + destruct o. + injection H. intros <- Hl. clear H. + destruct y as [z|z|]; simpl; trivial. elim (Hp _ H'). + + destruct (max_elt l). + injection H. intros <-. clear H. + destruct y as [z|z|]. + elim (Hp _ H'). + apply (IHl p z); trivial. + discriminate. + discriminate. + Qed. + +End PositiveSet. diff --git a/theories/MSets/MSetProperties.v b/theories/MSets/MSetProperties.v new file mode 100644 index 00000000..c0038a4f --- /dev/null +++ b/theories/MSets/MSetProperties.v @@ -0,0 +1,1176 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) +(* \VV/ *************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(***********************************************************************) + +(* $Id$ *) + +(** * Finite sets library *) + +(** This functor derives additional properties from [MSetInterface.S]. + Contrary to the functor in [MSetEqProperties] it uses + predicates over sets instead of sets operations, i.e. + [In x s] instead of [mem x s=true], + [Equal s s'] instead of [equal s s'=true], etc. *) + +Require Export MSetInterface. +Require Import DecidableTypeEx OrdersLists MSetFacts MSetDecide. +Set Implicit Arguments. +Unset Strict Implicit. + +Hint Unfold transpose. + +(** First, a functor for Weak Sets in functorial version. *) + +Module WPropertiesOn (Import E : DecidableType)(M : WSetsOn E). + Module Import Dec := WDecideOn E M. + Module Import FM := Dec.F (* MSetFacts.WFactsOn E M *). + Import M. + + Lemma In_dec : forall x s, {In x s} + {~ In x s}. + Proof. + intros; generalize (mem_iff s x); case (mem x s); intuition. + Qed. + + Definition Add x s s' := forall y, In y s' <-> E.eq x y \/ In y s. + + Lemma Add_Equal : forall x s s', Add x s s' <-> s' [=] add x s. + Proof. + unfold Add. + split; intros. + red; intros. + rewrite H; clear H. + fsetdec. + fsetdec. + Qed. + + Ltac expAdd := repeat rewrite Add_Equal. + + Section BasicProperties. + + Variable s s' s'' s1 s2 s3 : t. + Variable x x' : elt. + + Lemma equal_refl : s[=]s. + Proof. fsetdec. Qed. + + Lemma equal_sym : s[=]s' -> s'[=]s. + Proof. fsetdec. Qed. + + Lemma equal_trans : s1[=]s2 -> s2[=]s3 -> s1[=]s3. + Proof. fsetdec. Qed. + + Lemma subset_refl : s[<=]s. + Proof. fsetdec. Qed. + + Lemma subset_trans : s1[<=]s2 -> s2[<=]s3 -> s1[<=]s3. + Proof. fsetdec. Qed. + + Lemma subset_antisym : s[<=]s' -> s'[<=]s -> s[=]s'. + Proof. fsetdec. Qed. + + Lemma subset_equal : s[=]s' -> s[<=]s'. + Proof. fsetdec. Qed. + + Lemma subset_empty : empty[<=]s. + Proof. fsetdec. Qed. + + Lemma subset_remove_3 : s1[<=]s2 -> remove x s1 [<=] s2. + Proof. fsetdec. Qed. + + Lemma subset_diff : s1[<=]s3 -> diff s1 s2 [<=] s3. + Proof. fsetdec. Qed. + + Lemma subset_add_3 : In x s2 -> s1[<=]s2 -> add x s1 [<=] s2. + Proof. fsetdec. Qed. + + Lemma subset_add_2 : s1[<=]s2 -> s1[<=] add x s2. + Proof. fsetdec. Qed. + + Lemma in_subset : In x s1 -> s1[<=]s2 -> In x s2. + Proof. fsetdec. Qed. + + Lemma double_inclusion : s1[=]s2 <-> s1[<=]s2 /\ s2[<=]s1. + Proof. intuition fsetdec. Qed. + + Lemma empty_is_empty_1 : Empty s -> s[=]empty. + Proof. fsetdec. Qed. + + Lemma empty_is_empty_2 : s[=]empty -> Empty s. + Proof. fsetdec. Qed. + + Lemma add_equal : In x s -> add x s [=] s. + Proof. fsetdec. Qed. + + Lemma add_add : add x (add x' s) [=] add x' (add x s). + Proof. fsetdec. Qed. + + Lemma remove_equal : ~ In x s -> remove x s [=] s. + Proof. fsetdec. Qed. + + Lemma Equal_remove : s[=]s' -> remove x s [=] remove x s'. + Proof. fsetdec. Qed. + + Lemma add_remove : In x s -> add x (remove x s) [=] s. + Proof. fsetdec. Qed. + + Lemma remove_add : ~In x s -> remove x (add x s) [=] s. + Proof. fsetdec. Qed. + + Lemma singleton_equal_add : singleton x [=] add x empty. + Proof. fsetdec. Qed. + + Lemma remove_singleton_empty : + In x s -> remove x s [=] empty -> singleton x [=] s. + Proof. fsetdec. Qed. + + Lemma union_sym : union s s' [=] union s' s. + Proof. fsetdec. Qed. + + Lemma union_subset_equal : s[<=]s' -> union s s' [=] s'. + Proof. fsetdec. Qed. + + Lemma union_equal_1 : s[=]s' -> union s s'' [=] union s' s''. + Proof. fsetdec. Qed. + + Lemma union_equal_2 : s'[=]s'' -> union s s' [=] union s s''. + Proof. fsetdec. Qed. + + Lemma union_assoc : union (union s s') s'' [=] union s (union s' s''). + Proof. fsetdec. Qed. + + Lemma add_union_singleton : add x s [=] union (singleton x) s. + Proof. fsetdec. Qed. + + Lemma union_add : union (add x s) s' [=] add x (union s s'). + Proof. fsetdec. Qed. + + Lemma union_remove_add_1 : + union (remove x s) (add x s') [=] union (add x s) (remove x s'). + Proof. fsetdec. Qed. + + Lemma union_remove_add_2 : In x s -> + union (remove x s) (add x s') [=] union s s'. + Proof. fsetdec. Qed. + + Lemma union_subset_1 : s [<=] union s s'. + Proof. fsetdec. Qed. + + Lemma union_subset_2 : s' [<=] union s s'. + Proof. fsetdec. Qed. + + Lemma union_subset_3 : s[<=]s'' -> s'[<=]s'' -> union s s' [<=] s''. + Proof. fsetdec. Qed. + + Lemma union_subset_4 : s[<=]s' -> union s s'' [<=] union s' s''. + Proof. fsetdec. Qed. + + Lemma union_subset_5 : s[<=]s' -> union s'' s [<=] union s'' s'. + Proof. fsetdec. Qed. + + Lemma empty_union_1 : Empty s -> union s s' [=] s'. + Proof. fsetdec. Qed. + + Lemma empty_union_2 : Empty s -> union s' s [=] s'. + Proof. fsetdec. Qed. + + Lemma not_in_union : ~In x s -> ~In x s' -> ~In x (union s s'). + Proof. fsetdec. Qed. + + Lemma inter_sym : inter s s' [=] inter s' s. + Proof. fsetdec. Qed. + + Lemma inter_subset_equal : s[<=]s' -> inter s s' [=] s. + Proof. fsetdec. Qed. + + Lemma inter_equal_1 : s[=]s' -> inter s s'' [=] inter s' s''. + Proof. fsetdec. Qed. + + Lemma inter_equal_2 : s'[=]s'' -> inter s s' [=] inter s s''. + Proof. fsetdec. Qed. + + Lemma inter_assoc : inter (inter s s') s'' [=] inter s (inter s' s''). + Proof. fsetdec. Qed. + + Lemma union_inter_1 : inter (union s s') s'' [=] union (inter s s'') (inter s' s''). + Proof. fsetdec. Qed. + + Lemma union_inter_2 : union (inter s s') s'' [=] inter (union s s'') (union s' s''). + Proof. fsetdec. Qed. + + Lemma inter_add_1 : In x s' -> inter (add x s) s' [=] add x (inter s s'). + Proof. fsetdec. Qed. + + Lemma inter_add_2 : ~ In x s' -> inter (add x s) s' [=] inter s s'. + Proof. fsetdec. Qed. + + Lemma empty_inter_1 : Empty s -> Empty (inter s s'). + Proof. fsetdec. Qed. + + Lemma empty_inter_2 : Empty s' -> Empty (inter s s'). + Proof. fsetdec. Qed. + + Lemma inter_subset_1 : inter s s' [<=] s. + Proof. fsetdec. Qed. + + Lemma inter_subset_2 : inter s s' [<=] s'. + Proof. fsetdec. Qed. + + Lemma inter_subset_3 : + s''[<=]s -> s''[<=]s' -> s''[<=] inter s s'. + Proof. fsetdec. Qed. + + Lemma empty_diff_1 : Empty s -> Empty (diff s s'). + Proof. fsetdec. Qed. + + Lemma empty_diff_2 : Empty s -> diff s' s [=] s'. + Proof. fsetdec. Qed. + + Lemma diff_subset : diff s s' [<=] s. + Proof. fsetdec. Qed. + + Lemma diff_subset_equal : s[<=]s' -> diff s s' [=] empty. + Proof. fsetdec. Qed. + + Lemma remove_diff_singleton : + remove x s [=] diff s (singleton x). + Proof. fsetdec. Qed. + + Lemma diff_inter_empty : inter (diff s s') (inter s s') [=] empty. + Proof. fsetdec. Qed. + + Lemma diff_inter_all : union (diff s s') (inter s s') [=] s. + Proof. fsetdec. Qed. + + Lemma Add_add : Add x s (add x s). + Proof. expAdd; fsetdec. Qed. + + Lemma Add_remove : In x s -> Add x (remove x s) s. + Proof. expAdd; fsetdec. Qed. + + Lemma union_Add : Add x s s' -> Add x (union s s'') (union s' s''). + Proof. expAdd; fsetdec. Qed. + + Lemma inter_Add : + In x s'' -> Add x s s' -> Add x (inter s s'') (inter s' s''). + Proof. expAdd; fsetdec. Qed. + + Lemma union_Equal : + In x s'' -> Add x s s' -> union s s'' [=] union s' s''. + Proof. expAdd; fsetdec. Qed. + + Lemma inter_Add_2 : + ~In x s'' -> Add x s s' -> inter s s'' [=] inter s' s''. + Proof. expAdd; fsetdec. Qed. + + End BasicProperties. + + Hint Immediate equal_sym add_remove remove_add union_sym inter_sym: set. + Hint Resolve equal_refl equal_trans subset_refl subset_equal subset_antisym + subset_trans subset_empty subset_remove_3 subset_diff subset_add_3 + subset_add_2 in_subset empty_is_empty_1 empty_is_empty_2 add_equal + remove_equal singleton_equal_add union_subset_equal union_equal_1 + union_equal_2 union_assoc add_union_singleton union_add union_subset_1 + union_subset_2 union_subset_3 inter_subset_equal inter_equal_1 inter_equal_2 + inter_assoc union_inter_1 union_inter_2 inter_add_1 inter_add_2 + empty_inter_1 empty_inter_2 empty_union_1 empty_union_2 empty_diff_1 + empty_diff_2 union_Add inter_Add union_Equal inter_Add_2 not_in_union + inter_subset_1 inter_subset_2 inter_subset_3 diff_subset diff_subset_equal + remove_diff_singleton diff_inter_empty diff_inter_all Add_add Add_remove + Equal_remove add_add : set. + + (** * Properties of elements *) + + Lemma elements_Empty : forall s, Empty s <-> elements s = nil. + Proof. + intros. + unfold Empty. + split; intros. + assert (forall a, ~ List.In a (elements s)). + red; intros. + apply (H a). + rewrite elements_iff. + rewrite InA_alt; exists a; auto with relations. + destruct (elements s); auto. + elim (H0 e); simpl; auto. + red; intros. + rewrite elements_iff in H0. + rewrite InA_alt in H0; destruct H0. + rewrite H in H0; destruct H0 as (_,H0); inversion H0. + Qed. + + Lemma elements_empty : elements empty = nil. + Proof. + rewrite <-elements_Empty; auto with set. + Qed. + + (** * Conversions between lists and sets *) + + Definition of_list (l : list elt) := List.fold_right add empty l. + + Definition to_list := elements. + + Lemma of_list_1 : forall l x, In x (of_list l) <-> InA E.eq x l. + Proof. + induction l; simpl; intro x. + rewrite empty_iff, InA_nil. intuition. + rewrite add_iff, InA_cons, IHl. intuition. + Qed. + + Lemma of_list_2 : forall l, equivlistA E.eq (to_list (of_list l)) l. + Proof. + unfold to_list; red; intros. + rewrite <- elements_iff; apply of_list_1. + Qed. + + Lemma of_list_3 : forall s, of_list (to_list s) [=] s. + Proof. + unfold to_list; red; intros. + rewrite of_list_1; symmetry; apply elements_iff. + Qed. + + (** * Fold *) + + Section Fold. + + Notation NoDup := (NoDupA E.eq). + Notation InA := (InA E.eq). + + (** ** Induction principles for fold (contributed by S. Lescuyer) *) + + (** In the following lemma, the step hypothesis is deliberately restricted + to the precise set s we are considering. *) + + Theorem fold_rec : + forall (A:Type)(P : t -> A -> Type)(f : elt -> A -> A)(i:A)(s:t), + (forall s', Empty s' -> P s' i) -> + (forall x a s' s'', In x s -> ~In x s' -> Add x s' s'' -> + P s' a -> P s'' (f x a)) -> + 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)). + 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. + rewrite elements_iff, <- InA_rev; auto with *. + assert (Hdup : NoDup l) by + (unfold l; eauto using elements_3w, NoDupA_rev with *). + assert (Hsame : forall x, In x s <-> InA x l) by + (unfold l; intros; rewrite elements_iff, InA_rev; intuition). + clear Pstep; clearbody l; revert s Hsame; induction l. + (* empty *) + intros s Hsame; simpl. + apply Pempty. intro x. rewrite Hsame, InA_nil; intuition. + (* step *) + intros s Hsame; simpl. + apply Pstep' with (of_list l); auto with relations. + inversion_clear Hdup; rewrite of_list_1; auto. + red. intros. rewrite Hsame, of_list_1, InA_cons; intuition. + apply IHl. + intros; eapply Pstep'; eauto. + inversion_clear Hdup; auto. + exact (of_list_1 l). + Qed. + + (** Same, with [empty] and [add] instead of [Empty] and [Add]. In this + case, [P] must be compatible with equality of sets *) + + Theorem fold_rec_bis : + forall (A:Type)(P : t -> A -> Type)(f : elt -> A -> A)(i:A)(s:t), + (forall s s' a, s[=]s' -> P s a -> P s' a) -> + (P empty i) -> + (forall x a s', In x s -> ~In x s' -> P s' a -> P (add x s') (f x a)) -> + P s (fold f s i). + Proof. + intros A P f i s Pmorphism Pempty Pstep. + apply fold_rec; intros. + apply Pmorphism with empty; auto with set. + rewrite Add_Equal in H1; auto with set. + apply Pmorphism with (add x s'); auto with set. + Qed. + + Lemma fold_rec_nodep : + forall (A:Type)(P : A -> Type)(f : elt -> A -> A)(i:A)(s:t), + P i -> (forall x a, In x s -> P a -> P (f x a)) -> + P (fold f s i). + Proof. + intros; apply fold_rec_bis with (P:=fun _ => P); auto. + Qed. + + (** [fold_rec_weak] is a weaker principle than [fold_rec_bis] : + the step hypothesis must here be applicable to any [x]. + At the same time, it looks more like an induction principle, + and hence can be easier to use. *) + + Lemma fold_rec_weak : + forall (A:Type)(P : t -> A -> Type)(f : elt -> A -> A)(i:A), + (forall s s' a, s[=]s' -> P s a -> P s' a) -> + P empty i -> + (forall x a s, ~In x s -> P s a -> P (add x s) (f x a)) -> + forall s, P s (fold f s i). + Proof. + intros; apply fold_rec_bis; auto. + Qed. + + Lemma fold_rel : + forall (A B:Type)(R : A -> B -> Type) + (f : elt -> A -> A)(g : elt -> B -> B)(i : A)(j : B)(s : t), + R i j -> + (forall x a b, In x s -> R a b -> R (f x a) (g x b)) -> + 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)). + 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. + induction l; simpl; auto with relations. + Qed. + + (** From the induction principle on [fold], we can deduce some general + induction principles on sets. *) + + Lemma set_induction : + forall P : t -> Type, + (forall s, Empty s -> P s) -> + (forall s s', P s -> forall x, ~In x s -> Add x s s' -> P s') -> + forall s, P s. + Proof. + intros. apply (@fold_rec _ (fun s _ => P s) (fun _ _ => tt) tt s); eauto. + Qed. + + Lemma set_induction_bis : + forall P : t -> Type, + (forall s s', s [=] s' -> P s -> P s') -> + P empty -> + (forall x s, ~In x s -> P s -> P (add x s)) -> + forall s, P s. + Proof. + intros. + apply (@fold_rec_bis _ (fun s _ => P s) (fun _ _ => tt) tt s); eauto. + Qed. + + (** [fold] can be used to reconstruct the same initial set. *) + + Lemma fold_identity : forall s, fold add s empty [=] s. + Proof. + intros. + apply fold_rec with (P:=fun s acc => acc[=]s); auto with set. + intros. rewrite H2; rewrite Add_Equal in H1; auto with set. + Qed. + + (** ** Alternative (weaker) specifications for [fold] *) + + (** When [MSets] was first designed, the order in which Ocaml's [Set.fold] + takes the set elements was unspecified. This specification reflects + this fact: + *) + + Lemma fold_0 : + forall s (A : Type) (i : A) (f : elt -> A -> A), + exists l : list elt, + NoDup l /\ + (forall x : elt, In x s <-> InA x l) /\ + fold f s i = fold_right f i l. + Proof. + intros; exists (rev (elements s)); split. + apply NoDupA_rev; auto with *. + 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. + Qed. + + (** An alternate (and previous) specification for [fold] was based on + the recursive structure of a set. It is now lemmas [fold_1] and + [fold_2]. *) + + Lemma fold_1 : + forall s (A : Type) (eqA : A -> A -> Prop) + (st : Equivalence eqA) (i : A) (f : elt -> A -> A), + Empty s -> eqA (fold f s i) i. + Proof. + unfold Empty; intros; destruct (fold_0 s i f) as (l,(H1, (H2, H3))). + rewrite H3; clear H3. + generalize H H2; clear H H2; case l; simpl; intros. + reflexivity. + elim (H e). + elim (H2 e); intuition. + Qed. + + Lemma fold_2 : + forall s s' x (A : Type) (eqA : A -> A -> Prop) + (st : Equivalence eqA) (i : A) (f : elt -> A -> A), + Proper (E.eq==>eqA==>eqA) f -> + transpose eqA f -> + ~ In x s -> Add x s s' -> eqA (fold f s' i) (f x (fold f s i)). + Proof. + intros; destruct (fold_0 s i f) as (l,(Hl, (Hl1, Hl2))); + destruct (fold_0 s' i f) as (l',(Hl', (Hl'1, Hl'2))). + rewrite Hl2; rewrite Hl'2; clear Hl2 Hl'2. + apply fold_right_add with (eqA:=E.eq)(eqB:=eqA); auto. + eauto with *. + rewrite <- Hl1; auto. + intros a; rewrite InA_cons; rewrite <- Hl1; rewrite <- Hl'1; + rewrite (H2 a); intuition. + Qed. + + (** In fact, [fold] on empty sets is more than equivalent to + the initial element, it is Leibniz-equal to it. *) + + Lemma fold_1b : + forall s (A : Type)(i : A) (f : elt -> A -> A), + Empty s -> (fold f s i) = i. + Proof. + intros. + rewrite FM.fold_1. + rewrite elements_Empty in H; rewrite H; simpl; auto. + Qed. + + Section Fold_More. + + Variables (A:Type)(eqA:A->A->Prop)(st:Equivalence eqA). + Variables (f:elt->A->A)(Comp:Proper (E.eq==>eqA==>eqA) f)(Ass:transpose eqA f). + + Lemma fold_commutes : forall i s x, + eqA (fold f s (f x i)) (f x (fold f s i)). + Proof. + intros. + apply fold_rel with (R:=fun u v => eqA u (f x v)); intros. + reflexivity. + transitivity (f x0 (f x b)); auto. + apply Comp; auto with relations. + Qed. + + (** ** Fold is a morphism *) + + Lemma fold_init : forall i i' s, eqA i i' -> + eqA (fold f s i) (fold f s i'). + Proof. + intros. apply fold_rel with (R:=eqA); auto. + intros; apply Comp; auto with relations. + Qed. + + Lemma fold_equal : + forall i s s', s[=]s' -> eqA (fold f s i) (fold f s' i). + Proof. + intros i s; pattern s; apply set_induction; clear s; intros. + transitivity i. + apply fold_1; auto. + symmetry; apply fold_1; auto. + rewrite <- H0; auto. + transitivity (f x (fold f s i)). + apply fold_2 with (eqA := eqA); auto. + symmetry; apply fold_2 with (eqA := eqA); auto. + unfold Add in *; intros. + rewrite <- H2; auto. + Qed. + + (** ** Fold and other set operators *) + + Lemma fold_empty : forall i, fold f empty i = i. + Proof. + intros i; apply fold_1b; auto with set. + Qed. + + Lemma fold_add : forall i s x, ~In x s -> + eqA (fold f (add x s) i) (f x (fold f s i)). + Proof. + intros; apply fold_2 with (eqA := eqA); auto with set. + Qed. + + Lemma add_fold : forall i s x, In x s -> + eqA (fold f (add x s) i) (fold f s i). + Proof. + intros; apply fold_equal; auto with set. + Qed. + + Lemma remove_fold_1: forall i s x, In x s -> + eqA (f x (fold f (remove x s) i)) (fold f s i). + Proof. + intros. + symmetry. + apply fold_2 with (eqA:=eqA); auto with set relations. + Qed. + + Lemma remove_fold_2: forall i s x, ~In x s -> + eqA (fold f (remove x s) i) (fold f s i). + Proof. + intros. + apply fold_equal; auto with set. + Qed. + + Lemma fold_union_inter : forall i s s', + eqA (fold f (union s s') (fold f (inter s s') i)) + (fold f s (fold f s' i)). + Proof. + intros; pattern s; apply set_induction; clear s; intros. + transitivity (fold f s' (fold f (inter s s') i)). + apply fold_equal; auto with set. + transitivity (fold f s' i). + apply fold_init; auto. + apply fold_1; auto with set. + symmetry; apply fold_1; auto. + rename s'0 into s''. + destruct (In_dec x s'). + (* In x s' *) + transitivity (fold f (union s'' s') (f x (fold f (inter s s') i))); auto with set. + apply fold_init; auto. + apply fold_2 with (eqA:=eqA); auto with set. + rewrite inter_iff; intuition. + transitivity (f x (fold f s (fold f s' i))). + transitivity (fold f (union s s') (f x (fold f (inter s s') i))). + apply fold_equal; auto. + apply equal_sym; apply union_Equal with x; auto with set. + transitivity (f x (fold f (union s s') (fold f (inter s s') i))). + apply fold_commutes; auto. + apply Comp; auto with relations. + symmetry; apply fold_2 with (eqA:=eqA); auto. + (* ~(In x s') *) + transitivity (f x (fold f (union s s') (fold f (inter s'' s') i))). + apply fold_2 with (eqA:=eqA); auto with set. + transitivity (f x (fold f (union s s') (fold f (inter s s') i))). + apply Comp;auto with relations. + apply fold_init;auto. + apply fold_equal;auto. + apply equal_sym; apply inter_Add_2 with x; auto with set. + transitivity (f x (fold f s (fold f s' i))). + apply Comp; auto with relations. + symmetry; apply fold_2 with (eqA:=eqA); auto. + Qed. + + Lemma fold_diff_inter : forall i s s', + eqA (fold f (diff s s') (fold f (inter s s') i)) (fold f s i). + Proof. + intros. + transitivity (fold f (union (diff s s') (inter s s')) + (fold f (inter (diff s s') (inter s s')) i)). + symmetry; apply fold_union_inter; auto. + transitivity (fold f s (fold f (inter (diff s s') (inter s s')) i)). + apply fold_equal; auto with set. + apply fold_init; auto. + apply fold_1; auto with set. + Qed. + + Lemma fold_union: forall i s s', + (forall x, ~(In x s/\In x s')) -> + eqA (fold f (union s s') i) (fold f s (fold f s' i)). + Proof. + intros. + transitivity (fold f (union s s') (fold f (inter s s') i)). + apply fold_init; auto. + symmetry; apply fold_1; auto with set. + unfold Empty; intro a; generalize (H a); set_iff; tauto. + apply fold_union_inter; auto. + Qed. + + End Fold_More. + + Lemma fold_plus : + forall s p, fold (fun _ => S) s p = fold (fun _ => S) s 0 + p. + Proof. + intros. apply fold_rel with (R:=fun u v => u = v + p); simpl; auto. + Qed. + + End Fold. + + (** * Cardinal *) + + (** ** Characterization of cardinal in terms of fold *) + + Lemma cardinal_fold : forall s, cardinal s = fold (fun _ => S) s 0. + Proof. + intros; rewrite cardinal_1; rewrite FM.fold_1. + symmetry; apply fold_left_length; auto. + Qed. + + (** ** Old specifications for [cardinal]. *) + + Lemma cardinal_0 : + forall s, exists l : list elt, + NoDupA E.eq l /\ + (forall x : elt, In x s <-> InA E.eq x l) /\ + cardinal s = length l. + Proof. + intros; exists (elements s); intuition; apply cardinal_1. + Qed. + + Lemma cardinal_1 : forall s, Empty s -> cardinal s = 0. + Proof. + intros; rewrite cardinal_fold; apply fold_1; auto with *. + Qed. + + Lemma cardinal_2 : + forall s s' x, ~ In x s -> Add x s s' -> cardinal s' = S (cardinal s). + Proof. + intros; do 2 rewrite cardinal_fold. + change S with ((fun _ => S) x). + apply fold_2; auto. + split; congruence. + congruence. + Qed. + + (** ** Cardinal and (non-)emptiness *) + + Lemma cardinal_Empty : forall s, Empty s <-> cardinal s = 0. + Proof. + intros. + rewrite elements_Empty, FM.cardinal_1. + destruct (elements s); intuition; discriminate. + Qed. + + Lemma cardinal_inv_1 : forall s, cardinal s = 0 -> Empty s. + Proof. + intros; rewrite cardinal_Empty; auto. + Qed. + Hint Resolve cardinal_inv_1. + + Lemma cardinal_inv_2 : + forall s n, cardinal s = S n -> { x : elt | In x s }. + Proof. + intros; rewrite FM.cardinal_1 in H. + generalize (elements_2 (s:=s)). + destruct (elements s); try discriminate. + exists e; auto with relations. + Qed. + + Lemma cardinal_inv_2b : + forall s, cardinal s <> 0 -> { x : elt | In x s }. + Proof. + intro; generalize (@cardinal_inv_2 s); destruct cardinal; + [intuition|eauto]. + Qed. + + (** ** Cardinal is a morphism *) + + Lemma Equal_cardinal : forall s s', s[=]s' -> cardinal s = cardinal s'. + Proof. + symmetry. + remember (cardinal s) as n; symmetry in Heqn; revert s s' Heqn H. + induction n; intros. + apply cardinal_1; rewrite <- H; auto. + destruct (cardinal_inv_2 Heqn) as (x,H2). + revert Heqn. + rewrite (cardinal_2 (s:=remove x s) (s':=s) (x:=x)); + auto with set relations. + rewrite (cardinal_2 (s:=remove x s') (s':=s') (x:=x)); + eauto with set relations. + Qed. + + Instance cardinal_m : Proper (Equal==>Logic.eq) cardinal. + Proof. + exact Equal_cardinal. + Qed. + + Hint Resolve Add_add Add_remove Equal_remove cardinal_inv_1 Equal_cardinal. + + (** ** Cardinal and set operators *) + + Lemma empty_cardinal : cardinal empty = 0. + Proof. + rewrite cardinal_fold; apply fold_1; auto with *. + Qed. + + Hint Immediate empty_cardinal cardinal_1 : set. + + Lemma singleton_cardinal : forall x, cardinal (singleton x) = 1. + Proof. + intros. + rewrite (singleton_equal_add x). + replace 0 with (cardinal empty); auto with set. + apply cardinal_2 with x; auto with set. + Qed. + + Hint Resolve singleton_cardinal: set. + + Lemma diff_inter_cardinal : + forall s s', cardinal (diff s s') + cardinal (inter s s') = cardinal s . + Proof. + intros; do 3 rewrite cardinal_fold. + rewrite <- fold_plus. + apply fold_diff_inter with (eqA:=@Logic.eq nat); auto with *. + congruence. + Qed. + + Lemma union_cardinal: + forall s s', (forall x, ~(In x s/\In x s')) -> + cardinal (union s s')=cardinal s+cardinal s'. + Proof. + intros; do 3 rewrite cardinal_fold. + rewrite <- fold_plus. + apply fold_union; auto. + split; congruence. + congruence. + Qed. + + Lemma subset_cardinal : + forall s s', s[<=]s' -> cardinal s <= cardinal s' . + Proof. + intros. + rewrite <- (diff_inter_cardinal s' s). + rewrite (inter_sym s' s). + rewrite (inter_subset_equal H); auto with arith. + Qed. + + Lemma subset_cardinal_lt : + forall s s' x, s[<=]s' -> In x s' -> ~In x s -> cardinal s < cardinal s'. + Proof. + intros. + rewrite <- (diff_inter_cardinal s' s). + rewrite (inter_sym s' s). + rewrite (inter_subset_equal H). + generalize (@cardinal_inv_1 (diff s' s)). + destruct (cardinal (diff s' s)). + intro H2; destruct (H2 (refl_equal _) x). + set_iff; auto. + intros _. + change (0 + cardinal s < S n + cardinal s). + apply Plus.plus_lt_le_compat; auto with arith. + Qed. + + Theorem union_inter_cardinal : + forall s s', cardinal (union s s') + cardinal (inter s s') = cardinal s + cardinal s' . + Proof. + intros. + do 4 rewrite cardinal_fold. + do 2 rewrite <- fold_plus. + apply fold_union_inter with (eqA:=@Logic.eq nat); auto with *. + congruence. + Qed. + + Lemma union_cardinal_inter : + forall s s', cardinal (union s s') = cardinal s + cardinal s' - cardinal (inter s s'). + Proof. + intros. + rewrite <- union_inter_cardinal. + rewrite Plus.plus_comm. + auto with arith. + Qed. + + Lemma union_cardinal_le : + forall s s', cardinal (union s s') <= cardinal s + cardinal s'. + Proof. + intros; generalize (union_inter_cardinal s s'). + intros; rewrite <- H; auto with arith. + Qed. + + Lemma add_cardinal_1 : + forall s x, In x s -> cardinal (add x s) = cardinal s. + Proof. + auto with set. + Qed. + + Lemma add_cardinal_2 : + forall s x, ~In x s -> cardinal (add x s) = S (cardinal s). + Proof. + intros. + do 2 rewrite cardinal_fold. + change S with ((fun _ => S) x); + apply fold_add with (eqA:=@Logic.eq nat); auto with *. + congruence. + Qed. + + Lemma remove_cardinal_1 : + forall s x, In x s -> S (cardinal (remove x s)) = cardinal s. + Proof. + intros. + do 2 rewrite cardinal_fold. + change S with ((fun _ =>S) x). + apply remove_fold_1 with (eqA:=@Logic.eq nat); auto with *. + congruence. + Qed. + + Lemma remove_cardinal_2 : + forall s x, ~In x s -> cardinal (remove x s) = cardinal s. + Proof. + auto with set. + Qed. + + Hint Resolve subset_cardinal union_cardinal add_cardinal_1 add_cardinal_2. + +End WPropertiesOn. + +(** Now comes variants for self-contained weak sets and for full sets. + For these variants, only one argument is necessary. Thanks to + the subtyping [WS<=S], the [Properties] functor which is meant to be + used on modules [(M:S)] can simply be an alias of [WProperties]. *) + +Module WProperties (M:WSets) := WPropertiesOn M.E M. +Module Properties := WProperties. + + +(** Now comes some properties specific to the element ordering, + invalid for Weak Sets. *) + +Module OrdProperties (M:Sets). + Module Import ME:=OrderedTypeFacts(M.E). + Module Import ML:=OrderedTypeLists(M.E). + Module Import P := Properties M. + Import FM. + Import M.E. + Import M. + + Hint Resolve elements_spec2. + Hint Immediate + min_elt_spec1 min_elt_spec2 min_elt_spec3 + max_elt_spec1 max_elt_spec2 max_elt_spec3 : set. + + (** First, a specialized version of SortA_equivlistA_eqlistA: *) + Lemma sort_equivlistA_eqlistA : forall l l' : list elt, + sort E.lt l -> sort E.lt l' -> equivlistA E.eq l l' -> eqlistA E.eq l l'. + Proof. + apply SortA_equivlistA_eqlistA; eauto with *. + Qed. + + Definition gtb x y := match E.compare x y with Gt => true | _ => false end. + Definition leb x := fun y => negb (gtb x y). + + Definition elements_lt x s := List.filter (gtb x) (elements s). + Definition elements_ge x s := List.filter (leb x) (elements s). + + Lemma gtb_1 : forall x y, gtb x y = true <-> E.lt y x. + Proof. + intros; rewrite <- compare_gt_iff. unfold gtb. + destruct E.compare; intuition; try discriminate. + Qed. + + Lemma leb_1 : forall x y, leb x y = true <-> ~E.lt y x. + Proof. + intros; rewrite <- compare_gt_iff. unfold leb, gtb. + destruct E.compare; intuition; try discriminate. + Qed. + + Instance gtb_compat x : Proper (E.eq==>Logic.eq) (gtb x). + Proof. + intros a b H. unfold gtb. rewrite H; auto. + Qed. + + Instance leb_compat x : Proper (E.eq==>Logic.eq) (leb x). + Proof. + intros a b H; unfold leb. rewrite H; auto. + Qed. + Hint Resolve gtb_compat leb_compat. + + Lemma elements_split : forall x s, + elements s = elements_lt x s ++ elements_ge x s. + Proof. + unfold elements_lt, elements_ge, leb; intros. + eapply (@filter_split _ E.eq); eauto with *. + intros. + rewrite gtb_1 in H. + assert (~E.lt y x). + unfold gtb in *; elim_compare x y; intuition; + try discriminate; order. + order. + Qed. + + Lemma elements_Add : forall s s' x, ~In x s -> Add x s s' -> + eqlistA E.eq (elements s') (elements_lt x s ++ x :: elements_ge x s). + Proof. + intros; unfold elements_ge, elements_lt. + apply sort_equivlistA_eqlistA; auto with set. + apply (@SortA_app _ E.eq); auto with *. + apply (@filter_sort _ E.eq); auto with *; eauto with *. + constructor; auto. + apply (@filter_sort _ E.eq); auto with *; eauto with *. + rewrite Inf_alt by (apply (@filter_sort _ E.eq); eauto with *). + intros. + rewrite filter_InA in H1; auto with *; destruct H1. + rewrite leb_1 in H2. + rewrite <- elements_iff in H1. + assert (~E.eq x y). + contradict H; rewrite H; auto. + order. + intros. + rewrite filter_InA in H1; auto with *; destruct H1. + rewrite gtb_1 in H3. + inversion_clear H2. + order. + rewrite filter_InA in H4; auto with *; destruct H4. + rewrite leb_1 in H4. + order. + red; intros a. + rewrite InA_app_iff, InA_cons, !filter_InA, <-!elements_iff, + leb_1, gtb_1, (H0 a) by (auto with *). + intuition. + elim_compare a x; intuition. + right; right; split; auto. + order. + Qed. + + Definition Above x s := forall y, In y s -> E.lt y x. + Definition Below x s := forall y, In y s -> E.lt x y. + + Lemma elements_Add_Above : forall s s' x, + Above x s -> Add x s s' -> + eqlistA E.eq (elements s') (elements s ++ x::nil). + Proof. + intros. + apply sort_equivlistA_eqlistA; auto with set. + apply (@SortA_app _ E.eq); auto with *. + intros. + invlist InA. + rewrite <- elements_iff in H1. + setoid_replace y with x; auto. + red; intros a. + rewrite InA_app_iff, InA_cons, InA_nil, <-!elements_iff, (H0 a) + by (auto with *). + intuition. + Qed. + + Lemma elements_Add_Below : forall s s' x, + Below x s -> Add x s s' -> + eqlistA E.eq (elements s') (x::elements s). + Proof. + intros. + apply sort_equivlistA_eqlistA; auto with set. + change (sort E.lt ((x::nil) ++ elements s)). + apply (@SortA_app _ E.eq); auto with *. + intros. + invlist InA. + rewrite <- elements_iff in H2. + setoid_replace x0 with x; auto. + red; intros a. + rewrite InA_cons, <- !elements_iff, (H0 a); intuition. + Qed. + + (** Two other induction principles on sets: we can be more restrictive + on the element we add at each step. *) + + Lemma set_induction_max : + forall P : t -> Type, + (forall s : t, Empty s -> P s) -> + (forall s s', P s -> forall x, Above x s -> Add x s s' -> P s') -> + forall s : t, P s. + Proof. + intros; remember (cardinal s) as n; revert s Heqn; induction n; intros; auto. + case_eq (max_elt s); intros. + apply X0 with (remove e s) e; auto with set. + apply IHn. + assert (S n = S (cardinal (remove e s))). + rewrite Heqn; apply cardinal_2 with e; auto with set relations. + inversion H0; auto. + red; intros. + rewrite remove_iff in H0; destruct H0. + generalize (@max_elt_spec2 s e y H H0); order. + + assert (H0:=max_elt_spec3 H). + rewrite cardinal_Empty in H0; rewrite H0 in Heqn; inversion Heqn. + Qed. + + Lemma set_induction_min : + forall P : t -> Type, + (forall s : t, Empty s -> P s) -> + (forall s s', P s -> forall x, Below x s -> Add x s s' -> P s') -> + forall s : t, P s. + Proof. + intros; remember (cardinal s) as n; revert s Heqn; induction n; intros; auto. + case_eq (min_elt s); intros. + apply X0 with (remove e s) e; auto with set. + apply IHn. + assert (S n = S (cardinal (remove e s))). + rewrite Heqn; apply cardinal_2 with e; auto with set relations. + inversion H0; auto. + red; intros. + rewrite remove_iff in H0; destruct H0. + generalize (@min_elt_spec2 s e y H H0); order. + + assert (H0:=min_elt_spec3 H). + rewrite cardinal_Empty in H0; auto; rewrite H0 in Heqn; inversion Heqn. + Qed. + + (** More properties of [fold] : behavior with respect to Above/Below *) + + Lemma fold_3 : + forall s s' x (A : Type) (eqA : A -> A -> Prop) + (st : Equivalence eqA) (i : A) (f : elt -> A -> A), + Proper (E.eq==>eqA==>eqA) f -> + 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. + 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 *. + rewrite <- distr_rev. + apply eqlistA_rev. + apply elements_Add_Above; auto. + Qed. + + Lemma fold_4 : + forall s s' x (A : Type) (eqA : A -> A -> Prop) + (st : Equivalence eqA) (i : A) (f : elt -> A -> A), + Proper (E.eq==>eqA==>eqA) f -> + Below x s -> Add x s s' -> eqA (fold f s' i) (fold f s (f x i)). + Proof. + intros. + rewrite !FM.fold_1. + change (eqA (fold_left (flip f) (elements s') i) + (fold_left (flip f) (x::elements s) i)). + unfold flip; rewrite <-!fold_left_rev_right. + apply (@fold_right_eqlistA E.t E.eq A eqA st); auto. + apply eqlistA_rev. + apply elements_Add_Below; auto. + Qed. + + (** The following results have already been proved earlier, + but we can now prove them with one hypothesis less: + no need for [(transpose eqA f)]. *) + + Section FoldOpt. + Variables (A:Type)(eqA:A->A->Prop)(st:Equivalence eqA). + Variables (f:elt->A->A)(Comp:Proper (E.eq==>eqA==>eqA) f). + + Lemma fold_equal : + 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. + apply (@fold_right_eqlistA E.t E.eq A eqA st); auto. + apply eqlistA_rev. + apply sort_equivlistA_eqlistA; auto with set. + red; intro a; do 2 rewrite <- elements_iff; auto. + Qed. + + Lemma add_fold : forall i s x, In x s -> + eqA (fold f (add x s) i) (fold f s i). + Proof. + intros; apply fold_equal; auto with set. + Qed. + + Lemma remove_fold_2: forall i s x, ~In x s -> + eqA (fold f (remove x s) i) (fold f s i). + Proof. + intros. + apply fold_equal; auto with set. + Qed. + + End FoldOpt. + + (** An alternative version of [choose_3] *) + + Lemma choose_equal : forall s s', Equal s s' -> + match choose s, choose s' with + | Some x, Some x' => E.eq x x' + | None, None => True + | _, _ => False + end. + Proof. + intros s s' H; + generalize (@choose_spec1 s)(@choose_spec2 s) + (@choose_spec1 s')(@choose_spec2 s')(@choose_spec3 s s'); + destruct (choose s); destruct (choose s'); simpl; intuition. + apply H5 with e; rewrite <-H; auto. + apply H5 with e; rewrite H; auto. + Qed. + +End OrdProperties. diff --git a/theories/MSets/MSetToFiniteSet.v b/theories/MSets/MSetToFiniteSet.v new file mode 100644 index 00000000..f0b964cf --- /dev/null +++ b/theories/MSets/MSetToFiniteSet.v @@ -0,0 +1,158 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) +(* \VV/ *************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(***********************************************************************) + +(* $Id$ *) + +(** * Finite sets library : conversion to old [Finite_sets] *) + +Require Import Ensembles Finite_sets. +Require Import MSetInterface MSetProperties OrdersEx. + +(** * Going from [MSets] with usual Leibniz equality + to the good old [Ensembles] and [Finite_sets] theory. *) + +Module WS_to_Finite_set (U:UsualDecidableType)(M: WSetsOn U). + Module MP:= WPropertiesOn U M. + Import M MP FM Ensembles Finite_sets. + + Definition mkEns : M.t -> Ensemble M.elt := + fun s x => M.In x s. + + Notation " !! " := mkEns. + + Lemma In_In : forall s x, M.In x s <-> In _ (!!s) x. + Proof. + unfold In; compute; auto with extcore. + Qed. + + Lemma Subset_Included : forall s s', s[<=]s' <-> Included _ (!!s) (!!s'). + Proof. + unfold Subset, Included, In, mkEns; intuition. + Qed. + + Notation " a === b " := (Same_set M.elt a b) (at level 70, no associativity). + + Lemma Equal_Same_set : forall s s', s[=]s' <-> !!s === !!s'. + Proof. + intros. + rewrite double_inclusion. + unfold Subset, Included, Same_set, In, mkEns; intuition. + Qed. + + Lemma empty_Empty_Set : !!M.empty === Empty_set _. + Proof. + unfold Same_set, Included, mkEns, In. + split; intro; set_iff; inversion 1. + Qed. + + Lemma Empty_Empty_set : forall s, Empty s -> !!s === Empty_set _. + Proof. + unfold Same_set, Included, mkEns, In. + split; intros. + destruct(H x H0). + inversion H0. + Qed. + + Lemma singleton_Singleton : forall x, !!(M.singleton x) === Singleton _ x . + Proof. + unfold Same_set, Included, mkEns, In. + split; intro; set_iff; inversion 1; try constructor; auto. + Qed. + + Lemma union_Union : forall s s', !!(union s s') === Union _ (!!s) (!!s'). + Proof. + unfold Same_set, Included, mkEns, In. + split; intro; set_iff; inversion 1; [ constructor 1 | constructor 2 | | ]; auto. + Qed. + + Lemma inter_Intersection : forall s s', !!(inter s s') === Intersection _ (!!s) (!!s'). + Proof. + unfold Same_set, Included, mkEns, In. + split; intro; set_iff; inversion 1; try constructor; auto. + Qed. + + Lemma add_Add : forall x s, !!(add x s) === Add _ (!!s) x. + Proof. + unfold Same_set, Included, mkEns, In. + split; intro; set_iff; inversion 1; auto with sets. + inversion H0. + constructor 2; constructor. + constructor 1; auto. + Qed. + + Lemma Add_Add : forall x s s', MP.Add x s s' -> !!s' === Add _ (!!s) x. + Proof. + unfold Same_set, Included, mkEns, In. + split; intros. + red in H; rewrite H in H0. + destruct H0. + inversion H0. + constructor 2; constructor. + constructor 1; auto. + red in H; rewrite H. + inversion H0; auto. + inversion H1; auto. + Qed. + + Lemma remove_Subtract : forall x s, !!(remove x s) === Subtract _ (!!s) x. + Proof. + unfold Same_set, Included, mkEns, In. + split; intro; set_iff; inversion 1; auto with sets. + split; auto. + contradict H1. + inversion H1; auto. + Qed. + + Lemma mkEns_Finite : forall s, Finite _ (!!s). + Proof. + intro s; pattern s; apply set_induction; clear s; intros. + intros; replace (!!s) with (Empty_set elt); auto with sets. + symmetry; apply Extensionality_Ensembles. + apply Empty_Empty_set; auto. + replace (!!s') with (Add _ (!!s) x). + constructor 2; auto. + symmetry; apply Extensionality_Ensembles. + apply Add_Add; auto. + Qed. + + Lemma mkEns_cardinal : forall s, cardinal _ (!!s) (M.cardinal s). + Proof. + intro s; pattern s; apply set_induction; clear s; intros. + intros; replace (!!s) with (Empty_set elt); auto with sets. + rewrite MP.cardinal_1; auto with sets. + symmetry; apply Extensionality_Ensembles. + apply Empty_Empty_set; auto. + replace (!!s') with (Add _ (!!s) x). + rewrite (cardinal_2 H0 H1); auto with sets. + symmetry; apply Extensionality_Ensembles. + apply Add_Add; auto. + Qed. + + (** we can even build a function from Finite Ensemble to MSet + ... at least in Prop. *) + + Lemma Ens_to_MSet : forall e : Ensemble M.elt, Finite _ e -> + exists s:M.t, !!s === e. + Proof. + induction 1. + exists M.empty. + apply empty_Empty_Set. + destruct IHFinite as (s,Hs). + exists (M.add x s). + apply Extensionality_Ensembles in Hs. + rewrite <- Hs. + apply add_Add. + Qed. + +End WS_to_Finite_set. + + +Module S_to_Finite_set (U:UsualOrderedType)(M: SetsOn U) := + WS_to_Finite_set U M. + + diff --git a/theories/MSets/MSetWeakList.v b/theories/MSets/MSetWeakList.v new file mode 100644 index 00000000..945cb2dd --- /dev/null +++ b/theories/MSets/MSetWeakList.v @@ -0,0 +1,533 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) +(* \VV/ *************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(***********************************************************************) + +(* $Id$ *) + +(** * Finite sets library *) + +(** This file proposes an implementation of the non-dependant + interface [MSetWeakInterface.S] using lists without redundancy. *) + +Require Import MSetInterface. +Set Implicit Arguments. +Unset Strict Implicit. + +(** * Functions over lists + + First, we provide sets as lists which are (morally) without redundancy. + The specs are proved under the additional condition of no redundancy. + And the functions returning sets are proved to preserve this invariant. *) + + +(** ** The set operations. *) + +Module Ops (X: DecidableType) <: WOps X. + + Definition elt := X.t. + Definition t := list elt. + + Definition empty : t := nil. + + Definition is_empty (l : t) : bool := if l then true else false. + + Fixpoint mem (x : elt) (s : t) : bool := + match s with + | nil => false + | y :: l => + if X.eq_dec x y then true else mem x l + end. + + Fixpoint add (x : elt) (s : t) : t := + match s with + | nil => x :: nil + | y :: l => + if X.eq_dec x y then s else y :: add x l + end. + + Definition singleton (x : elt) : t := x :: nil. + + Fixpoint remove (x : elt) (s : t) : t := + match s with + | nil => nil + | y :: l => + if X.eq_dec x y then l else y :: remove x l + end. + + Definition fold (B : Type) (f : elt -> B -> B) (s : t) (i : B) : B := + fold_left (flip f) s i. + + Definition union (s : t) : t -> t := fold add s. + + Definition diff (s s' : t) : t := fold remove s' s. + + Definition inter (s s': t) : t := + fold (fun x s => if mem x s' then add x s else s) s nil. + + Definition subset (s s' : t) : bool := is_empty (diff s s'). + + Definition equal (s s' : t) : bool := andb (subset s s') (subset s' s). + + Fixpoint filter (f : elt -> bool) (s : t) : t := + match s with + | nil => nil + | x :: l => if f x then x :: filter f l else filter f l + end. + + Fixpoint for_all (f : elt -> bool) (s : t) : bool := + match s with + | nil => true + | x :: l => if f x then for_all f l else false + end. + + Fixpoint exists_ (f : elt -> bool) (s : t) : bool := + match s with + | nil => false + | x :: l => if f x then true else exists_ f l + end. + + Fixpoint partition (f : elt -> bool) (s : t) : t * t := + match s with + | nil => (nil, nil) + | x :: l => + let (s1, s2) := partition f l in + if f x then (x :: s1, s2) else (s1, x :: s2) + end. + + Definition cardinal (s : t) : nat := length s. + + Definition elements (s : t) : list elt := s. + + Definition choose (s : t) : option elt := + match s with + | nil => None + | x::_ => Some x + end. + +End Ops. + +(** ** Proofs of set operation specifications. *) + +Module MakeRaw (X:DecidableType) <: WRawSets X. + Include Ops X. + + Section ForNotations. + Notation NoDup := (NoDupA X.eq). + Notation In := (InA X.eq). + + (* TODO: modify proofs in order to avoid these hints *) + Hint Resolve (@Equivalence_Reflexive _ _ X.eq_equiv). + Hint Immediate (@Equivalence_Symmetric _ _ X.eq_equiv). + Hint Resolve (@Equivalence_Transitive _ _ X.eq_equiv). + + Definition IsOk := NoDup. + + Class Ok (s:t) : Prop := ok : NoDup s. + + Hint Unfold Ok. + Hint Resolve @ok. + + Instance NoDup_Ok s (nd : NoDup s) : Ok s := { ok := nd }. + + Ltac inv_ok := match goal with + | H:Ok (_ :: _) |- _ => inversion_clear H; inv_ok + | H:Ok nil |- _ => clear H; inv_ok + | H:NoDup ?l |- _ => change (Ok l) in H; inv_ok + | _ => idtac + end. + + Ltac inv := invlist InA; inv_ok. + Ltac constructors := repeat constructor. + + Fixpoint isok l := match l with + | nil => true + | a::l => negb (mem a l) && isok l + end. + + Definition Equal s s' := forall a : elt, In a s <-> In a s'. + Definition Subset s s' := forall a : elt, In a s -> In a s'. + Definition Empty s := forall a : elt, ~ In a s. + Definition For_all (P : elt -> Prop) s := forall x, In x s -> P x. + Definition Exists (P : elt -> Prop) s := exists x, In x s /\ P x. + + Lemma In_compat : Proper (X.eq==>eq==>iff) In. + Proof. + repeat red; intros. subst. rewrite H; auto. + Qed. + + Lemma mem_spec : forall s x `{Ok s}, + mem x s = true <-> In x s. + Proof. + induction s; intros. + split; intros; inv. discriminate. + simpl; destruct (X.eq_dec x a); split; intros; inv; auto. + right; rewrite <- IHs; auto. + rewrite IHs; auto. + Qed. + + Lemma isok_iff : forall l, Ok l <-> isok l = true. + Proof. + induction l. + intuition. + simpl. + rewrite andb_true_iff. + rewrite negb_true_iff. + rewrite <- IHl. + split; intros H. inv. + split; auto. + apply not_true_is_false. rewrite mem_spec; auto. + destruct H; constructors; auto. + rewrite <- mem_spec; auto; congruence. + Qed. + + Global Instance isok_Ok l : isok l = true -> Ok l | 10. + Proof. + intros. apply <- isok_iff; auto. + Qed. + + Lemma add_spec : + forall (s : t) (x y : elt) {Hs : Ok s}, + In y (add x s) <-> X.eq y x \/ In y s. + Proof. + induction s; simpl; intros. + intuition; inv; auto. + destruct X.eq_dec; inv; rewrite InA_cons, ?IHs; intuition. + left; eauto. + inv; auto. + Qed. + + Global Instance add_ok s x `(Ok s) : Ok (add x s). + Proof. + induction s. + simpl; intuition. + intros; inv. simpl. + destruct X.eq_dec; auto. + constructors; auto. + intro; inv; auto. + rewrite add_spec in *; intuition. + Qed. + + Lemma remove_spec : + forall (s : t) (x y : elt) {Hs : Ok s}, + In y (remove x s) <-> In y s /\ ~X.eq y x. + Proof. + induction s; simpl; intros. + intuition; inv; auto. + destruct X.eq_dec; inv; rewrite !InA_cons, ?IHs; intuition. + elim H. setoid_replace a with y; eauto. + elim H3. setoid_replace x with y; eauto. + elim n. eauto. + Qed. + + Global Instance remove_ok s x `(Ok s) : Ok (remove x s). + Proof. + induction s; simpl; intros. + auto. + destruct X.eq_dec; inv; auto. + constructors; auto. + rewrite remove_spec; intuition. + Qed. + + Lemma singleton_ok : forall x : elt, Ok (singleton x). + Proof. + unfold singleton; simpl; constructors; auto. intro; inv. + Qed. + + Lemma singleton_spec : forall x y : elt, In y (singleton x) <-> X.eq y x. + Proof. + unfold singleton; simpl; split; intros. inv; auto. left; auto. + Qed. + + Lemma empty_ok : Ok empty. + Proof. + unfold empty; constructors. + Qed. + + Lemma empty_spec : Empty empty. + Proof. + unfold Empty, empty; red; intros; inv. + Qed. + + Lemma is_empty_spec : forall s : t, is_empty s = true <-> Empty s. + Proof. + unfold Empty; destruct s; simpl; split; intros; auto. + intro; inv. + discriminate. + elim (H e); auto. + Qed. + + Lemma elements_spec1 : forall (s : t) (x : elt), In x (elements s) <-> In x s. + Proof. + unfold elements; intuition. + Qed. + + Lemma elements_spec2w : forall (s : t) {Hs : Ok s}, NoDup (elements s). + Proof. + unfold elements; auto. + Qed. + + Lemma fold_spec : + forall (s : t) (A : Type) (i : A) (f : elt -> A -> A), + fold f s i = fold_left (flip f) (elements s) i. + Proof. + reflexivity. + Qed. + + Global Instance union_ok : forall s s' `(Ok s, Ok s'), Ok (union s s'). + Proof. + induction s; simpl; auto; intros; inv; unfold flip; auto with *. + Qed. + + Lemma union_spec : + forall (s s' : t) (x : elt) {Hs : Ok s} {Hs' : Ok s'}, + In x (union s s') <-> In x s \/ In x s'. + Proof. + induction s; simpl in *; unfold flip; intros; auto; inv. + intuition; inv. + rewrite IHs, add_spec, InA_cons; intuition. + Qed. + + Global Instance inter_ok s s' `(Ok s, Ok s') : Ok (inter s s'). + Proof. + unfold inter, fold, flip. + set (acc := nil (A:=elt)). + assert (Hacc : Ok acc) by constructors. + clearbody acc; revert acc Hacc. + induction s; simpl; auto; intros. inv. + apply IHs; auto. + destruct (mem a s'); auto with *. + Qed. + + Lemma inter_spec : + forall (s s' : t) (x : elt) {Hs : Ok s} {Hs' : Ok s'}, + In x (inter s s') <-> In x s /\ In x s'. + Proof. + unfold inter, fold, flip; intros. + set (acc := nil (A:=elt)) in *. + assert (Hacc : Ok acc) by constructors. + assert (IFF : (In x s /\ In x s') <-> (In x s /\ In x s') \/ In x acc). + intuition; unfold acc in *; inv. + rewrite IFF; clear IFF. clearbody acc. + revert acc Hacc x s' Hs Hs'. + induction s; simpl; intros. + intuition; inv. + inv. + case_eq (mem a s'); intros Hm. + rewrite IHs, add_spec, InA_cons; intuition. + rewrite mem_spec in Hm; auto. + left; split; auto. rewrite H1; auto. + rewrite IHs, InA_cons; intuition. + rewrite H2, <- mem_spec in H3; auto. congruence. + Qed. + + Global Instance diff_ok : forall s s' `(Ok s, Ok s'), Ok (diff s s'). + Proof. + unfold diff; intros s s'; revert s. + induction s'; simpl; unfold flip; auto; intros. inv; auto with *. + Qed. + + Lemma diff_spec : + forall (s s' : t) (x : elt) {Hs : Ok s} {Hs' : Ok s'}, + In x (diff s s') <-> In x s /\ ~In x s'. + Proof. + unfold diff; intros s s'; revert s. + induction s'; simpl; unfold flip. + intuition; inv. + intros. inv. + rewrite IHs', remove_spec, InA_cons; intuition. + Qed. + + Lemma subset_spec : + forall (s s' : t) {Hs : Ok s} {Hs' : Ok s'}, + subset s s' = true <-> Subset s s'. + Proof. + unfold subset, Subset; intros. + rewrite is_empty_spec. + unfold Empty; intros. + intuition. + specialize (H a). rewrite diff_spec in H; intuition. + rewrite <- (mem_spec a) in H |- *. destruct (mem a s'); intuition. + rewrite diff_spec in H0; intuition. + Qed. + + Lemma equal_spec : + forall (s s' : t) {Hs : Ok s} {Hs' : Ok s'}, + equal s s' = true <-> Equal s s'. + Proof. + unfold Equal, equal; intros. + rewrite andb_true_iff, !subset_spec. + unfold Subset; intuition. rewrite <- H; auto. rewrite H; auto. + Qed. + + Definition choose_spec1 : + forall (s : t) (x : elt), choose s = Some x -> In x s. + Proof. + destruct s; simpl; intros; inversion H; auto. + Qed. + + Definition choose_spec2 : forall s : t, choose s = None -> Empty s. + Proof. + destruct s; simpl; intros. + intros x H0; inversion H0. + inversion H. + Qed. + + Lemma cardinal_spec : + forall (s : t) {Hs : Ok s}, cardinal s = length (elements s). + Proof. + auto. + Qed. + + Lemma filter_spec' : forall s x f, + In x (filter f s) -> In x s. + Proof. + induction s; simpl. + intuition; inv. + intros; destruct (f a); inv; intuition; right; eauto. + Qed. + + Lemma filter_spec : + forall (s : t) (x : elt) (f : elt -> bool), + Proper (X.eq==>eq) f -> + (In x (filter f s) <-> In x s /\ f x = true). + Proof. + induction s; simpl. + intuition; inv. + intros. + destruct (f a) as [ ]_eqn:E; rewrite ?InA_cons, IHs; intuition. + setoid_replace x with a; auto. + setoid_replace a with x in E; auto. congruence. + Qed. + + Global Instance filter_ok s f `(Ok s) : Ok (filter f s). + Proof. + induction s; simpl. + auto. + intros; inv. + case (f a); auto. + constructors; auto. + contradict H0. + eapply filter_spec'; eauto. + Qed. + + Lemma for_all_spec : + forall (s : t) (f : elt -> bool), + Proper (X.eq==>eq) f -> + (for_all f s = true <-> For_all (fun x => f x = true) s). + Proof. + unfold For_all; induction s; simpl. + intuition. inv. + intros; inv. + destruct (f a) as [ ]_eqn:F. + rewrite IHs; intuition. inv; auto. + setoid_replace x with a; auto. + split; intros H'; try discriminate. + intros. + rewrite <- F, <- (H' a); auto. + Qed. + + Lemma exists_spec : + forall (s : t) (f : elt -> bool), + Proper (X.eq==>eq) f -> + (exists_ f s = true <-> Exists (fun x => f x = true) s). + Proof. + unfold Exists; induction s; simpl. + split; [discriminate| intros (x & Hx & _); inv]. + intros. + destruct (f a) as [ ]_eqn:F. + split; auto. + exists a; auto. + rewrite IHs; firstorder. + inv. + setoid_replace a with x in F; auto; congruence. + exists x; auto. + Qed. + + Lemma partition_spec1 : + forall (s : t) (f : elt -> bool), + Proper (X.eq==>eq) f -> + Equal (fst (partition f s)) (filter f s). + Proof. + simple induction s; simpl; auto; unfold Equal. + firstorder. + intros x l Hrec f Hf. + generalize (Hrec f Hf); clear Hrec. + case (partition f l); intros s1 s2; simpl; intros. + case (f x); simpl; firstorder; inversion H0; intros; firstorder. + Qed. + + Lemma partition_spec2 : + forall (s : t) (f : elt -> bool), + Proper (X.eq==>eq) f -> + Equal (snd (partition f s)) (filter (fun x => negb (f x)) s). + Proof. + simple induction s; simpl; auto; unfold Equal. + firstorder. + intros x l Hrec f Hf. + generalize (Hrec f Hf); clear Hrec. + case (partition f l); intros s1 s2; simpl; intros. + case (f x); simpl; firstorder; inversion H0; intros; firstorder. + Qed. + + Lemma partition_ok1' : + forall (s : t) {Hs : Ok s} (f : elt -> bool)(x:elt), + In x (fst (partition f s)) -> In x s. + Proof. + induction s; simpl; auto; intros. inv. + generalize (IHs H1 f x). + destruct (f a); destruct (partition f s); simpl in *; auto. + inversion_clear H; auto. + Qed. + + Lemma partition_ok2' : + forall (s : t) {Hs : Ok s} (f : elt -> bool)(x:elt), + In x (snd (partition f s)) -> In x s. + Proof. + induction s; simpl; auto; intros. inv. + generalize (IHs H1 f x). + destruct (f a); destruct (partition f s); simpl in *; auto. + inversion_clear H; auto. + Qed. + + Global Instance partition_ok1 : forall s f `(Ok s), Ok (fst (partition f s)). + Proof. + simple induction s; simpl. + auto. + intros x l Hrec f Hs; inv. + generalize (@partition_ok1' _ _ f x). + generalize (Hrec f H0). + case (f x); case (partition f l); simpl; constructors; auto. + Qed. + + Global Instance partition_ok2 : forall s f `(Ok s), Ok (snd (partition f s)). + Proof. + simple induction s; simpl. + auto. + intros x l Hrec f Hs; inv. + generalize (@partition_ok2' _ _ f x). + generalize (Hrec f H0). + case (f x); case (partition f l); simpl; constructors; auto. + Qed. + + End ForNotations. + + Definition In := InA X.eq. + Definition eq := Equal. + Instance eq_equiv : Equivalence eq. + +End MakeRaw. + +(** * Encapsulation + + Now, in order to really provide a functor implementing [S], we + need to encapsulate everything into a type of lists without redundancy. *) + +Module Make (X: DecidableType) <: WSets with Module E := X. + Module Raw := MakeRaw X. + Include WRaw2Sets X Raw. +End Make. + diff --git a/theories/MSets/MSets.v b/theories/MSets/MSets.v new file mode 100644 index 00000000..958e9861 --- /dev/null +++ b/theories/MSets/MSets.v @@ -0,0 +1,23 @@ +(***********************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) +(* \VV/ *************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(***********************************************************************) + +(* $Id$ *) + +Require Export Orders. +Require Export OrdersEx. +Require Export OrdersAlt. +Require Export Equalities. +Require Export MSetInterface. +Require Export MSetFacts. +Require Export MSetDecide. +Require Export MSetProperties. +Require Export MSetEqProperties. +Require Export MSetWeakList. +Require Export MSetList. +Require Export MSetPositive. +Require Export MSetAVL.
\ No newline at end of file diff --git a/theories/MSets/vo.itarget b/theories/MSets/vo.itarget new file mode 100644 index 00000000..14429b81 --- /dev/null +++ b/theories/MSets/vo.itarget @@ -0,0 +1,11 @@ +MSetAVL.vo +MSetDecide.vo +MSetEqProperties.vo +MSetFacts.vo +MSetInterface.vo +MSetList.vo +MSetProperties.vo +MSets.vo +MSetToFiniteSet.vo +MSetWeakList.vo +MSetPositive.vo
\ No newline at end of file |