summaryrefslogtreecommitdiff
path: root/theories/FSets
diff options
context:
space:
mode:
Diffstat (limited to 'theories/FSets')
-rw-r--r--theories/FSets/FMapAVL.v2773
-rw-r--r--theories/FSets/FMapFacts.v1170
-rw-r--r--theories/FSets/FMapFullAVL.v823
-rw-r--r--theories/FSets/FMapIntMap.v622
-rw-r--r--theories/FSets/FMapInterface.v233
-rw-r--r--theories/FSets/FMapList.v120
-rw-r--r--theories/FSets/FMapPositive.v102
-rw-r--r--theories/FSets/FMapWeak.v15
-rw-r--r--theories/FSets/FMapWeakFacts.v599
-rw-r--r--theories/FSets/FMapWeakInterface.v201
-rw-r--r--theories/FSets/FMapWeakList.v127
-rw-r--r--theories/FSets/FMaps.v14
-rw-r--r--theories/FSets/FSetAVL.v3351
-rw-r--r--theories/FSets/FSetBridge.v104
-rw-r--r--theories/FSets/FSetDecide.v841
-rw-r--r--theories/FSets/FSetEqProperties.v167
-rw-r--r--theories/FSets/FSetFacts.v120
-rw-r--r--theories/FSets/FSetFullAVL.v1125
-rw-r--r--theories/FSets/FSetInterface.v265
-rw-r--r--theories/FSets/FSetList.v36
-rw-r--r--theories/FSets/FSetProperties.v806
-rw-r--r--theories/FSets/FSetToFiniteSet.v46
-rw-r--r--theories/FSets/FSetWeak.v16
-rw-r--r--theories/FSets/FSetWeakFacts.v421
-rw-r--r--theories/FSets/FSetWeakInterface.v251
-rw-r--r--theories/FSets/FSetWeakList.v114
-rw-r--r--theories/FSets/FSetWeakProperties.v896
-rw-r--r--theories/FSets/FSets.v8
-rw-r--r--theories/FSets/OrderedType.v36
-rw-r--r--theories/FSets/OrderedTypeAlt.v24
-rw-r--r--theories/FSets/OrderedTypeEx.v27
31 files changed, 8092 insertions, 7361 deletions
diff --git a/theories/FSets/FMapAVL.v b/theories/FSets/FMapAVL.v
index 4807ed66..8cb1236e 100644
--- a/theories/FSets/FMapAVL.v
+++ b/theories/FSets/FMapAVL.v
@@ -9,35 +9,35 @@
(* Finite map library. *)
-(* $Id: FMapAVL.v 9862 2007-05-25 16:57:06Z letouzey $ *)
+(* $Id: FMapAVL.v 11033 2008-06-01 22:56:50Z letouzey $ *)
-(** This module implements map using AVL trees.
- It follows the implementation from Ocaml's standard library. *)
+(** * FMapAVL *)
-Require Import FSetInterface.
-Require Import FMapInterface.
-Require Import FMapList.
+(** This module implements maps using AVL trees.
+ It follows the implementation from Ocaml's standard library.
+
+ See the comments at the beginning of FSetAVL for more details.
+*)
-Require Import ZArith.
-Require Import Int.
+Require Import FMapInterface FMapList ZArith Int.
-Set Firstorder Depth 3.
Set Implicit Arguments.
Unset Strict Implicit.
+(** Notations and helper lemma about pairs *)
-Module Raw (I:Int)(X: OrderedType).
-Import I.
-Module II:=MoreInt(I).
-Import II.
-Open Local Scope Int_scope.
+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.
-Module E := X.
-Module MX := OrderedTypeFacts X.
-Module PX := KeyOrderedType X.
-Module L := FMapList.Raw X.
-Import MX.
-Import PX.
+(** * The Raw functor
+
+ Functor of pure functions + separate proofs of invariant
+ preservation *)
+
+Module Raw (Import I:Int)(X: OrderedType).
+Open Local Scope pair_scope.
+Open Local Scope lazy_bool_scope.
+Open Local Scope Int_scope.
Definition key := X.t.
@@ -45,30 +45,391 @@ Definition key := X.t.
Section Elt.
-Variable elt : Set.
+Variable elt : Type.
-(* Now in KeyOrderedType:
-Definition eqk (p p':key*elt) := X.eq (fst p) (fst p').
-Definition eqke (p p':key*elt) :=
- X.eq (fst p) (fst p') /\ (snd p) = (snd p').
-Definition ltk (p p':key*elt) := X.lt (fst p) (fst p').
-*)
+(** * Trees
-Notation eqk := (eqk (elt:= elt)).
-Notation eqke := (eqke (elt:= elt)).
-Notation ltk := (ltk (elt:= elt)).
+ The fifth field of [Node] is the height of the tree *)
-Inductive tree : Set :=
+Inductive tree :=
| Leaf : tree
| Node : tree -> key -> elt -> tree -> int -> tree.
Notation t := tree.
-(** The Sixth field of [Node] is the height of the tree *)
+(** * Basic functions on trees: height and cardinal *)
+
+Definition height (m : t) : int :=
+ match m with
+ | Leaf => 0
+ | Node _ _ _ _ h => h
+ end.
+
+Fixpoint cardinal (m : t) : nat :=
+ match m with
+ | Leaf => 0%nat
+ | Node l _ _ r _ => S (cardinal l + cardinal r)
+ end.
+
+(** * Empty Map *)
+
+Definition empty := Leaf.
+
+(** * Emptyness test *)
+
+Definition is_empty m := match m with Leaf => true | _ => false end.
+
+(** * Appartness *)
+
+(** The [mem] function is deciding appartness. It exploits the [bst] property
+ to achieve logarithmic complexity. *)
+
+Fixpoint mem x m : bool :=
+ match m 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.
+
+Fixpoint find x m : option elt :=
+ match m with
+ | Leaf => None
+ | Node l y d r _ => match X.compare x y with
+ | LT _ => find x l
+ | EQ _ => Some d
+ | GT _ => find x r
+ end
+ end.
+
+(** * Helper functions *)
-(** * Occurrence in a tree *)
+(** [create l x r] creates a node, assuming [l] and [r]
+ to be balanced and [|height l - height r| <= 2]. *)
-Inductive MapsTo (x : key)(e : elt) : tree -> Prop :=
+Definition create l x e r :=
+ Node l x e r (max (height l) (height r) + 1).
+
+(** [bal l x e r] acts as [create], but performs one step of
+ rebalancing if necessary, i.e. assumes [|height l - height r| <= 3]. *)
+
+Definition assert_false := create.
+
+Fixpoint bal l x d 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 d r
+ | Node ll lx ld lr _ =>
+ if ge_lt_dec (height ll) (height lr) then
+ create ll lx ld (create lr x d r)
+ else
+ match lr with
+ | Leaf => assert_false l x d r
+ | Node lrl lrx lrd lrr _ =>
+ create (create ll lx ld lrl) lrx lrd (create lrr x d r)
+ end
+ end
+ else
+ if gt_le_dec hr (hl+2) then
+ match r with
+ | Leaf => assert_false l x d r
+ | Node rl rx rd rr _ =>
+ if ge_lt_dec (height rr) (height rl) then
+ create (create l x d rl) rx rd rr
+ else
+ match rl with
+ | Leaf => assert_false l x d r
+ | Node rll rlx rld rlr _ =>
+ create (create l x d rll) rlx rld (create rlr rx rd rr)
+ end
+ end
+ else
+ create l x d r.
+
+(** * Insertion *)
+
+Fixpoint add x d m :=
+ match m with
+ | Leaf => Node Leaf x d Leaf 1
+ | Node l y d' r h =>
+ match X.compare x y with
+ | LT _ => bal (add x d l) y d' r
+ | EQ _ => Node l y d r h
+ | GT _ => bal l y d' (add x d r)
+ end
+ end.
+
+(** * Extraction of minimum binding
+
+ Morally, [remove_min] is to be applied to a non-empty tree
+ [t = Node l x e 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 d r : t*(key*elt) :=
+ match l with
+ | Leaf => (r,(x,d))
+ | Node ll lx ld lr lh =>
+ let (l',m) := remove_min ll lx ld lr in
+ (bal l' x d 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].
+*)
+
+Fixpoint merge s1 s2 := match s1,s2 with
+ | Leaf, _ => s2
+ | _, Leaf => s1
+ | _, Node l2 x2 d2 r2 h2 =>
+ match remove_min l2 x2 d2 r2 with
+ (s2',(x,d)) => bal s1 x d s2'
+ end
+end.
+
+(** * Deletion *)
+
+Fixpoint remove x m := match m with
+ | Leaf => Leaf
+ | Node l y d r h =>
+ match X.compare x y with
+ | LT _ => bal (remove x l) y d r
+ | EQ _ => merge l r
+ | GT _ => bal l y d (remove x r)
+ end
+ end.
+
+(** * join
+
+ Same as [bal] but does not assume anything regarding heights of [l]
+ and [r].
+*)
+
+Fixpoint join l : key -> elt -> t -> t :=
+ match l with
+ | Leaf => add
+ | Node ll lx ld lr lh => fun x d =>
+ fix join_aux (r:t) : t := match r with
+ | Leaf => add x d l
+ | Node rl rx rd rr rh =>
+ if gt_le_dec lh (rh+2) then bal ll lx ld (join lr x d r)
+ else if gt_le_dec rh (lh+2) then bal (join_aux rl) rx rd rr
+ else create l x d r
+ end
+ end.
+
+(** * Splitting
+
+ [split x m] returns a triple [(l, o, r)] where
+ - [l] is the set of elements of [m] that are [< x]
+ - [r] is the set of elements of [m] that are [> x]
+ - [o] is the result of [find x m].
+*)
+
+Record triple := mktriple { t_left:t; t_opt:option elt; t_right:t }.
+Notation "<< l , b , r >>" := (mktriple l b r) (at level 9).
+
+Fixpoint split x m : triple := match m with
+ | Leaf => << Leaf, None, Leaf >>
+ | Node l y d r h =>
+ match X.compare x y with
+ | LT _ => let (ll,o,rl) := split x l in << ll, o, join rl y d r >>
+ | EQ _ => << l, Some d, r >>
+ | GT _ => let (rl,o,rr) := split x r in << join l y d rl, o, rr >>
+ end
+ end.
+
+(** * Concatenation
+
+ Same as [merge] but does not assume anything about heights.
+*)
+
+Definition concat m1 m2 :=
+ match m1, m2 with
+ | Leaf, _ => m2
+ | _ , Leaf => m1
+ | _, Node l2 x2 d2 r2 _ =>
+ let (m2',xd) := remove_min l2 x2 d2 r2 in
+ join m1 xd#1 xd#2 m2'
+ end.
+
+(** * Elements *)
+
+(** [elements_tree_aux acc t] catenates the elements of [t] in infix
+ order to the list [acc] *)
+
+Fixpoint elements_aux (acc : list (key*elt)) m : list (key*elt) :=
+ match m with
+ | Leaf => acc
+ | Node l x d r _ => elements_aux ((x,d) :: elements_aux acc r) l
+ end.
+
+(** then [elements] is an instanciation with an empty [acc] *)
+
+Definition elements := elements_aux nil.
+
+(** * Fold *)
+
+Fixpoint fold (A : Type) (f : key -> elt -> A -> A) (m : t) : A -> A :=
+ fun a => match m with
+ | Leaf => a
+ | Node l x d r _ => fold f r (f x d (fold f l a))
+ end.
+
+(** * Comparison *)
+
+Variable cmp : elt->elt->bool.
+
+(** ** Enumeration of the elements of a tree *)
+
+Inductive enumeration :=
+ | End : enumeration
+ | More : key -> elt -> t -> enumeration -> enumeration.
+
+(** [cons m e] adds the elements of tree [m] on the head of
+ enumeration [e]. *)
+
+Fixpoint cons m e : enumeration :=
+ match m with
+ | Leaf => e
+ | Node l x d r h => cons l (More x d r e)
+ end.
+
+(** One step of comparison of elements *)
+
+Definition equal_more x1 d1 (cont:enumeration->bool) e2 :=
+ match e2 with
+ | End => false
+ | More x2 d2 r2 e2 =>
+ match X.compare x1 x2 with
+ | EQ _ => cmp d1 d2 &&& cont (cons r2 e2)
+ | _ => false
+ end
+ end.
+
+(** Comparison of left tree, middle element, then right tree *)
+
+Fixpoint equal_cont m1 (cont:enumeration->bool) e2 :=
+ match m1 with
+ | Leaf => cont e2
+ | Node l1 x1 d1 r1 _ =>
+ equal_cont l1 (equal_more x1 d1 (equal_cont r1 cont)) e2
+ end.
+
+(** Initial continuation *)
+
+Definition equal_end e2 := match e2 with End => true | _ => false end.
+
+(** The complete comparison *)
+
+Definition equal m1 m2 := equal_cont m1 equal_end (cons m2 End).
+
+End Elt.
+Notation t := tree.
+Notation "<< l , b , r >>" := (mktriple l b r) (at level 9).
+Notation "t #l" := (t_left t) (at level 9, format "t '#l'").
+Notation "t #o" := (t_opt t) (at level 9, format "t '#o'").
+Notation "t #r" := (t_right t) (at level 9, format "t '#r'").
+
+
+(** * Map *)
+
+Fixpoint map (elt elt' : Type)(f : elt -> elt')(m : t elt) : t elt' :=
+ match m with
+ | Leaf => Leaf _
+ | Node l x d r h => Node (map f l) x (f d) (map f r) h
+ end.
+
+(* * Mapi *)
+
+Fixpoint mapi (elt elt' : Type)(f : key -> elt -> elt')(m : t elt) : t elt' :=
+ match m with
+ | Leaf => Leaf _
+ | Node l x d r h => Node (mapi f l) x (f x d) (mapi f r) h
+ end.
+
+(** * Map with removal *)
+
+Fixpoint map_option (elt elt' : Type)(f : key -> elt -> option elt')(m : t elt)
+ : t elt' :=
+ match m with
+ | Leaf => Leaf _
+ | Node l x d r h =>
+ match f x d with
+ | Some d' => join (map_option f l) x d' (map_option f r)
+ | None => concat (map_option f l) (map_option f r)
+ end
+ end.
+
+(** * Optimized map2
+
+ Suggestion by B. Gregoire: a [map2] function with specialized
+ arguments allowing to bypass some tree traversal. Instead of one
+ [f0] of type [key -> option elt -> option elt' -> option elt''],
+ we ask here for:
+ - [f] which is a specialisation of [f0] when first option isn't [None]
+ - [mapl] treats a [tree elt] with [f0] when second option is [None]
+ - [mapr] treats a [tree elt'] with [f0] when first option is [None]
+
+ The idea is that [mapl] and [mapr] can be instantaneous (e.g.
+ the identity or some constant function).
+*)
+
+Section Map2_opt.
+Variable elt elt' elt'' : Type.
+Variable f : key -> elt -> option elt' -> option elt''.
+Variable mapl : t elt -> t elt''.
+Variable mapr : t elt' -> t elt''.
+
+Fixpoint map2_opt m1 m2 :=
+ match m1, m2 with
+ | Leaf, _ => mapr m2
+ | _, Leaf => mapl m1
+ | Node l1 x1 d1 r1 h1, _ =>
+ let (l2',o2,r2') := split x1 m2 in
+ match f x1 d1 o2 with
+ | Some e => join (map2_opt l1 l2') x1 e (map2_opt r1 r2')
+ | None => concat (map2_opt l1 l2') (map2_opt r1 r2')
+ end
+ end.
+
+End Map2_opt.
+
+(** * Map2
+
+ The [map2] function of the Map interface can be implemented
+ via [map2_opt] and [map_option].
+*)
+
+Section Map2.
+Variable elt elt' elt'' : Type.
+Variable f : option elt -> option elt' -> option elt''.
+
+Definition map2 : t elt -> t elt' -> t elt'' :=
+ map2_opt
+ (fun _ d o => f (Some d) o)
+ (map_option (fun _ d => f (Some d) None))
+ (map_option (fun _ d' => f None (Some d'))).
+
+End Map2.
+
+
+
+(** * Invariants *)
+
+Section Invariants.
+Variable elt : Type.
+
+(** ** Occurrence in a tree *)
+
+Inductive MapsTo (x : key)(e : elt) : t elt -> Prop :=
| MapsRoot : forall l r h y,
X.eq x y -> MapsTo x e (Node l y e r h)
| MapsLeft : forall l r h y e',
@@ -76,7 +437,7 @@ Inductive MapsTo (x : key)(e : elt) : tree -> Prop :=
| MapsRight : forall l r h y e',
MapsTo x e r -> MapsTo x e (Node l y e' r h).
-Inductive In (x : key) : tree -> Prop :=
+Inductive In (x : key) : t elt -> Prop :=
| InRoot : forall l r h y e,
X.eq x y -> In x (Node l y e r h)
| InLeft : forall l r h y e',
@@ -84,58 +445,66 @@ Inductive In (x : key) : tree -> Prop :=
| InRight : forall l r h y e',
In x r -> In x (Node l y e' r h).
-Definition In0 (k:key)(m:t) : Prop := exists e:elt, MapsTo k e m.
+Definition In0 k m := exists e:elt, MapsTo k e m.
-(** * Binary search trees *)
+(** ** 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:key, In y s -> X.lt y x.
-Definition gt_tree x s := forall y:key, In y s -> X.lt x y.
+Definition lt_tree x m := forall y, In y m -> X.lt y x.
+Definition gt_tree x m := forall y, In y m -> X.lt x y.
(** [bst t] : [t] is a binary search tree *)
-Inductive bst : tree -> Prop :=
- | BSLeaf : bst Leaf
- | BSNode : forall x e l r h,
- bst l -> bst r -> lt_tree x l -> gt_tree x r -> bst (Node l x e r h).
+Inductive bst : t elt -> Prop :=
+ | BSLeaf : bst (Leaf _)
+ | BSNode : forall x e l r h, bst l -> bst r ->
+ lt_tree x l -> gt_tree x r -> bst (Node l x e r h).
-(** * AVL trees *)
+End Invariants.
-(** [avl s] : [s] is a properly balanced AVL tree,
- i.e. for any node the heights of the two children
- differ by at most 2 *)
-Definition height (s : tree) : int :=
- match s with
- | Leaf => 0
- | Node _ _ _ _ h => h
- end.
+(** * Correctness proofs, isolated in a sub-module *)
-Inductive avl : tree -> Prop :=
- | RBLeaf : avl Leaf
- | RBNode : forall x e l r h,
- avl l ->
- avl r ->
- -(2) <= height l - height r <= 2 ->
- h = max (height l) (height r) + 1 ->
- avl (Node l x e r h).
+Module Proofs.
+ Module MX := OrderedTypeFacts X.
+ Module PX := KeyOrderedType X.
+ Module L := FMapList.Raw X.
-(* We should end this section before the big proofs that follows,
- otherwise the discharge takes a lot of time. *)
-End Elt.
+Functional Scheme mem_ind := Induction for mem Sort Prop.
+Functional Scheme find_ind := Induction for find Sort Prop.
+Functional Scheme bal_ind := Induction for bal Sort Prop.
+Functional Scheme add_ind := Induction for add Sort Prop.
+Functional Scheme remove_min_ind := Induction for remove_min Sort Prop.
+Functional Scheme merge_ind := Induction for merge Sort Prop.
+Functional Scheme remove_ind := Induction for remove Sort Prop.
+Functional Scheme concat_ind := Induction for concat Sort Prop.
+Functional Scheme split_ind := Induction for split Sort Prop.
+Functional Scheme map_option_ind := Induction for map_option Sort Prop.
+Functional Scheme map2_opt_ind := Induction for map2_opt Sort Prop.
-(** Some helpful hints and tactics. *)
+(** * Automation and dedicated tactics. *)
-Notation t := tree.
-Hint Constructors tree.
-Hint Constructors MapsTo.
-Hint Constructors In.
-Hint Constructors bst.
-Hint Constructors avl.
+Hint Constructors tree MapsTo In bst.
Hint Unfold lt_tree gt_tree.
+Tactic Notation "factornode" ident(l) ident(x) ident(d) ident(r) ident(h)
+ "as" ident(s) :=
+ set (s:=Node l x d r h) in *; clearbody s; clear l x d r h.
+
+(** A tactic for cleaning hypothesis after use of functional induction. *)
+
+Ltac clearf :=
+ match goal with
+ | H : (@Logic.eq (Compare _ _ _ _) _ _) |- _ => clear H; clearf
+ | H : (@Logic.eq (sumbool _ _) _ _) |- _ => clear H; clearf
+ | _ => idtac
+ end.
+
+(** A tactic to repeat [inversion_clear] on all hyps of the
+ form [(f (Node ...))] *)
+
Ltac inv f :=
match goal with
| H:f (Leaf _) |- _ => inversion_clear H; inv f
@@ -149,14 +518,6 @@ Ltac inv f :=
| _ => idtac
end.
-Ltac safe_inv f := match goal with
- | H:f (Node _ _ _ _ _) |- _ =>
- generalize H; inversion_clear H; safe_inv f
- | H:f _ (Node _ _ _ _ _) |- _ =>
- generalize H; inversion_clear H; safe_inv f
- | _ => intros
- end.
-
Ltac inv_all f :=
match goal with
| H: f _ |- _ => inversion_clear H; inv f
@@ -166,55 +527,54 @@ Ltac inv_all f :=
| _ => idtac
end.
+(** Helper tactic concerning order of elements. *)
+
Ltac order := match goal with
- | H: lt_tree ?x ?s, H1: In ?y ?s |- _ => generalize (H _ H1); clear H; order
- | H: gt_tree ?x ?s, H1: In ?y ?s |- _ => generalize (H _ H1); clear H; order
+ | U: lt_tree _ ?s, V: In _ ?s |- _ => generalize (U _ V); clear U; order
+ | U: gt_tree _ ?s, V: In _ ?s |- _ => generalize (U _ V); clear U; order
| _ => MX.order
end.
Ltac intuition_in := repeat progress (intuition; inv In; inv MapsTo).
-Ltac firstorder_in := repeat progress (firstorder; inv In; inv MapsTo).
-
-Lemma height_non_negative : forall elt (s : t elt), avl s -> height s >= 0.
-Proof.
- induction s; simpl; intros; auto with zarith.
- inv avl; intuition; omega_max.
-Qed.
-
-Ltac avl_nn_hyp H :=
- let nz := fresh "nz" in assert (nz := height_non_negative H).
-Ltac avl_nn h :=
- let t := type of h in
- match type of t with
- | Prop => avl_nn_hyp h
- | _ => match goal with H : avl h |- _ => avl_nn_hyp H end
- end.
-
-(* Repeat the previous tactic.
- Drawback: need to clear the [avl _] hyps ... Thank you Ltac *)
+(* Function/Functional Scheme can't deal with internal fix.
+ Let's do its job by hand: *)
+
+Ltac join_tac :=
+ intros l; induction l as [| ll _ lx ld lr Hlr lh];
+ [ | intros x d r; induction r as [| rl Hrl rx rd rr _ rh]; unfold join;
+ [ | destruct (gt_le_dec lh (rh+2));
+ [ match goal with |- context [ bal ?u ?v ?w ?z ] =>
+ replace (bal u v w z)
+ with (bal ll lx ld (join lr x d (Node rl rx rd rr rh))); [ | auto]
+ end
+ | destruct (gt_le_dec rh (lh+2));
+ [ match goal with |- context [ bal ?u ?v ?w ?z ] =>
+ replace (bal u v w z)
+ with (bal (join (Node ll lx ld lr lh) x d rl) rx rd rr); [ | auto]
+ end
+ | ] ] ] ]; intros.
-Ltac avl_nns :=
- match goal with
- | H:avl _ |- _ => avl_nn_hyp H; clear H; avl_nns
- | _ => idtac
- end.
+Section Elt.
+Variable elt:Type.
+Implicit Types m r : t elt.
+(** * Basic results about [MapsTo], [In], [lt_tree], [gt_tree], [height] *)
(** Facts about [MapsTo] and [In]. *)
-Lemma MapsTo_In : forall elt k e (m:t elt), MapsTo k e m -> In k m.
+Lemma MapsTo_In : forall k e m, MapsTo k e m -> In k m.
Proof.
induction 1; auto.
Qed.
Hint Resolve MapsTo_In.
-Lemma In_MapsTo : forall elt k (m:t elt), In k m -> exists e, MapsTo k e m.
+Lemma In_MapsTo : forall k m, In k m -> exists e, MapsTo k e m.
Proof.
induction 1; try destruct IHIn as (e,He); exists e; auto.
Qed.
-Lemma In_alt : forall elt k (m:t elt), In0 k m <-> In k m.
+Lemma In_alt : forall k m, In0 k m <-> In k m.
Proof.
split.
intros (e,H); eauto.
@@ -222,64 +582,70 @@ Proof.
Qed.
Lemma MapsTo_1 :
- forall elt (m:t elt) x y e, X.eq x y -> MapsTo x e m -> MapsTo y e m.
+ forall m x y e, X.eq x y -> MapsTo x e m -> MapsTo y e m.
Proof.
induction m; simpl; intuition_in; eauto.
Qed.
Hint Immediate MapsTo_1.
Lemma In_1 :
- forall elt (m:t elt) x y, X.eq x y -> In x m -> In y m.
+ forall m x y, X.eq x y -> In x m -> In y m.
Proof.
- intros elt m x y; induction m; simpl; intuition_in; eauto.
+ intros m x y; induction m; simpl; intuition_in; eauto.
Qed.
+Lemma In_node_iff :
+ forall l x e r h y,
+ In y (Node l x e r h) <-> In y l \/ X.eq y x \/ In y r.
+Proof.
+ intuition_in.
+Qed.
(** Results about [lt_tree] and [gt_tree] *)
-Lemma lt_leaf : forall elt x, lt_tree x (Leaf elt).
+Lemma lt_leaf : forall x, lt_tree x (Leaf elt).
Proof.
unfold lt_tree in |- *; intros; intuition_in.
Qed.
-Lemma gt_leaf : forall elt x, gt_tree x (Leaf elt).
+Lemma gt_leaf : forall x, gt_tree x (Leaf elt).
Proof.
unfold gt_tree in |- *; intros; intuition_in.
Qed.
-Lemma lt_tree_node : forall elt x y (l:t elt) r e h,
+Lemma lt_tree_node : forall x y l r e h,
lt_tree x l -> lt_tree x r -> X.lt y x -> lt_tree x (Node l y e r h).
Proof.
- unfold lt_tree in *; firstorder_in; order.
+ unfold lt_tree in *; intuition_in; order.
Qed.
-Lemma gt_tree_node : forall elt x y (l:t elt) r e h,
+Lemma gt_tree_node : forall x y l r e h,
gt_tree x l -> gt_tree x r -> X.lt x y -> gt_tree x (Node l y e r h).
Proof.
- unfold gt_tree in *; firstorder_in; order.
+ unfold gt_tree in *; intuition_in; order.
Qed.
Hint Resolve lt_leaf gt_leaf lt_tree_node gt_tree_node.
-Lemma lt_left : forall elt x y (l: t elt) r e h,
+Lemma lt_left : forall x y l r e h,
lt_tree x (Node l y e r h) -> lt_tree x l.
Proof.
intuition_in.
Qed.
-Lemma lt_right : forall elt x y (l:t elt) r e h,
+Lemma lt_right : forall x y l r e h,
lt_tree x (Node l y e r h) -> lt_tree x r.
Proof.
intuition_in.
Qed.
-Lemma gt_left : forall elt x y (l:t elt) r e h,
+Lemma gt_left : forall x y l r e h,
gt_tree x (Node l y e r h) -> gt_tree x l.
Proof.
intuition_in.
Qed.
-Lemma gt_right : forall elt x y (l:t elt) r e h,
+Lemma gt_right : forall x y l r e h,
gt_tree x (Node l y e r h) -> gt_tree x r.
Proof.
intuition_in.
@@ -288,731 +654,639 @@ Qed.
Hint Resolve lt_left lt_right gt_left gt_right.
Lemma lt_tree_not_in :
- forall elt x (t : t elt), lt_tree x t -> ~ In x t.
+ forall x m, lt_tree x m -> ~ In x m.
Proof.
intros; intro; generalize (H _ H0); order.
Qed.
Lemma lt_tree_trans :
- forall elt x y, X.lt x y -> forall (t:t elt), lt_tree x t -> lt_tree y t.
+ forall x y, X.lt x y -> forall m, lt_tree x m -> lt_tree y m.
Proof.
- firstorder eauto.
+ eauto.
Qed.
Lemma gt_tree_not_in :
- forall elt x (t : t elt), gt_tree x t -> ~ In x t.
+ forall x m, gt_tree x m -> ~ In x m.
Proof.
intros; intro; generalize (H _ H0); order.
Qed.
Lemma gt_tree_trans :
- forall elt x y, X.lt y x -> forall (t:t elt), gt_tree x t -> gt_tree y t.
+ forall x y, X.lt y x -> forall m, gt_tree x m -> gt_tree y m.
Proof.
- firstorder eauto.
+ eauto.
Qed.
Hint Resolve lt_tree_not_in lt_tree_trans gt_tree_not_in gt_tree_trans.
-(** Results about [avl] *)
+(** * Empty map *)
-Lemma avl_node : forall elt x e (l:t elt) r,
- avl l ->
- avl r ->
- -(2) <= height l - height r <= 2 ->
- avl (Node l x e r (max (height l) (height r) + 1)).
+Definition Empty m := forall (a:key)(e:elt) , ~ MapsTo a e m.
+
+Lemma empty_bst : bst (empty elt).
Proof.
- intros; auto.
+ unfold empty; auto.
Qed.
-Hint Resolve avl_node.
-(** * Helper functions *)
-
-(** [create l x r] creates a node, assuming [l] and [r]
- to be balanced and [|height l - height r| <= 2]. *)
+Lemma empty_1 : Empty (empty elt).
+Proof.
+ unfold empty, Empty; intuition_in.
+Qed.
-Definition create elt (l:t elt) x e r :=
- Node l x e r (max (height l) (height r) + 1).
+(** * Emptyness test *)
-Lemma create_bst :
- forall elt (l:t elt) x e r, bst l -> bst r -> lt_tree x l -> gt_tree x r ->
- bst (create l x e r).
+Lemma is_empty_1 : forall m, Empty m -> is_empty m = true.
Proof.
- unfold create; auto.
+ destruct m as [|r x e l h]; simpl; auto.
+ intro H; elim (H x e); auto.
Qed.
-Hint Resolve create_bst.
-Lemma create_avl :
- forall elt (l:t elt) x e r, avl l -> avl r -> -(2) <= height l - height r <= 2 ->
- avl (create l x e r).
-Proof.
- unfold create; auto.
+Lemma is_empty_2 : forall m, is_empty m = true -> Empty m.
+Proof.
+ destruct m; simpl; intros; try discriminate; red; intuition_in.
Qed.
-Lemma create_height :
- forall elt (l:t elt) x e r, avl l -> avl r -> -(2) <= height l - height r <= 2 ->
- height (create l x e r) = max (height l) (height r) + 1.
-Proof.
- unfold create; intros; auto.
+(** * Appartness *)
+
+Lemma mem_1 : forall m x, bst m -> In x m -> mem x m = true.
+Proof.
+ intros m x; functional induction (mem x m); auto; intros; clearf;
+ inv bst; intuition_in; order.
Qed.
-Lemma create_in :
- forall elt (l:t elt) x e r y, In y (create l x e r) <-> X.eq y x \/ In y l \/ In y r.
-Proof.
- unfold create; split; [ inversion_clear 1 | ]; intuition.
+Lemma mem_2 : forall m x, mem x m = true -> In x m.
+Proof.
+ intros m x; functional induction (mem x m); auto; intros; discriminate.
Qed.
-(** trick for emulating [assert false] in Coq *)
+Lemma find_1 : forall m x e, bst m -> MapsTo x e m -> find x m = Some e.
+Proof.
+ intros m x; functional induction (find x m); auto; intros; clearf;
+ inv bst; intuition_in; simpl; auto;
+ try solve [order | absurd (X.lt x y); eauto | absurd (X.lt y x); eauto].
+Qed.
-Notation assert_false := Leaf.
+Lemma find_2 : forall m x e, find x m = Some e -> MapsTo x e m.
+Proof.
+ intros m x; functional induction (find x m); subst; intros; clearf;
+ try discriminate.
+ constructor 2; auto.
+ inversion H; auto.
+ constructor 3; auto.
+Qed.
-(** [bal l x e r] acts as [create], but performs one step of
- rebalancing if necessary, i.e. assumes [|height l - height r| <= 3]. *)
+Lemma find_iff : forall m x e, bst m ->
+ (find x m = Some e <-> MapsTo x e m).
+Proof.
+ split; auto using find_1, find_2.
+Qed.
-Definition bal elt (l: tree elt) x e 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 _
- | Node ll lx le lr _ =>
- if ge_lt_dec (height ll) (height lr) then
- create ll lx le (create lr x e r)
- else
- match lr with
- | Leaf => assert_false _
- | Node lrl lrx lre lrr _ =>
- create (create ll lx le lrl) lrx lre (create lrr x e r)
- end
- end
- else
- if gt_le_dec hr (hl+2) then
- match r with
- | Leaf => assert_false _
- | Node rl rx re rr _ =>
- if ge_lt_dec (height rr) (height rl) then
- create (create l x e rl) rx re rr
- else
- match rl with
- | Leaf => assert_false _
- | Node rll rlx rle rlr _ =>
- create (create l x e rll) rlx rle (create rlr rx re rr)
- end
- end
- else
- create l x e r.
-
-Ltac bal_tac :=
- intros elt l x e r;
- unfold bal;
- destruct (gt_le_dec (height l) (height r + 2));
- [ destruct l as [ |ll lx le lr lh];
- [ | destruct (ge_lt_dec (height ll) (height lr));
- [ | destruct lr ] ]
- | destruct (gt_le_dec (height r) (height l + 2));
- [ destruct r as [ |rl rx re rr rh];
- [ | destruct (ge_lt_dec (height rr) (height rl));
- [ | destruct rl ] ]
- | ] ]; intros.
-
-Ltac bal_tac_imp := match goal with
- | |- context [ assert_false ] =>
- inv avl; avl_nns; simpl in *; false_omega
- | _ => idtac
-end.
+Lemma find_in : forall m x, find x m <> None -> In x m.
+Proof.
+ intros.
+ case_eq (find x m); [intros|congruence].
+ apply MapsTo_In with e; apply find_2; auto.
+Qed.
-Lemma bal_bst : forall elt (l:t elt) x e r, bst l -> bst r ->
- lt_tree x l -> gt_tree x r -> bst (bal l x e r).
+Lemma in_find : forall m x, bst m -> In x m -> find x m <> None.
Proof.
- bal_tac;
- inv bst; repeat apply create_bst; auto; unfold create;
- apply lt_tree_node || apply gt_tree_node; auto;
- eapply lt_tree_trans || eapply gt_tree_trans || eauto; eauto.
+ intros.
+ destruct (In_MapsTo H0) as (d,Hd).
+ rewrite (find_1 H Hd); discriminate.
Qed.
-Lemma bal_avl : forall elt (l:t elt) x e r, avl l -> avl r ->
- -(3) <= height l - height r <= 3 -> avl (bal l x e r).
+Lemma find_in_iff : forall m x, bst m ->
+ (find x m <> None <-> In x m).
Proof.
- bal_tac; inv avl; repeat apply create_avl; simpl in *; auto; omega_max.
+ split; auto using find_in, in_find.
Qed.
-Lemma bal_height_1 : forall elt (l:t elt) x e r, avl l -> avl r ->
- -(3) <= height l - height r <= 3 ->
- 0 <= height (bal l x e r) - max (height l) (height r) <= 1.
+Lemma not_find_iff : forall m x, bst m ->
+ (find x m = None <-> ~In x m).
Proof.
- bal_tac; inv avl; avl_nns; simpl in *; omega_max.
+ split; intros.
+ red; intros.
+ elim (in_find H H1 H0).
+ case_eq (find x m); [ intros | auto ].
+ elim H0; apply find_in; congruence.
Qed.
-Lemma bal_height_2 :
- forall elt (l:t elt) x e r, avl l -> avl r -> -(2) <= height l - height r <= 2 ->
- height (bal l x e r) == max (height l) (height r) +1.
+Lemma find_find : forall m m' x,
+ find x m = find x m' <->
+ (forall d, find x m = Some d <-> find x m' = Some d).
Proof.
- bal_tac; inv avl; simpl in *; omega_max.
+ intros; destruct (find x m); destruct (find x m'); split; intros;
+ try split; try congruence.
+ rewrite H; auto.
+ symmetry; rewrite <- H; auto.
+ rewrite H; auto.
Qed.
-Lemma bal_in : forall elt (l:t elt) x e r y, avl l -> avl r ->
- (In y (bal l x e r) <-> X.eq y x \/ In y l \/ In y r).
+Lemma find_mapsto_equiv : forall m m' x, bst m -> bst m' ->
+ (find x m = find x m' <->
+ (forall d, MapsTo x d m <-> MapsTo x d m')).
Proof.
- bal_tac; bal_tac_imp; repeat rewrite create_in; intuition_in.
+ intros m m' x Hm Hm'.
+ rewrite find_find.
+ split; intros H d; specialize H with d.
+ rewrite <- 2 find_iff; auto.
+ rewrite 2 find_iff; auto.
Qed.
-Lemma bal_mapsto : forall elt (l:t elt) x e r y e', avl l -> avl r ->
- (MapsTo y e' (bal l x e r) <-> MapsTo y e' (create l x e r)).
+Lemma find_in_equiv : forall m m' x, bst m -> bst m' ->
+ find x m = find x m' ->
+ (In x m <-> In x m').
Proof.
- bal_tac; bal_tac_imp; unfold create; intuition_in.
+ split; intros; apply find_in; [ rewrite <- H1 | rewrite H1 ];
+ apply in_find; auto.
Qed.
-Ltac omega_bal := match goal with
- | H:avl ?l, H':avl ?r |- context [ bal ?l ?x ?e ?r ] =>
- generalize (bal_height_1 x e H H') (bal_height_2 x e H H');
- omega_max
- end.
+(** * Helper functions *)
-(** * Insertion *)
+Lemma create_bst :
+ forall l x e r, bst l -> bst r -> lt_tree x l -> gt_tree x r ->
+ bst (create l x e r).
+Proof.
+ unfold create; auto.
+Qed.
+Hint Resolve create_bst.
-Function add (elt:Set)(x:key)(e:elt)(s:t elt) { struct s } : t elt := match s with
- | Leaf => Node (Leaf _) x e (Leaf _) 1
- | Node l y e' r h =>
- match X.compare x y with
- | LT _ => bal (add x e l) y e' r
- | EQ _ => Node l y e r h
- | GT _ => bal l y e' (add x e r)
- end
- end.
+Lemma create_in :
+ forall l x e r y,
+ In y (create l x e r) <-> X.eq y x \/ In y l \/ In y r.
+Proof.
+ unfold create; split; [ inversion_clear 1 | ]; intuition.
+Qed.
-Lemma add_avl_1 : forall elt (m:t elt) x e, avl m ->
- avl (add x e m) /\ 0 <= height (add x e m) - height m <= 1.
-Proof.
- intros elt m x e; functional induction (add x e m); intros; inv avl; simpl in *.
- intuition; try constructor; simpl; auto; try omega_max.
- (* LT *)
- destruct IHt; auto.
- split.
- apply bal_avl; auto; omega_max.
- omega_bal.
- (* EQ *)
- intuition; omega_max.
- (* GT *)
- destruct IHt; auto.
- split.
- apply bal_avl; auto; omega_max.
- omega_bal.
+Lemma bal_bst : forall l x e r, bst l -> bst r ->
+ lt_tree x l -> gt_tree x r -> bst (bal l x e r).
+Proof.
+ intros l x e r; functional induction (bal l x e r); intros; clearf;
+ inv bst; repeat apply create_bst; auto; unfold create; try constructor;
+ (apply lt_tree_node || apply gt_tree_node); auto;
+ (eapply lt_tree_trans || eapply gt_tree_trans); eauto.
Qed.
+Hint Resolve bal_bst.
-Lemma add_avl : forall elt (m:t elt) x e, avl m -> avl (add x e m).
+Lemma bal_in : forall l x e r y,
+ In y (bal l x e r) <-> X.eq y x \/ In y l \/ In y r.
Proof.
- intros; generalize (add_avl_1 x e H); intuition.
+ intros l x e r; functional induction (bal l x e r); intros; clearf;
+ rewrite !create_in; intuition_in.
Qed.
-Hint Resolve add_avl.
-Lemma add_in : forall elt (m:t elt) x y e, avl m ->
- (In y (add x e m) <-> X.eq y x \/ In y m).
+Lemma bal_mapsto : forall l x e r y e',
+ MapsTo y e' (bal l x e r) <-> MapsTo y e' (create l x e r).
Proof.
- intros elt m x y e; functional induction (add x e m); auto; intros.
- intuition_in.
- (* LT *)
- inv avl.
- rewrite bal_in; auto.
- rewrite (IHt H0); intuition_in.
- (* EQ *)
- inv avl.
- firstorder_in.
- eapply In_1; eauto.
- (* GT *)
- inv avl.
- rewrite bal_in; auto.
- rewrite (IHt H1); intuition_in.
+ intros l x e r; functional induction (bal l x e r); intros; clearf;
+ unfold assert_false, create; intuition_in.
Qed.
-Lemma add_bst : forall elt (m:t elt) x e, bst m -> avl m -> bst (add x e m).
-Proof.
- intros elt m x e; functional induction (add x e m);
- intros; inv bst; inv avl; auto; apply bal_bst; auto.
- (* lt_tree -> lt_tree (add ...) *)
- red; red in H4.
- intros.
- rewrite (add_in x y0 e H) in H0.
- intuition.
- eauto.
- (* gt_tree -> gt_tree (add ...) *)
- red; red in H4.
- intros.
- rewrite (add_in x y0 e H5) in H0.
- intuition.
- apply lt_eq with x; auto.
+Lemma bal_find : forall l x e r y,
+ bst l -> bst r -> lt_tree x l -> gt_tree x r ->
+ find y (bal l x e r) = find y (create l x e r).
+Proof.
+ intros; rewrite find_mapsto_equiv; auto; intros; apply bal_mapsto.
Qed.
-Lemma add_1 : forall elt (m:t elt) x y e, avl m -> X.eq x y -> MapsTo y e (add x e m).
+(** * Insertion *)
+
+Lemma add_in : forall m x y e,
+ In y (add x e m) <-> X.eq y x \/ In y m.
+Proof.
+ intros m x y e; functional induction (add x e m); auto; intros;
+ try (rewrite bal_in, IHt); intuition_in.
+ apply In_1 with x; auto.
+Qed.
+
+Lemma add_bst : forall m x e, bst m -> bst (add x e m).
+Proof.
+ intros m x e; functional induction (add x e m); intros;
+ inv bst; try apply bal_bst; auto;
+ intro z; rewrite add_in; intuition.
+ apply MX.eq_lt with x; auto.
+ apply MX.lt_eq with x; auto.
+Qed.
+Hint Resolve add_bst.
+
+Lemma add_1 : forall m x y e, X.eq x y -> MapsTo y e (add x e m).
Proof.
- intros elt m x y e; functional induction (add x e m);
- intros; inv bst; inv avl; try rewrite bal_mapsto; unfold create; eauto.
-Qed.
+ intros m x y e; functional induction (add x e m);
+ intros; inv bst; try rewrite bal_mapsto; unfold create; eauto.
+Qed.
-Lemma add_2 : forall elt (m:t elt) x y e e', avl m -> ~X.eq x y ->
+Lemma add_2 : forall m x y e e', ~X.eq x y ->
MapsTo y e m -> MapsTo y e (add x e' m).
Proof.
- intros elt m x y e e'; induction m; simpl; auto.
+ intros m x y e e'; induction m; simpl; auto.
destruct (X.compare x k);
- intros; inv bst; inv avl; try rewrite bal_mapsto; unfold create; auto;
+ intros; inv bst; try rewrite bal_mapsto; unfold create; auto;
inv MapsTo; auto; order.
Qed.
-Lemma add_3 : forall elt (m:t elt) x y e e', avl m -> ~X.eq x y ->
+Lemma add_3 : forall m x y e e', ~X.eq x y ->
MapsTo y e (add x e' m) -> MapsTo y e m.
Proof.
- intros elt m x y e e'; induction m; simpl; auto.
- intros; inv avl; inv MapsTo; auto; order.
- destruct (X.compare x k); intro; inv avl;
+ intros m x y e e'; induction m; simpl; auto.
+ intros; inv MapsTo; auto; order.
+ destruct (X.compare x k); intro;
try rewrite bal_mapsto; auto; unfold create; intros; inv MapsTo; auto;
- order.
+ order.
Qed.
-
-(** * Extraction of minimum binding
-
- morally, [remove_min] is to be applied to a non-empty tree
- [t = Node l x e r h]. Since we can't deal here with [assert false]
- for [t=Leaf], we pre-unpack [t] (and forget about [h]).
-*)
-
-Function remove_min (elt:Set)(l:t elt)(x:key)(e:elt)(r:t elt) { struct l } : t elt*(key*elt) :=
- match l with
- | Leaf => (r,(x,e))
- | Node ll lx le lr lh => let (l',m) := (remove_min ll lx le lr : t elt*(key*elt)) in (bal l' x e r, m)
- end.
-
-Lemma remove_min_avl_1 : forall elt (l:t elt) x e r h, avl (Node l x e r h) ->
- avl (fst (remove_min l x e r)) /\
- 0 <= height (Node l x e r h) - height (fst (remove_min l x e r)) <= 1.
+Lemma add_find : forall m x y e, bst m ->
+ find y (add x e m) =
+ match X.compare y x with EQ _ => Some e | _ => find y m end.
Proof.
- intros elt l x e r; functional induction (remove_min l x e r); simpl in *; intros.
- inv avl; simpl in *; split; auto.
- avl_nns; omega_max.
- (* l = Node *)
- inversion_clear H.
- destruct (IHp lh); auto.
- split; simpl in *.
- rewrite_all e1. simpl in *.
- apply bal_avl; subst;auto; omega_max.
- rewrite_all e1;simpl in *;omega_bal.
+ intros.
+ assert (~X.eq x y -> find y (add x e m) = find y m).
+ intros; rewrite find_mapsto_equiv; auto.
+ split; eauto using add_2, add_3.
+ destruct X.compare; try (apply H0; order).
+ auto using find_1, add_1.
Qed.
-Lemma remove_min_avl : forall elt (l:t elt) x e r h, avl (Node l x e r h) ->
- avl (fst (remove_min l x e r)).
-Proof.
- intros; generalize (remove_min_avl_1 H); intuition.
-Qed.
+(** * Extraction of minimum binding *)
-Lemma remove_min_in : forall elt (l:t elt) x e r h y, avl (Node l x e r h) ->
- (In y (Node l x e r h) <->
- X.eq y (fst (snd (remove_min l x e r))) \/ In y (fst (remove_min l x e r))).
+Lemma remove_min_in : forall l x e r h y,
+ In y (Node l x e r h) <->
+ X.eq y (remove_min l x e r)#2#1 \/ In y (remove_min l x e r)#1.
Proof.
- intros elt l x e r; functional induction (remove_min l x e r); simpl in *; intros.
+ intros l x e r; functional induction (remove_min l x e r); simpl in *; intros.
intuition_in.
- (* l = Node *)
- inversion_clear H.
- generalize (remove_min_avl H0).
-
- rewrite_all e1; simpl; intros.
- rewrite bal_in; auto.
- generalize (IHp lh y H0).
- intuition.
- inversion_clear H7; intuition.
+ rewrite e0 in *; simpl; intros.
+ rewrite bal_in, In_node_iff, IHp; intuition.
Qed.
-Lemma remove_min_mapsto : forall elt (l:t elt) x e r h y e', avl (Node l x e r h) ->
- (MapsTo y e' (Node l x e r h) <->
- ((X.eq y (fst (snd (remove_min l x e r))) /\ e' = (snd (snd (remove_min l x e r))))
- \/ MapsTo y e' (fst (remove_min l x e r)))).
+Lemma remove_min_mapsto : forall l x e r h y e',
+ MapsTo y e' (Node l x e r h) <->
+ ((X.eq y (remove_min l x e r)#2#1) /\ e' = (remove_min l x e r)#2#2)
+ \/ MapsTo y e' (remove_min l x e r)#1.
Proof.
- intros elt l x e r; functional induction (remove_min l x e r); simpl in *; intros.
+ intros l x e r; functional induction (remove_min l x e r); simpl in *; intros.
intuition_in; subst; auto.
- (* l = Node *)
- inversion_clear H.
- generalize (remove_min_avl H0).
- rewrite_all e1; simpl; intros.
+ rewrite e0 in *; simpl; intros.
rewrite bal_mapsto; auto; unfold create.
- simpl in *;destruct (IHp lh y e').
- auto.
+ simpl in *;destruct (IHp _x y e').
intuition.
- inversion_clear H2; intuition.
- inversion_clear H9; intuition.
+ inversion_clear H1; intuition.
+ inversion_clear H3; intuition.
Qed.
-Lemma remove_min_bst : forall elt (l:t elt) x e r h,
- bst (Node l x e r h) -> avl (Node l x e r h) -> bst (fst (remove_min l x e r)).
+Lemma remove_min_bst : forall l x e r h,
+ bst (Node l x e r h) -> bst (remove_min l x e r)#1.
Proof.
- intros elt l x e r; functional induction (remove_min l x e r); simpl in *; intros.
+ intros l x e r; functional induction (remove_min l x e r); simpl in *; intros.
inv bst; auto.
inversion_clear H; inversion_clear H0.
apply bal_bst; auto.
- rewrite_all e1;simpl in *;firstorder.
+ rewrite e0 in *; simpl in *; apply (IHp _x); auto.
intro; intros.
- generalize (remove_min_in y H).
- rewrite_all e1; simpl in *.
+ generalize (remove_min_in ll lx ld lr _x y).
+ rewrite e0; simpl in *.
destruct 1.
- apply H3; intuition.
+ apply H2; intuition.
Qed.
+Hint Resolve remove_min_bst.
-Lemma remove_min_gt_tree : forall elt (l:t elt) x e r h,
- bst (Node l x e r h) -> avl (Node l x e r h) ->
- gt_tree (fst (snd (remove_min l x e r))) (fst (remove_min l x e r)).
+Lemma remove_min_gt_tree : forall l x e r h,
+ bst (Node l x e r h) ->
+ gt_tree (remove_min l x e r)#2#1 (remove_min l x e r)#1.
Proof.
- intros elt l x e r; functional induction (remove_min l x e r); simpl in *; intros.
+ intros l x e r; functional induction (remove_min l x e r); simpl in *; intros.
inv bst; auto.
- inversion_clear H; inversion_clear H0.
+ inversion_clear H.
intro; intro.
- rewrite_all e1;simpl in *.
- generalize (IHp lh H1 H); clear H7 H6 IHp.
- generalize (remove_min_avl H).
- generalize (remove_min_in (fst m) H).
- rewrite e1; simpl; intros.
- rewrite (bal_in x e y H7 H5) in H0.
- destruct H6.
- firstorder.
- apply lt_eq with x; auto.
- apply X.lt_trans with x; auto.
-Qed.
-
-(** * 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].
-*)
-
-Function merge (elt:Set) (s1 s2 : t elt) : tree elt := match s1,s2 with
- | Leaf, _ => s2
- | _, Leaf => s1
- | _, Node l2 x2 e2 r2 h2 =>
- match remove_min l2 x2 e2 r2 with
- (s2',(x,e)) => bal s1 x e s2'
- end
-end.
-
-Lemma merge_avl_1 : forall elt (s1 s2:t elt), avl s1 -> avl s2 ->
- -(2) <= height s1 - height s2 <= 2 ->
- avl (merge s1 s2) /\
- 0<= height (merge s1 s2) - max (height s1) (height s2) <=1.
-Proof.
- intros elt s1 s2; functional induction (merge s1 s2); simpl in *; intros.
- split; auto; avl_nns; omega_max.
- destruct s1;try contradiction;clear y.
- split; auto; avl_nns; simpl in *; omega_max.
- destruct s1;try contradiction;clear y.
- generalize (remove_min_avl_1 H0).
- rewrite e3; simpl;destruct 1.
- split.
- apply bal_avl; auto.
- simpl; omega_max.
- omega_bal.
+ rewrite e0 in *;simpl in *.
+ generalize (IHp _x H0).
+ generalize (remove_min_in ll lx ld lr _x m#1).
+ rewrite e0; simpl; intros.
+ rewrite (bal_in l' x d r y) in H.
+ assert (In m#1 (Node ll lx ld lr _x)) by (rewrite H4; auto); clear H4.
+ assert (X.lt m#1 x) by order.
+ decompose [or] H; order.
+Qed.
+Hint Resolve remove_min_gt_tree.
+
+Lemma remove_min_find : forall l x e r h y,
+ bst (Node l x e r h) ->
+ find y (Node l x e r h) =
+ match X.compare y (remove_min l x e r)#2#1 with
+ | LT _ => None
+ | EQ _ => Some (remove_min l x e r)#2#2
+ | GT _ => find y (remove_min l x e r)#1
+ end.
+Proof.
+ intros.
+ destruct X.compare.
+ rewrite not_find_iff; auto.
+ rewrite remove_min_in; red; destruct 1 as [H'|H']; [ order | ].
+ generalize (remove_min_gt_tree H H'); order.
+ apply find_1; auto.
+ rewrite remove_min_mapsto; auto.
+ rewrite find_mapsto_equiv; eauto; intros.
+ rewrite remove_min_mapsto; intuition; order.
Qed.
-Lemma merge_avl : forall elt (s1 s2:t elt), avl s1 -> avl s2 ->
- -(2) <= height s1 - height s2 <= 2 -> avl (merge s1 s2).
-Proof.
- intros; generalize (merge_avl_1 H H0 H1); intuition.
-Qed.
+(** * Merging two trees *)
-Lemma merge_in : forall elt (s1 s2:t elt) y, bst s1 -> avl s1 -> bst s2 -> avl s2 ->
- (In y (merge s1 s2) <-> In y s1 \/ In y s2).
+Lemma merge_in : forall m1 m2 y, bst m1 -> bst m2 ->
+ (In y (merge m1 m2) <-> In y m1 \/ In y m2).
Proof.
- intros elt s1 s2; functional induction (merge s1 s2);intros.
+ intros m1 m2; functional induction (merge m1 m2);intros;
+ try factornode _x _x0 _x1 _x2 _x3 as m1.
intuition_in.
intuition_in.
- destruct s1;try contradiction;clear y.
-(* rewrite H_eq_2; rewrite H_eq_2 in H_eq_1; clear H_eq_2. *)
- replace s2' with (fst (remove_min l2 x2 e2 r2)); [|rewrite e3; auto].
- rewrite bal_in; auto.
- generalize (remove_min_avl H2); rewrite e3; simpl; auto.
- generalize (remove_min_in y0 H2); rewrite e3; simpl; intro.
- rewrite H3; intuition.
+ rewrite bal_in, remove_min_in, e1; simpl; intuition.
Qed.
-Lemma merge_mapsto : forall elt (s1 s2:t elt) y e, bst s1 -> avl s1 -> bst s2 -> avl s2 ->
- (MapsTo y e (merge s1 s2) <-> MapsTo y e s1 \/ MapsTo y e s2).
+Lemma merge_mapsto : forall m1 m2 y e, bst m1 -> bst m2 ->
+ (MapsTo y e (merge m1 m2) <-> MapsTo y e m1 \/ MapsTo y e m2).
Proof.
- intros elt s1 s2; functional induction (@merge elt s1 s2); intros.
+ intros m1 m2; functional induction (merge m1 m2); intros;
+ try factornode _x _x0 _x1 _x2 _x3 as m1.
intuition_in.
intuition_in.
- destruct s1;try contradiction;clear y.
- replace s2' with (fst (remove_min l2 x2 e2 r2)); [|rewrite e3; auto].
- rewrite bal_mapsto; auto; unfold create.
- generalize (remove_min_avl H2); rewrite e3; simpl; auto.
- generalize (remove_min_mapsto y0 e H2); rewrite e3; simpl; intro.
- rewrite H3; intuition (try subst; auto).
- inversion_clear H3; intuition.
+ rewrite bal_mapsto, remove_min_mapsto, e1; simpl; auto.
+ unfold create.
+ intuition; subst; auto.
+ inversion_clear H1; intuition.
Qed.
-Lemma merge_bst : forall elt (s1 s2:t elt), bst s1 -> avl s1 -> bst s2 -> avl s2 ->
- (forall y1 y2 : key, In y1 s1 -> In y2 s2 -> X.lt y1 y2) ->
- bst (merge s1 s2).
+Lemma merge_bst : forall m1 m2, bst m1 -> bst m2 ->
+ (forall y1 y2 : key, In y1 m1 -> In y2 m2 -> X.lt y1 y2) ->
+ bst (merge m1 m2).
Proof.
- intros elt s1 s2; functional induction (@merge elt s1 s2); intros; auto.
-
+ intros m1 m2; functional induction (merge m1 m2); intros; auto;
+ try factornode _x _x0 _x1 _x2 _x3 as m1.
apply bal_bst; auto.
- destruct s1;try contradiction.
- generalize (remove_min_bst H1); rewrite e3; simpl in *; auto.
- destruct s1;try contradiction.
+ generalize (remove_min_bst H0); rewrite e1; simpl in *; auto.
intro; intro.
- apply H3; auto.
- generalize (remove_min_in x H2); rewrite e3; simpl; intuition.
- destruct s1;try contradiction.
- generalize (remove_min_gt_tree H1); rewrite e3; simpl; auto.
-Qed.
-
-(** * Deletion *)
-
-Function remove (elt:Set)(x:key)(s:t elt) { struct s } : t elt := match s with
- | Leaf => Leaf _
- | Node l y e r h =>
- match X.compare x y with
- | LT _ => bal (remove x l) y e r
- | EQ _ => merge l r
- | GT _ => bal l y e (remove x r)
- end
- end.
-
-Lemma remove_avl_1 : forall elt (s:t elt) x, avl s ->
- avl (remove x s) /\ 0 <= height s - height (remove x s) <= 1.
-Proof.
- intros elt s x; functional induction (@remove elt x s); intros.
- split; auto; omega_max.
- (* LT *)
- inv avl.
- destruct (IHt H0).
- split.
- apply bal_avl; auto.
- omega_max.
- omega_bal.
- (* EQ *)
- inv avl.
- generalize (merge_avl_1 H0 H1 H2).
- intuition omega_max.
- (* GT *)
- inv avl.
- destruct (IHt H1).
- split.
- apply bal_avl; auto.
- omega_max.
- omega_bal.
+ apply H1; auto.
+ generalize (remove_min_in l2 x2 d2 r2 _x4 x); rewrite e1; simpl; intuition.
+ generalize (remove_min_gt_tree H0); rewrite e1; simpl; auto.
Qed.
-Lemma remove_avl : forall elt (s:t elt) x, avl s -> avl (remove x s).
-Proof.
- intros; generalize (remove_avl_1 x H); intuition.
-Qed.
-Hint Resolve remove_avl.
+(** * Deletion *)
-Lemma remove_in : forall elt (s:t elt) x y, bst s -> avl s ->
- (In y (remove x s) <-> ~ X.eq y x /\ In y s).
+Lemma remove_in : forall m x y, bst m ->
+ (In y (remove x m) <-> ~ X.eq y x /\ In y m).
Proof.
- intros elt s x; functional induction (@remove elt x s); simpl; intros.
+ intros m x; functional induction (remove x m); simpl; intros.
intuition_in.
(* LT *)
- inv avl; inv bst; clear e1.
+ inv bst; clear e0.
rewrite bal_in; auto.
generalize (IHt y0 H0); intuition; [ order | order | intuition_in ].
(* EQ *)
- inv avl; inv bst; clear e1.
+ inv bst; clear e0.
rewrite merge_in; intuition; [ order | order | intuition_in ].
- elim H9; eauto.
+ elim H4; eauto.
(* GT *)
- inv avl; inv bst; clear e1.
+ inv bst; clear e0.
rewrite bal_in; auto.
- generalize (IHt y0 H5); intuition; [ order | order | intuition_in ].
+ generalize (IHt y0 H1); intuition; [ order | order | intuition_in ].
Qed.
-Lemma remove_bst : forall elt (s:t elt) x, bst s -> avl s -> bst (remove x s).
+Lemma remove_bst : forall m x, bst m -> bst (remove x m).
Proof.
- intros elt s x; functional induction (@remove elt x s); simpl; intros.
+ intros m x; functional induction (remove x m); simpl; intros.
auto.
(* LT *)
- inv avl; inv bst.
+ inv bst.
apply bal_bst; auto.
intro; intro.
rewrite (remove_in x y0 H0) in H; auto.
destruct H; eauto.
(* EQ *)
- inv avl; inv bst.
+ inv bst.
apply merge_bst; eauto.
(* GT *)
- inv avl; inv bst.
+ inv bst.
apply bal_bst; auto.
intro; intro.
- rewrite (remove_in x y0 H5) in H; auto.
+ rewrite (remove_in x y0 H1) in H; auto.
destruct H; eauto.
Qed.
-Lemma remove_1 : forall elt (m:t elt) x y, bst m -> avl m -> X.eq x y -> ~ In y (remove x m).
+Lemma remove_1 : forall m x y, bst m -> X.eq x y -> ~ In y (remove x m).
Proof.
intros; rewrite remove_in; intuition.
-Qed.
+Qed.
-Lemma remove_2 : forall elt (m:t elt) x y e, bst m -> avl m -> ~X.eq x y ->
+Lemma remove_2 : forall m x y e, bst m -> ~X.eq x y ->
MapsTo y e m -> MapsTo y e (remove x m).
Proof.
- intros elt m x y e; induction m; simpl; auto.
+ intros m x y e; induction m; simpl; auto.
destruct (X.compare x k);
- intros; inv bst; inv avl; try rewrite bal_mapsto; unfold create; auto;
+ intros; inv bst; try rewrite bal_mapsto; unfold create; auto;
try solve [inv MapsTo; auto].
rewrite merge_mapsto; auto.
inv MapsTo; auto; order.
Qed.
-Lemma remove_3 : forall elt (m:t elt) x y e, bst m -> avl m ->
+Lemma remove_3 : forall m x y e, bst m ->
MapsTo y e (remove x m) -> MapsTo y e m.
Proof.
- intros elt m x y e; induction m; simpl; auto.
- destruct (X.compare x k); intros Bs Av; inv avl; inv bst;
+ intros m x y e; induction m; simpl; auto.
+ destruct (X.compare x k); intros Bs; inv bst;
try rewrite bal_mapsto; auto; unfold create.
- intros; inv MapsTo; auto.
+ intros; inv MapsTo; auto.
rewrite merge_mapsto; intuition.
intros; inv MapsTo; auto.
Qed.
-Section Elt2.
+(** * join *)
-Variable elt:Set.
+Lemma join_in : forall l x d r y,
+ In y (join l x d r) <-> X.eq y x \/ In y l \/ In y r.
+Proof.
+ join_tac.
+ simpl.
+ rewrite add_in; intuition_in.
+ rewrite add_in; intuition_in.
+ rewrite bal_in, Hlr; clear Hlr Hrl; intuition_in.
+ rewrite bal_in, Hrl; clear Hlr Hrl; intuition_in.
+ apply create_in.
+Qed.
-Notation eqk := (eqk (elt:= elt)).
-Notation eqke := (eqke (elt:= elt)).
-Notation ltk := (ltk (elt:= elt)).
+Lemma join_bst : forall l x d r, bst l -> bst r ->
+ lt_tree x l -> gt_tree x r -> bst (join l x d r).
+Proof.
+ join_tac; auto; try (simpl; auto; fail); inv bst; apply bal_bst; auto;
+ clear Hrl Hlr z; intro; intros; rewrite join_in in *.
+ intuition; [ apply MX.lt_eq with x | ]; eauto.
+ intuition; [ apply MX.eq_lt with x | ]; eauto.
+Qed.
+Hint Resolve join_bst.
-(** * Empty map *)
+Lemma join_find : forall l x d r y,
+ bst l -> bst r -> lt_tree x l -> gt_tree x r ->
+ find y (join l x d r) = find y (create l x d r).
+Proof.
+ join_tac; auto; inv bst;
+ simpl (join (Leaf elt));
+ try (assert (X.lt lx x) by auto);
+ try (assert (X.lt x rx) by auto);
+ rewrite ?add_find, ?bal_find; auto.
-Definition Empty m := forall (a : key)(e:elt) , ~ MapsTo a e m.
+ simpl; destruct X.compare; auto.
+ rewrite not_find_iff; auto; intro; order.
-Definition empty := (Leaf elt).
+ simpl; repeat (destruct X.compare; auto); try (order; fail).
+ rewrite not_find_iff by auto; intro.
+ assert (X.lt y x) by auto; order.
-Lemma empty_bst : bst empty.
+ simpl; rewrite Hlr; simpl; auto.
+ repeat (destruct X.compare; auto); order.
+ intros u Hu; rewrite join_in in Hu.
+ destruct Hu as [Hu|[Hu|Hu]]; try generalize (H2 _ Hu); order.
+
+ simpl; rewrite Hrl; simpl; auto.
+ repeat (destruct X.compare; auto); order.
+ intros u Hu; rewrite join_in in Hu.
+ destruct Hu as [Hu|[Hu|Hu]]; order.
+Qed.
+
+(** * split *)
+
+Lemma split_in_1 : forall m x, bst m -> forall y,
+ (In y (split x m)#l <-> In y m /\ X.lt y x).
Proof.
- unfold empty; auto.
+ intros m x; functional induction (split x m); simpl; intros;
+ inv bst; try clear e0.
+ intuition_in.
+ rewrite e1 in IHt; simpl in IHt; rewrite IHt; intuition_in; order.
+ intuition_in; order.
+ rewrite join_in.
+ rewrite e1 in IHt; simpl in IHt; rewrite IHt; intuition_in; order.
Qed.
-Lemma empty_avl : avl empty.
+Lemma split_in_2 : forall m x, bst m -> forall y,
+ (In y (split x m)#r <-> In y m /\ X.lt x y).
Proof.
- unfold empty; auto.
+ intros m x; functional induction (split x m); subst; simpl; intros;
+ inv bst; try clear e0.
+ intuition_in.
+ rewrite join_in.
+ rewrite e1 in IHt; simpl in IHt; rewrite IHt; intuition_in; order.
+ intuition_in; order.
+ rewrite e1 in IHt; simpl in IHt; rewrite IHt; intuition_in; order.
Qed.
-Lemma empty_1 : Empty empty.
+Lemma split_in_3 : forall m x, bst m ->
+ (split x m)#o = find x m.
Proof.
- unfold empty, Empty; intuition_in.
+ intros m x; functional induction (split x m); subst; simpl; auto;
+ intros; inv bst; try clear e0;
+ destruct X.compare; try (order;fail); rewrite <-IHt, e1; auto.
Qed.
-(** * Emptyness test *)
-
-Definition is_empty (s:t elt) := match s with Leaf => true | _ => false end.
+Lemma split_bst : forall m x, bst m ->
+ bst (split x m)#l /\ bst (split x m)#r.
+Proof.
+ intros m x; functional induction (split x m); subst; simpl; intros;
+ inv bst; try clear e0; try rewrite e1 in *; simpl in *; intuition;
+ apply join_bst; auto.
+ intros y0.
+ generalize (split_in_2 x H0 y0); rewrite e1; simpl; intuition.
+ intros y0.
+ generalize (split_in_1 x H1 y0); rewrite e1; simpl; intuition.
+Qed.
-Lemma is_empty_1 : forall s, Empty s -> is_empty s = true.
+Lemma split_lt_tree : forall m x, bst m -> lt_tree x (split x m)#l.
Proof.
- destruct s as [|r x e l h]; simpl; auto.
- intro H; elim (H x e); auto.
+ intros m x B y Hy; rewrite split_in_1 in Hy; intuition.
Qed.
-Lemma is_empty_2 : forall s, is_empty s = true -> Empty s.
-Proof.
- destruct s; simpl; intros; try discriminate; red; intuition_in.
+Lemma split_gt_tree : forall m x, bst m -> gt_tree x (split x m)#r.
+Proof.
+ intros m x B y Hy; rewrite split_in_2 in Hy; intuition.
Qed.
-(** * Appartness *)
+Lemma split_find : forall m x y, bst m ->
+ find y m = match X.compare y x with
+ | LT _ => find y (split x m)#l
+ | EQ _ => (split x m)#o
+ | GT _ => find y (split x m)#r
+ end.
+Proof.
+ intros m x; functional induction (split x m); subst; simpl; intros;
+ inv bst; try clear e0; try rewrite e1 in *; simpl in *;
+ [ destruct X.compare; auto | .. ];
+ try match goal with E:split ?x ?t = _, B:bst ?t |- _ =>
+ generalize (split_in_1 x B)(split_in_2 x B)(split_bst x B);
+ rewrite E; simpl; destruct 3 end.
-(** The [mem] function is deciding appartness. It exploits the [bst] property
- to achieve logarithmic complexity. *)
+ rewrite join_find, IHt; auto; clear IHt; simpl.
+ repeat (destruct X.compare; auto); order.
+ intro y1; rewrite H4; intuition.
-Function mem (x:key)(m:t elt) { struct m } : bool :=
- match m with
- | Leaf => false
- | Node l y e r _ => match X.compare x y with
- | LT _ => mem x l
- | EQ _ => true
- | GT _ => mem x r
- end
- end.
-Implicit Arguments mem.
+ repeat (destruct X.compare; auto); order.
-Lemma mem_1 : forall s x, bst s -> In x s -> mem x s = true.
-Proof.
- intros s x.
- functional induction (mem x s); inversion_clear 1; auto.
- intuition_in.
- intuition_in; firstorder; absurd (X.lt x y); eauto.
- intuition_in; firstorder; absurd (X.lt y x); eauto.
+ rewrite join_find, IHt; auto; clear IHt; simpl.
+ repeat (destruct X.compare; auto); order.
+ intros y1; rewrite H; intuition.
Qed.
-Lemma mem_2 : forall s x, mem x s = true -> In x s.
-Proof.
- intros s x.
- functional induction (mem x s); firstorder; intros; try discriminate.
-Qed.
-
-Function find (x:key)(m:t elt) { struct m } : option elt :=
- match m with
- | Leaf => None
- | Node l y e r _ => match X.compare x y with
- | LT _ => find x l
- | EQ _ => Some e
- | GT _ => find x r
- end
- end.
+(** * Concatenation *)
-Lemma find_1 : forall m x e, bst m -> MapsTo x e m -> find x m = Some e.
-Proof.
- intros m x e.
- functional induction (find x m); inversion_clear 1; auto.
+Lemma concat_in : forall m1 m2 y,
+ In y (concat m1 m2) <-> In y m1 \/ In y m2.
+Proof.
+ intros m1 m2; functional induction (concat m1 m2); intros;
+ try factornode _x _x0 _x1 _x2 _x3 as m1.
+ intuition_in.
intuition_in.
- intuition_in; firstorder; absurd (X.lt x y); eauto.
- intuition_in; auto.
- absurd (X.lt x y); eauto.
- absurd (X.lt y x); eauto.
- intuition_in; firstorder; absurd (X.lt y x); eauto.
+ rewrite join_in, remove_min_in, e1; simpl; intuition.
Qed.
-Lemma find_2 : forall m x e, find x m = Some e -> MapsTo x e m.
-Proof.
- intros m x.
- functional induction (find x m); subst;firstorder; intros; try discriminate.
- inversion H; subst; auto.
-Qed.
-
-(** An all-in-one spec for [add] used later in the naive [map2] *)
-
-Lemma add_spec : forall m x y e , bst m -> avl m ->
- find x (add y e m) = if eq_dec x y then Some e else find x m.
-Proof.
-intros.
-destruct (eq_dec x y).
-apply find_1.
-apply add_bst; auto.
-eapply MapsTo_1 with y; eauto.
-apply add_1; auto.
-case_eq (find x m); intros.
-apply find_1.
-apply add_bst; auto.
-apply add_2; auto.
-apply find_2; auto.
-case_eq (find x (add y e m)); auto; intros.
-rewrite <- H1; symmetry.
-apply find_1; auto.
-eapply add_3; eauto.
-apply find_2; eauto.
+Lemma concat_bst : forall m1 m2, bst m1 -> bst m2 ->
+ (forall y1 y2, In y1 m1 -> In y2 m2 -> X.lt y1 y2) ->
+ bst (concat m1 m2).
+Proof.
+ intros m1 m2; functional induction (concat m1 m2); intros; auto;
+ try factornode _x _x0 _x1 _x2 _x3 as m1.
+ apply join_bst; auto.
+ change (bst (m2',xd)#1); rewrite <-e1; eauto.
+ intros y Hy.
+ apply H1; auto.
+ rewrite remove_min_in, e1; simpl; auto.
+ change (gt_tree (m2',xd)#2#1 (m2',xd)#1); rewrite <-e1; eauto.
Qed.
+Hint Resolve concat_bst.
-(** * Elements *)
+Lemma concat_find : forall m1 m2 y, bst m1 -> bst m2 ->
+ (forall y1 y2, In y1 m1 -> In y2 m2 -> X.lt y1 y2) ->
+ find y (concat m1 m2) =
+ match find y m2 with Some d => Some d | None => find y m1 end.
+Proof.
+ intros m1 m2; functional induction (concat m1 m2); intros; auto;
+ try factornode _x _x0 _x1 _x2 _x3 as m1.
+ simpl; destruct (find y m2); auto.
-(** [elements_tree_aux acc t] catenates the elements of [t] in infix
- order to the list [acc] *)
+ generalize (remove_min_find y H0)(remove_min_in l2 x2 d2 r2 _x4)
+ (remove_min_bst H0)(remove_min_gt_tree H0);
+ rewrite e1; simpl fst; simpl snd; intros.
+
+ inv bst.
+ rewrite H2, join_find; auto; clear H2.
+ simpl; destruct X.compare; simpl; auto.
+ destruct (find y m2'); auto.
+ symmetry; rewrite not_find_iff; auto; intro.
+ apply (MX.lt_not_gt l); apply H1; auto; rewrite H3; auto.
-Fixpoint elements_aux (acc : list (key*elt)) (t : t elt) {struct t} : list (key*elt) :=
- match t with
- | Leaf => acc
- | Node l x e r _ => elements_aux ((x,e) :: elements_aux acc r) l
- end.
+ intros z Hz; apply H1; auto; rewrite H3; auto.
+Qed.
-(** then [elements] is an instanciation with an empty [acc] *)
-Definition elements := elements_aux nil.
+(** * Elements *)
-Lemma elements_aux_mapsto : forall s acc x e,
+Notation eqk := (PX.eqk (elt:= elt)).
+Notation eqke := (PX.eqke (elt:= elt)).
+Notation ltk := (PX.ltk (elt:= elt)).
+
+Lemma elements_aux_mapsto : forall (s:t elt) acc x e,
InA eqke (x,e) (elements_aux acc s) <-> MapsTo x e s \/ InA eqke (x,e) acc.
Proof.
induction s as [ | l Hl x e r Hr h ]; simpl; auto.
@@ -1025,13 +1299,13 @@ Proof.
destruct H0; simpl in *; subst; intuition.
Qed.
-Lemma elements_mapsto : forall s x e, InA eqke (x,e) (elements s) <-> MapsTo x e s.
+Lemma elements_mapsto : forall (s:t elt) x e, InA eqke (x,e) (elements s) <-> MapsTo x e s.
Proof.
intros; generalize (elements_aux_mapsto s nil x e); intuition.
inversion_clear H0.
Qed.
-Lemma elements_in : forall s x, L.PX.In x (elements s) <-> In x s.
+Lemma elements_in : forall (s:t elt) x, L.PX.In x (elements s) <-> In x s.
Proof.
intros.
unfold L.PX.In.
@@ -1043,7 +1317,7 @@ Proof.
unfold L.PX.MapsTo; rewrite elements_mapsto; auto.
Qed.
-Lemma elements_aux_sort : forall s acc, bst s -> sort ltk acc ->
+Lemma elements_aux_sort : forall (s:t elt) acc, bst s -> sort ltk acc ->
(forall x e y, InA eqke (x,e) acc -> In y s -> X.lt y x) ->
sort ltk (elements_aux acc s).
Proof.
@@ -1052,7 +1326,7 @@ Proof.
apply Hl; auto.
constructor.
apply Hr; eauto.
- apply (InA_InfA (eqke_refl (elt:=elt))); intros (y',e') H6.
+ apply (InA_InfA (PX.eqke_refl (elt:=elt))); intros (y',e') H6.
destruct (elements_aux_mapsto r acc y' e'); intuition.
red; simpl; eauto.
red; simpl; eauto.
@@ -1070,20 +1344,49 @@ Proof.
Qed.
Hint Resolve elements_sort.
+Lemma elements_nodup : forall s : t elt, bst s -> NoDupA eqk (elements s).
+Proof.
+ intros; apply PX.Sort_NoDupA; auto.
+Qed.
-(** * Fold *)
+Lemma elements_aux_cardinal :
+ forall (m:t elt) acc, (length acc + cardinal m)%nat = length (elements_aux acc m).
+Proof.
+ simple induction m; simpl; intuition.
+ rewrite <- H; simpl.
+ rewrite <- H0; omega.
+Qed.
-Fixpoint fold (A : Set) (f : key -> elt -> A -> A)(s : t elt) {struct s} : A -> A :=
- fun a => match s with
- | Leaf => a
- | Node l x e r _ => fold f r (f x e (fold f l a))
- end.
+Lemma elements_cardinal : forall (m:t elt), cardinal m = length (elements m).
+Proof.
+ exact (fun m => elements_aux_cardinal m nil).
+Qed.
-Definition fold' (A : Set) (f : key -> elt -> A -> A)(s : t elt) :=
+Lemma elements_app :
+ forall (s:t elt) 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 (t1 t2:t elt) x e z l,
+ elements t1 ++ (x,e) :: elements t2 ++ l =
+ elements (Node t1 x e t2 z) ++ l.
+Proof.
+ unfold elements; simpl; intros.
+ rewrite !elements_app, <- !app_nil_end, !app_ass; auto.
+Qed.
+
+(** * Fold *)
+
+Definition fold' (A : Type) (f : key -> elt -> A -> A)(s : t elt) :=
L.fold f (elements s).
Lemma fold_equiv_aux :
- forall (A : Set) (s : t elt) (f : key -> elt -> A -> A) (a : A) acc,
+ forall (A : Type) (s : t elt) (f : key -> elt -> A -> A) (a : A) acc,
L.fold f (elements_aux acc s) a = L.fold f acc (fold f s a).
Proof.
simple induction s.
@@ -1095,7 +1398,7 @@ Proof.
Qed.
Lemma fold_equiv :
- forall (A : Set) (s : t elt) (f : key -> elt -> A -> A) (a : A),
+ forall (A : Type) (s : t elt) (f : key -> elt -> A -> A) (a : A),
fold f s a = fold' f s a.
Proof.
unfold fold', elements in |- *.
@@ -1106,8 +1409,8 @@ Proof.
Qed.
Lemma fold_1 :
- forall (s:t elt)(Hs:bst s)(A : Set)(i:A)(f : key -> elt -> A -> A),
- fold f s i = fold_left (fun a p => f (fst p) (snd p) a) (elements s) i.
+ forall (s:t elt)(Hs:bst s)(A : Type)(i:A)(f : key -> elt -> A -> A),
+ fold f s i = fold_left (fun a p => f p#1 p#2 a) (elements s) i.
Proof.
intros.
rewrite fold_equiv.
@@ -1118,288 +1421,93 @@ Qed.
(** * Comparison *)
-Definition Equal (cmp:elt->elt->bool) m m' :=
- (forall k, In k m <-> In k m') /\
- (forall k e e', MapsTo k e m -> MapsTo k e' m' -> cmp e e' = true).
-
-(** ** Enumeration of the elements of a tree *)
-
-Inductive enumeration : Set :=
- | End : enumeration
- | More : key -> elt -> t elt -> enumeration -> enumeration.
-
-(** [flatten_e e] returns the list of elements of [e] i.e. the list
- of elements actually compared *)
+(** [flatten_e e] returns the list of elements of the enumeration [e]
+ i.e. the list of elements actually compared *)
-Fixpoint flatten_e (e : enumeration) : list (key*elt) := match e with
+Fixpoint flatten_e (e : enumeration elt) : list (key*elt) := match e with
| End => nil
| More x e t r => (x,e) :: elements t ++ flatten_e r
end.
-(** [sorted_e e] expresses that elements in the enumeration [e] are
- sorted, and that all trees in [e] are binary search trees. *)
-
-Inductive In_e (p:key*elt) : enumeration -> Prop :=
- | InEHd1 :
- forall (y : key)(d:elt) (s : t elt) (e : enumeration),
- eqke p (y,d) -> In_e p (More y d s e)
- | InEHd2 :
- forall (y : key) (d:elt) (s : t elt) (e : enumeration),
- MapsTo (fst p) (snd p) s -> In_e p (More y d s e)
- | InETl :
- forall (y : key) (d:elt) (s : t elt) (e : enumeration),
- In_e p e -> In_e p (More y d s e).
-
-Hint Constructors In_e.
-
-Inductive sorted_e : enumeration -> Prop :=
- | SortedEEnd : sorted_e End
- | SortedEMore :
- forall (x : key) (d:elt) (s : t elt) (e : enumeration),
- bst s ->
- (gt_tree x s) ->
- sorted_e e ->
- (forall p, In_e p e -> ltk (x,d) p) ->
- (forall p,
- MapsTo (fst p) (snd p) s -> forall q, In_e q e -> ltk p q) ->
- sorted_e (More x d s e).
-
-Hint Constructors sorted_e.
-
-Lemma in_flatten_e :
- forall p e, InA eqke p (flatten_e e) -> In_e p e.
-Proof.
- simple induction e; simpl in |- *; intuition.
- inversion_clear H.
- inversion_clear H0; auto.
- elim (InA_app H1); auto.
- destruct (elements_mapsto t a b); auto.
+Lemma flatten_e_elements :
+ forall (l:t elt) r x d z e,
+ elements l ++ flatten_e (More x d r e) =
+ elements (Node l x d r z) ++ flatten_e e.
+Proof.
+ intros; simpl; apply elements_node.
Qed.
-Lemma sorted_flatten_e :
- forall e : enumeration, sorted_e e -> sort ltk (flatten_e e).
+Lemma cons_1 : forall (s:t elt) e,
+ flatten_e (cons s e) = elements s ++ flatten_e e.
Proof.
- simple induction e; simpl in |- *; intuition.
- apply cons_sort.
- apply (SortA_app (eqke_refl (elt:=elt))); inversion_clear H0; auto.
- intros; apply H5; auto.
- rewrite <- elements_mapsto; auto; destruct x; auto.
- apply in_flatten_e; auto.
- inversion_clear H0.
- apply In_InfA; intros.
- intros; elim (in_app_or _ _ _ H0); intuition.
- generalize (In_InA (eqke_refl (elt:=elt)) H6).
- destruct y; rewrite elements_mapsto; eauto.
- apply H4; apply in_flatten_e; auto.
- apply In_InA; auto.
+ induction s; simpl; auto; intros.
+ rewrite IHs1; apply flatten_e_elements; auto.
Qed.
-Lemma elements_app :
- forall s acc, elements_aux acc s = elements s ++ acc.
+(** Proof of correction for the comparison *)
+
+Variable cmp : elt->elt->bool.
+
+Definition IfEq b l1 l2 := L.equal cmp l1 l2 = b.
+
+Lemma cons_IfEq : forall b x1 x2 d1 d2 l1 l2,
+ X.eq x1 x2 -> cmp d1 d2 = true ->
+ IfEq b l1 l2 ->
+ IfEq b ((x1,d1)::l1) ((x2,d2)::l2).
Proof.
- simple induction s; simpl in |- *; intuition.
- rewrite H0.
- rewrite H.
- unfold elements; simpl.
- do 2 rewrite H.
- rewrite H0.
- repeat rewrite <- app_nil_end.
- repeat rewrite app_ass; auto.
+ unfold IfEq; destruct b; simpl; intros; destruct X.compare; simpl;
+ try rewrite H0; auto; order.
Qed.
-Lemma compare_flatten_1 :
- forall t1 t2 x e z l,
- elements t1 ++ (x,e) :: elements t2 ++ l =
- elements (Node t1 x e t2 z) ++ l.
+Lemma equal_end_IfEq : forall e2,
+ IfEq (equal_end e2) nil (flatten_e e2).
Proof.
- simpl in |- *; unfold elements in |- *; simpl in |- *; intuition.
- repeat rewrite elements_app.
- repeat rewrite <- app_nil_end.
- repeat rewrite app_ass; auto.
+ destruct e2; red; auto.
Qed.
-(** key lemma for correctness *)
-
-Lemma flatten_e_elements :
- forall l r x d z e,
- elements l ++ flatten_e (More x d r e) =
- elements (Node l x d r z) ++ flatten_e e.
+Lemma equal_more_IfEq :
+ forall x1 d1 (cont:enumeration elt -> bool) x2 d2 r2 e2 l,
+ IfEq (cont (cons r2 e2)) l (elements r2 ++ flatten_e e2) ->
+ IfEq (equal_more cmp x1 d1 cont (More x2 d2 r2 e2)) ((x1,d1)::l)
+ (flatten_e (More x2 d2 r2 e2)).
Proof.
- intros; simpl.
- apply compare_flatten_1.
+ unfold IfEq; simpl; intros; destruct X.compare; simpl; auto.
+ rewrite <-andb_lazy_alt; f_equal; auto.
Qed.
-Open Local Scope Z_scope.
-
-(** termination of [compare_aux] *)
-
-Fixpoint measure_e_t (s : t elt) : Z := match s with
- | Leaf => 0
- | Node l _ _ r _ => 1 + measure_e_t l + measure_e_t r
- end.
-
-Fixpoint measure_e (e : enumeration) : Z := match e with
- | End => 0
- | More _ _ s r => 1 + measure_e_t s + measure_e r
- end.
-
-Ltac Measure_e_t := unfold measure_e_t in |- *; fold measure_e_t in |- *.
-Ltac Measure_e := unfold measure_e in |- *; fold measure_e in |- *.
-
-Lemma measure_e_t_0 : forall s : t elt, measure_e_t s >= 0.
+Lemma equal_cont_IfEq : forall m1 cont e2 l,
+ (forall e, IfEq (cont e) l (flatten_e e)) ->
+ IfEq (equal_cont cmp m1 cont e2) (elements m1 ++ l) (flatten_e e2).
Proof.
- simple induction s.
- simpl in |- *; omega.
- intros.
- Measure_e_t; omega.
+ induction m1 as [|l1 Hl1 x1 d1 r1 Hr1 h1]; simpl; intros; auto.
+ rewrite <- elements_node; simpl.
+ apply Hl1; auto.
+ clear e2; intros [|x2 d2 r2 e2].
+ simpl; red; auto.
+ apply equal_more_IfEq.
+ rewrite <- cons_1; auto.
Qed.
-Ltac Measure_e_t_0 s := generalize (@measure_e_t_0 s); intro.
-
-Lemma measure_e_0 : forall e : enumeration, measure_e e >= 0.
+Lemma equal_IfEq : forall (m1 m2:t elt),
+ IfEq (equal cmp m1 m2) (elements m1) (elements m2).
Proof.
- simple induction e.
- simpl in |- *; omega.
+ intros; unfold equal.
+ rewrite (app_nil_end (elements m1)).
+ replace (elements m2) with (flatten_e (cons m2 (End _)))
+ by (rewrite cons_1; simpl; rewrite <-app_nil_end; auto).
+ apply equal_cont_IfEq.
intros.
- Measure_e; Measure_e_t_0 t; omega.
+ apply equal_end_IfEq; auto.
Qed.
-Ltac Measure_e_0 e := generalize (@measure_e_0 e); intro.
-
-(** Induction principle over the sum of the measures for two lists *)
+Definition Equivb m m' :=
+ (forall k, In k m <-> In k m') /\
+ (forall k e e', MapsTo k e m -> MapsTo k e' m' -> cmp e e' = true).
-Definition compare_rec2 :
- forall P : enumeration -> enumeration -> Set,
- (forall x x' : enumeration,
- (forall y y' : enumeration,
- measure_e y + measure_e y' < measure_e x + measure_e x' -> P y y') ->
- P x x') ->
- forall x x' : enumeration, P x x'.
+Lemma Equivb_elements : forall s s',
+ Equivb s s' <-> L.Equivb cmp (elements s) (elements s').
Proof.
- intros P H x x'.
- apply well_founded_induction_type_2
- with (R := fun yy' xx' : enumeration * enumeration =>
- measure_e (fst yy') + measure_e (snd yy') <
- measure_e (fst xx') + measure_e (snd xx')); auto.
- apply Wf_nat.well_founded_lt_compat
- with (f := fun xx' : enumeration * enumeration =>
- Zabs_nat (measure_e (fst xx') + measure_e (snd xx'))).
- intros; apply Zabs.Zabs_nat_lt.
- Measure_e_0 (fst x0); Measure_e_0 (snd x0); Measure_e_0 (fst y);
- Measure_e_0 (snd y); intros; omega.
-Qed.
-
-(** [cons t e] adds the elements of tree [t] on the head of
- enumeration [e]. Code:
-
-let rec cons s e = match s with
- | Empty -> e
- | Node(l, k, d, r, _) -> cons l (More(k, d, r, e))
-*)
-
-Definition cons : forall s e, bst s -> sorted_e e ->
- (forall x y, MapsTo (fst x) (snd x) s -> In_e y e -> ltk x y) ->
- { r : enumeration
- | sorted_e r /\
- measure_e r = measure_e_t s + measure_e e /\
- flatten_e r = elements s ++ flatten_e e
- }.
-Proof.
- simple induction s; intuition.
- (* s = Leaf *)
- exists e; intuition.
- (* s = Node t k e t0 z *)
- clear H0.
- case (H (More k e t0 e0)); clear H; intuition.
- inv bst; auto.
- constructor; inversion_clear H1; auto.
- inversion_clear H0; inv bst; intuition.
- destruct y; red; red in H4; simpl in *; intuition.
- apply lt_eq with k; eauto.
- destruct y; red; simpl in *; intuition.
- apply X.lt_trans with k; eauto.
- exists x; intuition.
- generalize H4; Measure_e; intros; Measure_e_t; omega.
- rewrite H5.
- apply flatten_e_elements.
-Qed.
-
-Definition equal_aux :
- forall (cmp: elt -> elt -> bool)(e1 e2:enumeration),
- sorted_e e1 -> sorted_e e2 ->
- { L.Equal cmp (flatten_e e1) (flatten_e e2) } +
- { ~ L.Equal cmp (flatten_e e1) (flatten_e e2) }.
-Proof.
- intros cmp e1 e2; pattern e1, e2 in |- *; apply compare_rec2.
- simple destruct x; simple destruct x'; intuition.
- (* x = x' = End *)
- left; unfold L.Equal in |- *; intuition.
- inversion H2.
- (* x = End x' = More *)
- right; simpl in |- *; auto.
- destruct 1.
- destruct (H2 k).
- destruct H5; auto.
- exists e; auto.
- inversion H5.
- (* x = More x' = End *)
- right; simpl in |- *; auto.
- destruct 1.
- destruct (H2 k).
- destruct H4; auto.
- exists e; auto.
- inversion H4.
- (* x = More k e t e0, x' = More k0 e3 t0 e4 *)
- case (X.compare k k0); intro.
- (* k < k0 *)
- right.
- destruct 1.
- clear H3 H.
- assert (L.PX.In k (flatten_e (More k0 e3 t0 e4))).
- destruct (H2 k).
- apply H; simpl; exists e; auto.
- destruct H.
- generalize (Sort_In_cons_2 (sorted_flatten_e H1) (InA_eqke_eqk H)).
- compute.
- intuition order.
- (* k = k0 *)
- case_eq (cmp e e3).
- intros EQ.
- destruct (@cons t e0) as [c1 (H2,(H3,H4))]; try inversion_clear H0; auto.
- destruct (@cons t0 e4) as [c2 (H5,(H6,H7))]; try inversion_clear H1; auto.
- destruct (H c1 c2); clear H; intuition.
- Measure_e; omega.
- left.
- rewrite H4 in e6; rewrite H7 in e6.
- simpl; rewrite <- L.equal_cons; auto.
- apply (sorted_flatten_e H0).
- apply (sorted_flatten_e H1).
- right.
- simpl; rewrite <- L.equal_cons; auto.
- apply (sorted_flatten_e H0).
- apply (sorted_flatten_e H1).
- swap f.
- rewrite H4; rewrite H7; auto.
- right.
- destruct 1.
- rewrite (H4 k) in H2; try discriminate; simpl; auto.
- (* k > k0 *)
- right.
- destruct 1.
- clear H3 H.
- assert (L.PX.In k0 (flatten_e (More k e t e0))).
- destruct (H2 k0).
- apply H3; simpl; exists e3; auto.
- destruct H.
- generalize (Sort_In_cons_2 (sorted_flatten_e H0) (InA_eqke_eqk H)).
- compute.
- intuition order.
-Qed.
-
-Lemma Equal_elements : forall cmp s s',
- Equal cmp s s' <-> L.Equal cmp (elements s) (elements s').
-Proof.
-unfold Equal, L.Equal; split; split; intros.
+unfold Equivb, L.Equivb; split; split; intros.
do 2 rewrite elements_in; firstorder.
destruct H.
apply (H2 k); rewrite <- elements_mapsto; auto.
@@ -1408,95 +1516,46 @@ destruct H.
apply (H2 k); unfold L.PX.MapsTo; rewrite elements_mapsto; auto.
Qed.
-Definition equal : forall cmp s s', bst s -> bst s' ->
- {Equal cmp s s'} + {~ Equal cmp s s'}.
+Lemma equal_Equivb : forall (s s': t elt), bst s -> bst s' ->
+ (equal cmp s s' = true <-> Equivb s s').
Proof.
- intros cmp s1 s2 s1_bst s2_bst; simpl.
- destruct (@cons s1 End); auto.
- inversion_clear 2.
- destruct (@cons s2 End); auto.
- inversion_clear 2.
- simpl in a; rewrite <- app_nil_end in a.
- simpl in a0; rewrite <- app_nil_end in a0.
- destruct (@equal_aux cmp x x0); intuition.
- left.
- rewrite H4 in e; rewrite H5 in e.
- rewrite Equal_elements; auto.
- right.
- swap n.
- rewrite H4; rewrite H5.
- rewrite <- Equal_elements; auto.
+ intros s s' B B'.
+ rewrite Equivb_elements, <- equal_IfEq.
+ split; [apply L.equal_2|apply L.equal_1]; auto.
Qed.
-End Elt2.
-
-Section Elts.
-
-Variable elt elt' elt'' : Set.
+End Elt.
Section Map.
+Variable elt elt' : Type.
Variable f : elt -> elt'.
-Fixpoint map (m:t elt) {struct m} : t elt' :=
- match m with
- | Leaf => Leaf _
- | Node l v d r h => Node (map l) v (f d) (map r) h
- end.
-
-Lemma map_height : forall m, height (map m) = height m.
-Proof.
-destruct m; simpl; auto.
-Qed.
-
-Lemma map_avl : forall m, avl m -> avl (map m).
-Proof.
-induction m; simpl; auto.
-inversion_clear 1; constructor; auto; do 2 rewrite map_height; auto.
-Qed.
-
-Lemma map_1 : forall (m: tree elt)(x:key)(e:elt),
- MapsTo x e m -> MapsTo x (f e) (map m).
+Lemma map_1 : forall (m: t elt)(x:key)(e:elt),
+ MapsTo x e m -> MapsTo x (f e) (map f m).
Proof.
induction m; simpl; inversion_clear 1; auto.
Qed.
Lemma map_2 : forall (m: t elt)(x:key),
- In x (map m) -> In x m.
+ In x (map f m) -> In x m.
Proof.
induction m; simpl; inversion_clear 1; auto.
Qed.
-Lemma map_bst : forall m, bst m -> bst (map m).
+Lemma map_bst : forall m, bst m -> bst (map f m).
Proof.
induction m; simpl; auto.
-inversion_clear 1; constructor; auto.
-red; intros; apply H2; apply map_2; auto.
-red; intros; apply H3; apply map_2; auto.
+inversion_clear 1; constructor; auto;
+ red; auto using map_2.
Qed.
End Map.
-Section Mapi.
-Variable f : key -> elt -> elt'.
-
-Fixpoint mapi (m:t elt) {struct m} : t elt' :=
- match m with
- | Leaf => Leaf _
- | Node l v d r h => Node (mapi l) v (f v d) (mapi r) h
- end.
-
-Lemma mapi_height : forall m, height (mapi m) = height m.
-Proof.
-destruct m; simpl; auto.
-Qed.
-
-Lemma mapi_avl : forall m, avl m -> avl (mapi m).
-Proof.
-induction m; simpl; auto.
-inversion_clear 1; constructor; auto; do 2 rewrite mapi_height; auto.
-Qed.
+Section Mapi.
+Variable elt elt' : Type.
+Variable f : key -> elt -> elt'.
Lemma mapi_1 : forall (m: tree elt)(x:key)(e:elt),
- MapsTo x e m -> exists y, X.eq y x /\ MapsTo x (f y e) (mapi m).
+ MapsTo x e m -> exists y, X.eq y x /\ MapsTo x (f y e) (mapi f m).
Proof.
induction m; simpl; inversion_clear 1; auto.
exists k; auto.
@@ -1507,198 +1566,242 @@ exists x0; intuition.
Qed.
Lemma mapi_2 : forall (m: t elt)(x:key),
- In x (mapi m) -> In x m.
+ In x (mapi f m) -> In x m.
Proof.
induction m; simpl; inversion_clear 1; auto.
Qed.
-Lemma mapi_bst : forall m, bst m -> bst (mapi m).
+Lemma mapi_bst : forall m, bst m -> bst (mapi f m).
Proof.
induction m; simpl; auto.
-inversion_clear 1; constructor; auto.
-red; intros; apply H2; apply mapi_2; auto.
-red; intros; apply H3; apply mapi_2; auto.
+inversion_clear 1; constructor; auto;
+ red; auto using mapi_2.
Qed.
End Mapi.
-Section Map2.
-Variable f : option elt -> option elt' -> option elt''.
-
-(* Not exactly pretty nor perfect, but should suffice as a first naive implem.
- Anyway, map2 isn't in Ocaml...
-*)
-
-Definition anti_elements (l:list (key*elt'')) := L.fold (@add _) l (empty _).
+Section Map_option.
+Variable elt elt' : Type.
+Variable f : key -> elt -> option elt'.
+Hypothesis f_compat : forall x x' d, X.eq x x' -> f x d = f x' d.
-Definition map2 (m:t elt)(m':t elt') : t elt'' :=
- anti_elements (L.map2 f (elements m) (elements m')).
-
-Lemma anti_elements_avl_aux : forall (l:list (key*elt''))(m:t elt''),
- avl m -> avl (L.fold (@add _) l m).
+Lemma map_option_2 : forall (m:t elt)(x:key),
+ In x (map_option f m) -> exists d, MapsTo x d m /\ f x d <> None.
Proof.
-unfold anti_elements; induction l.
-simpl; auto.
-simpl; destruct a; intros.
-apply IHl.
-apply add_avl; auto.
+intros m; functional induction (map_option f m); simpl; auto; intros.
+inversion H.
+rewrite join_in in H; destruct H as [H|[H|H]].
+exists d; split; auto; rewrite (f_compat d H), e0; discriminate.
+destruct (IHt _ H) as (d0 & ? & ?); exists d0; auto.
+destruct (IHt0 _ H) as (d0 & ? & ?); exists d0; auto.
+rewrite concat_in in H; destruct H as [H|H].
+destruct (IHt _ H) as (d0 & ? & ?); exists d0; auto.
+destruct (IHt0 _ H) as (d0 & ? & ?); exists d0; auto.
Qed.
-Lemma anti_elements_avl : forall l, avl (anti_elements l).
+Lemma map_option_bst : forall m, bst m -> bst (map_option f m).
Proof.
-unfold anti_elements, empty; intros; apply anti_elements_avl_aux; auto.
-Qed.
+intros m; functional induction (map_option f m); simpl; auto; intros;
+ inv bst.
+apply join_bst; auto; intros y H;
+ destruct (map_option_2 H) as (d0 & ? & ?); eauto using MapsTo_In.
+apply concat_bst; auto; intros y y' H H'.
+destruct (map_option_2 H) as (d0 & ? & ?).
+destruct (map_option_2 H') as (d0' & ? & ?).
+eapply X.lt_trans with x; eauto using MapsTo_In.
+Qed.
+Hint Resolve map_option_bst.
+
+Ltac nonify e :=
+ replace e with (@None elt) by
+ (symmetry; rewrite not_find_iff; auto; intro; order).
+
+Lemma map_option_find : forall (m:t elt)(x:key),
+ bst m ->
+ find x (map_option f m) =
+ match (find x m) with Some d => f x d | None => None end.
+Proof.
+intros m; functional induction (map_option f m); simpl; auto; intros;
+ inv bst; rewrite join_find || rewrite concat_find; auto; simpl;
+ try destruct X.compare; simpl; auto.
+rewrite (f_compat d e); auto.
+intros y H;
+ destruct (map_option_2 H) as (? & ? & ?); eauto using MapsTo_In.
+intros y H;
+ destruct (map_option_2 H) as (? & ? & ?); eauto using MapsTo_In.
+
+rewrite <- IHt, IHt0; auto; nonify (find x0 r); auto.
+rewrite IHt, IHt0; auto; nonify (find x0 r); nonify (find x0 l); auto.
+rewrite (f_compat d e); auto.
+rewrite <- IHt0, IHt; auto; nonify (find x0 l); auto.
+ destruct (find x0 (map_option f r)); auto.
+
+intros y y' H H'.
+destruct (map_option_2 H) as (? & ? & ?).
+destruct (map_option_2 H') as (? & ? & ?).
+eapply X.lt_trans with x; eauto using MapsTo_In.
+Qed.
+
+End Map_option.
+
+Section Map2_opt.
+Variable elt elt' elt'' : Type.
+Variable f0 : key -> option elt -> option elt' -> option elt''.
+Variable f : key -> elt -> option elt' -> option elt''.
+Variable mapl : t elt -> t elt''.
+Variable mapr : t elt' -> t elt''.
+Hypothesis f0_f : forall x d o, f x d o = f0 x (Some d) o.
+Hypothesis mapl_bst : forall m, bst m -> bst (mapl m).
+Hypothesis mapr_bst : forall m', bst m' -> bst (mapr m').
+Hypothesis mapl_f0 : forall x m, bst m ->
+ find x (mapl m) =
+ match find x m with Some d => f0 x (Some d) None | None => None end.
+Hypothesis mapr_f0 : forall x m', bst m' ->
+ find x (mapr m') =
+ match find x m' with Some d' => f0 x None (Some d') | None => None end.
+Hypothesis f0_compat : forall x x' o o', X.eq x x' -> f0 x o o' = f0 x' o o'.
+
+Notation map2_opt := (map2_opt f mapl mapr).
+
+Lemma map2_opt_2 : forall m m' y, bst m -> bst m' ->
+ In y (map2_opt m m') -> In y m \/ In y m'.
+Proof.
+intros m m'; functional induction (map2_opt m m'); intros;
+ auto; try factornode _x0 _x1 _x2 _x3 _x4 as m2;
+ try (generalize (split_in_1 x1 H0 y)(split_in_2 x1 H0 y)
+ (split_bst x1 H0); rewrite e1; simpl; destruct 3; inv bst).
+
+right; apply find_in.
+generalize (in_find (mapr_bst H0) H1); rewrite mapr_f0; auto.
+destruct (find y m2); auto; intros; discriminate.
+
+factornode l1 x1 d1 r1 _x as m1.
+left; apply find_in.
+generalize (in_find (mapl_bst H) H1); rewrite mapl_f0; auto.
+destruct (find y m1); auto; intros; discriminate.
+
+rewrite join_in in H1; destruct H1 as [H'|[H'|H']]; auto.
+destruct (IHt1 y H6 H4 H'); intuition.
+destruct (IHt0 y H7 H5 H'); intuition.
+
+rewrite concat_in in H1; destruct H1 as [H'|H']; auto.
+destruct (IHt1 y H6 H4 H'); intuition.
+destruct (IHt0 y H7 H5 H'); intuition.
+Qed.
+
+Lemma map2_opt_bst : forall m m', bst m -> bst m' ->
+ bst (map2_opt m m').
+Proof.
+intros m m'; functional induction (map2_opt m m'); intros;
+ auto; try factornode _x0 _x1 _x2 _x3 _x4 as m2; inv bst;
+ generalize (split_in_1 x1 H0)(split_in_2 x1 H0)(split_bst x1 H0);
+ rewrite e1; simpl in *; destruct 3.
+
+apply join_bst; auto.
+intros y Hy; specialize H with y.
+destruct (map2_opt_2 H1 H6 Hy); intuition.
+intros y Hy; specialize H5 with y.
+destruct (map2_opt_2 H2 H7 Hy); intuition.
+
+apply concat_bst; auto.
+intros y y' Hy Hy'; specialize H with y; specialize H5 with y'.
+apply X.lt_trans with x1.
+destruct (map2_opt_2 H1 H6 Hy); intuition.
+destruct (map2_opt_2 H2 H7 Hy'); intuition.
+Qed.
+Hint Resolve map2_opt_bst.
+
+Ltac map2_aux :=
+ match goal with
+ | H : In ?x _ \/ In ?x ?m,
+ H' : find ?x ?m = find ?x ?m', B:bst ?m, B':bst ?m' |- _ =>
+ destruct H; [ intuition_in; order |
+ rewrite <-(find_in_equiv B B' H'); auto ]
+ end.
-Lemma anti_elements_bst_aux : forall (l:list (key*elt''))(m:t elt''),
- bst m -> avl m -> bst (L.fold (@add _) l m).
-Proof.
-induction l.
-simpl; auto.
-simpl; destruct a; intros.
-apply IHl.
-apply add_bst; auto.
-apply add_avl; auto.
-Qed.
+Ltac nonify t :=
+ match t with (find ?y (map2_opt ?m ?m')) =>
+ replace t with (@None elt'');
+ [ | symmetry; rewrite not_find_iff; auto; intro;
+ destruct (@map2_opt_2 m m' y); auto; order ]
+ end.
-Lemma anti_elements_bst : forall l, bst (anti_elements l).
-Proof.
-unfold anti_elements, empty; intros; apply anti_elements_bst_aux; auto.
-Qed.
+Lemma map2_opt_1 : forall m m' y, bst m -> bst m' ->
+ In y m \/ In y m' ->
+ find y (map2_opt m m') = f0 y (find y m) (find y m').
+Proof.
+intros m m'; functional induction (map2_opt m m'); intros;
+ auto; try factornode _x0 _x1 _x2 _x3 _x4 as m2;
+ try (generalize (split_in_1 x1 H0)(split_in_2 x1 H0)
+ (split_in_3 x1 H0)(split_bst x1 H0)(split_find x1 y H0)
+ (split_lt_tree (x:=x1) H0)(split_gt_tree (x:=x1) H0);
+ rewrite e1; simpl in *; destruct 4; intros; inv bst;
+ subst o2; rewrite H7, ?join_find, ?concat_find; auto).
+
+simpl; destruct H1; [ inversion_clear H1 | ].
+rewrite mapr_f0; auto.
+generalize (in_find H0 H1); destruct (find y m2); intuition.
+
+factornode l1 x1 d1 r1 _x as m1.
+destruct H1; [ | inversion_clear H1 ].
+rewrite mapl_f0; auto.
+generalize (in_find H H1); destruct (find y m1); intuition.
+
+simpl; destruct X.compare; auto.
+apply IHt1; auto; map2_aux.
+rewrite (@f0_compat y x1), <- f0_f; auto.
+apply IHt0; auto; map2_aux.
+intros z Hz; destruct (@map2_opt_2 l1 l2' z); auto.
+intros z Hz; destruct (@map2_opt_2 r1 r2' z); auto.
+
+destruct X.compare.
+nonify (find y (map2_opt r1 r2')).
+apply IHt1; auto; map2_aux.
+nonify (find y (map2_opt r1 r2')).
+nonify (find y (map2_opt l1 l2')).
+rewrite (@f0_compat y x1), <- f0_f; auto.
+nonify (find y (map2_opt l1 l2')).
+rewrite IHt0; auto; [ | map2_aux ].
+destruct (f0 y (find y r1) (find y r2')); auto.
+intros y1 y2 Hy1 Hy2; apply X.lt_trans with x1.
+ destruct (@map2_opt_2 l1 l2' y1); auto.
+ destruct (@map2_opt_2 r1 r2' y2); auto.
+Qed.
+
+End Map2_opt.
-Lemma anti_elements_mapsto_aux : forall (l:list (key*elt'')) m k e,
- bst m -> avl m -> NoDupA (eqk (elt:=elt'')) l ->
- (forall x, L.PX.In x l -> In x m -> False) ->
- (MapsTo k e (L.fold (@add _) l m) <-> L.PX.MapsTo k e l \/ MapsTo k e m).
-Proof.
-induction l.
-simpl; auto.
-intuition.
-inversion H4.
-simpl; destruct a; intros.
-rewrite IHl; clear IHl.
-apply add_bst; auto.
-apply add_avl; auto.
-inversion H1; auto.
-intros.
-inversion_clear H1.
-assert (~X.eq x k).
- swap H5.
- destruct H3.
- apply InA_eqA with (x,x0); eauto.
-apply (H2 x).
-destruct H3; exists x0; auto.
-revert H4; do 2 rewrite <- In_alt; destruct 1; exists x0; auto.
-eapply add_3; eauto.
-intuition.
-assert (find k0 (add k e m) = Some e0).
- apply find_1; auto.
- apply add_bst; auto.
-clear H4.
-rewrite add_spec in H3; auto.
-destruct (eq_dec k0 k).
-inversion_clear H3; subst; auto.
-right; apply find_2; auto.
-inversion_clear H4; auto.
-compute in H3; destruct H3.
-subst; right; apply add_1; auto.
-inversion_clear H1.
-destruct (eq_dec k0 k).
-destruct (H2 k); eauto.
-right; apply add_2; auto.
-Qed.
-
-Lemma anti_elements_mapsto : forall l k e, NoDupA (eqk (elt:=elt'')) l ->
- (MapsTo k e (anti_elements l) <-> L.PX.MapsTo k e l).
-Proof.
-intros.
-unfold anti_elements.
-rewrite anti_elements_mapsto_aux; auto; unfold empty; auto.
-inversion 2.
-intuition.
-inversion H1.
-Qed.
+Section Map2.
+Variable elt elt' elt'' : Type.
+Variable f : option elt -> option elt' -> option elt''.
-Lemma map2_avl : forall (m: t elt)(m': t elt'), avl (map2 m m').
+Lemma map2_bst : forall m m', bst m -> bst m' -> bst (map2 f m m').
Proof.
-unfold map2; intros; apply anti_elements_avl; auto.
+unfold map2; intros.
+apply map2_opt_bst with (fun _ => f); auto using map_option_bst;
+ intros; rewrite map_option_find; auto.
Qed.
-Lemma map2_bst : forall (m: t elt)(m': t elt'), bst (map2 m m').
+Lemma map2_1 : forall m m' y, bst m -> bst m' ->
+ In y m \/ In y m' -> find y (map2 f m m') = f (find y m) (find y m').
Proof.
-unfold map2; intros; apply anti_elements_bst; auto.
-Qed.
-
-Lemma find_elements : forall (elt:Set)(m: t elt) x, bst m ->
- L.find x (elements m) = find x m.
-Proof.
-intros.
-case_eq (find x m); intros.
-apply L.find_1.
-apply elements_sort; auto.
-red; rewrite elements_mapsto.
-apply find_2; auto.
-case_eq (L.find x (elements m)); auto; intros.
-rewrite <- H0; symmetry.
-apply find_1; auto.
-rewrite <- elements_mapsto.
-apply L.find_2; auto.
-Qed.
-
-Lemma find_anti_elements : forall (l: list (key*elt'')) x, sort (@ltk _) l ->
- find x (anti_elements l) = L.find x l.
-Proof.
-intros.
-case_eq (L.find x l); intros.
-apply find_1.
-apply anti_elements_bst; auto.
-rewrite anti_elements_mapsto.
-apply L.PX.Sort_NoDupA; auto.
-apply L.find_2; auto.
-case_eq (find x (anti_elements l)); auto; intros.
-rewrite <- H0; symmetry.
-apply L.find_1; auto.
-rewrite <- anti_elements_mapsto.
-apply L.PX.Sort_NoDupA; auto.
-apply find_2; auto.
-Qed.
-
-Lemma map2_1 : forall (m: t elt)(m': t elt')(x:key), bst m -> bst m' ->
- In x m \/ In x m' -> find x (map2 m m') = f (find x m) (find x m').
-Proof.
unfold map2; intros.
-rewrite find_anti_elements; auto.
-rewrite <- find_elements; auto.
-rewrite <- find_elements; auto.
-apply L.map2_1; auto.
-apply elements_sort; auto.
-apply elements_sort; auto.
-do 2 rewrite elements_in; auto.
-apply L.map2_sorted; auto.
-apply elements_sort; auto.
-apply elements_sort; auto.
+rewrite (map2_opt_1 (f0:=fun _ => f));
+ auto using map_option_bst; intros; rewrite map_option_find; auto.
Qed.
-Lemma map2_2 : forall (m: t elt)(m': t elt')(x:key), bst m -> bst m' ->
- In x (map2 m m') -> In x m \/ In x m'.
+Lemma map2_2 : forall m m' y, bst m -> bst m' ->
+ In y (map2 f m m') -> In y m \/ In y m'.
Proof.
unfold map2; intros.
-do 2 rewrite <- elements_in.
-apply L.map2_2 with (f:=f); auto.
-apply elements_sort; auto.
-apply elements_sort; auto.
-revert H1.
-rewrite <- In_alt.
-destruct 1.
-exists x0.
-rewrite <- anti_elements_mapsto; auto.
-apply L.PX.Sort_NoDupA; auto.
-apply L.map2_sorted; auto.
-apply elements_sort; auto.
-apply elements_sort; auto.
+eapply map2_opt_2 with (f0:=fun _ => f); eauto; intros.
+ apply map_option_bst; auto.
+ apply map_option_bst; auto.
+ rewrite map_option_find; auto.
+ rewrite map_option_find; auto.
Qed.
End Map2.
-End Elts.
+End Proofs.
End Raw.
(** * Encapsulation
@@ -1710,178 +1813,184 @@ Module IntMake (I:Int)(X: OrderedType) <: S with Module E := X.
Module E := X.
Module Raw := Raw I X.
+ Import Raw.Proofs.
- Record bbst (elt:Set) : Set :=
- Bbst {this :> Raw.tree elt; is_bst : Raw.bst this; is_avl: Raw.avl this}.
+ Record bst (elt:Type) :=
+ Bst {this :> Raw.tree elt; is_bst : Raw.bst this}.
- Definition t := bbst.
+ Definition t := bst.
Definition key := E.t.
Section Elt.
- Variable elt elt' elt'': Set.
+ Variable elt elt' elt'': Type.
Implicit Types m : t elt.
Implicit Types x y : key.
Implicit Types e : elt.
- Definition empty : t elt := Bbst (Raw.empty_bst elt) (Raw.empty_avl elt).
+ Definition empty : t elt := Bst (empty_bst elt).
Definition is_empty m : bool := Raw.is_empty m.(this).
- Definition add x e m : t elt :=
- Bbst (Raw.add_bst x e m.(is_bst) m.(is_avl)) (Raw.add_avl x e m.(is_avl)).
- Definition remove x m : t elt :=
- Bbst (Raw.remove_bst x m.(is_bst) m.(is_avl)) (Raw.remove_avl x m.(is_avl)).
+ Definition add x e m : t elt := Bst (add_bst x e m.(is_bst)).
+ Definition remove x m : t elt := Bst (remove_bst x m.(is_bst)).
Definition mem x m : bool := Raw.mem x m.(this).
Definition find x m : option elt := Raw.find x m.(this).
- Definition map f m : t elt' :=
- Bbst (Raw.map_bst f m.(is_bst)) (Raw.map_avl f m.(is_avl)).
+ Definition map f m : t elt' := Bst (map_bst f m.(is_bst)).
Definition mapi (f:key->elt->elt') m : t elt' :=
- Bbst (Raw.mapi_bst f m.(is_bst)) (Raw.mapi_avl f m.(is_avl)).
+ Bst (mapi_bst f m.(is_bst)).
Definition map2 f m (m':t elt') : t elt'' :=
- Bbst (Raw.map2_bst f m m') (Raw.map2_avl f m m').
+ Bst (map2_bst f m.(is_bst) m'.(is_bst)).
Definition elements m : list (key*elt) := Raw.elements m.(this).
- Definition fold (A:Set) (f:key->elt->A->A) m i := Raw.fold (A:=A) f m.(this) i.
- Definition equal cmp m m' : bool :=
- if (Raw.equal cmp m.(is_bst) m'.(is_bst)) then true else false.
+ Definition cardinal m := Raw.cardinal m.(this).
+ Definition fold (A:Type) (f:key->elt->A->A) m i := Raw.fold (A:=A) f m.(this) i.
+ Definition equal cmp m m' : bool := Raw.equal cmp m.(this) m'.(this).
Definition MapsTo x e m : Prop := Raw.MapsTo x e m.(this).
Definition In x m : Prop := Raw.In0 x m.(this).
- Definition Empty m : Prop := Raw.Empty m.(this).
+ Definition Empty m : Prop := Empty m.(this).
- Definition eq_key : (key*elt) -> (key*elt) -> Prop := @Raw.PX.eqk elt.
- Definition eq_key_elt : (key*elt) -> (key*elt) -> Prop := @Raw.PX.eqke elt.
- Definition lt_key : (key*elt) -> (key*elt) -> Prop := @Raw.PX.ltk elt.
+ Definition eq_key : (key*elt) -> (key*elt) -> Prop := @PX.eqk elt.
+ Definition eq_key_elt : (key*elt) -> (key*elt) -> Prop := @PX.eqke elt.
+ Definition lt_key : (key*elt) -> (key*elt) -> Prop := @PX.ltk elt.
Lemma MapsTo_1 : forall m x y e, E.eq x y -> MapsTo x e m -> MapsTo y e m.
- Proof. intros m; exact (@Raw.MapsTo_1 _ m.(this)). Qed.
+ Proof. intros m; exact (@MapsTo_1 _ m.(this)). Qed.
Lemma mem_1 : forall m x, In x m -> mem x m = true.
Proof.
- unfold In, mem; intros m x; rewrite Raw.In_alt; simpl; apply Raw.mem_1; auto.
+ unfold In, mem; intros m x; rewrite In_alt; simpl; apply mem_1; auto.
apply m.(is_bst).
Qed.
Lemma mem_2 : forall m x, mem x m = true -> In x m.
Proof.
- unfold In, mem; intros m x; rewrite Raw.In_alt; simpl; apply Raw.mem_2; auto.
+ unfold In, mem; intros m x; rewrite In_alt; simpl; apply mem_2; auto.
Qed.
Lemma empty_1 : Empty empty.
- Proof. exact (@Raw.empty_1 elt). Qed.
+ Proof. exact (@empty_1 elt). Qed.
Lemma is_empty_1 : forall m, Empty m -> is_empty m = true.
- Proof. intros m; exact (@Raw.is_empty_1 _ m.(this)). Qed.
+ Proof. intros m; exact (@is_empty_1 _ m.(this)). Qed.
Lemma is_empty_2 : forall m, is_empty m = true -> Empty m.
- Proof. intros m; exact (@Raw.is_empty_2 _ m.(this)). Qed.
+ Proof. intros m; exact (@is_empty_2 _ m.(this)). Qed.
Lemma add_1 : forall m x y e, E.eq x y -> MapsTo y e (add x e m).
- Proof. intros m x y e; exact (@Raw.add_1 elt _ x y e m.(is_avl)). Qed.
+ Proof. intros m x y e; exact (@add_1 elt _ x y e). Qed.
Lemma add_2 : forall m x y e e', ~ E.eq x y -> MapsTo y e m -> MapsTo y e (add x e' m).
- Proof. intros m x y e e'; exact (@Raw.add_2 elt _ x y e e' m.(is_avl)). Qed.
+ Proof. intros m x y e e'; exact (@add_2 elt _ x y e e'). Qed.
Lemma add_3 : forall m x y e e', ~ E.eq x y -> MapsTo y e (add x e' m) -> MapsTo y e m.
- Proof. intros m x y e e'; exact (@Raw.add_3 elt _ x y e e' m.(is_avl)). Qed.
+ Proof. intros m x y e e'; exact (@add_3 elt _ x y e e'). Qed.
Lemma remove_1 : forall m x y, E.eq x y -> ~ In y (remove x m).
Proof.
- unfold In, remove; intros m x y; rewrite Raw.In_alt; simpl; apply Raw.remove_1; auto.
+ unfold In, remove; intros m x y; rewrite In_alt; simpl; apply remove_1; auto.
apply m.(is_bst).
- apply m.(is_avl).
Qed.
Lemma remove_2 : forall m x y e, ~ E.eq x y -> MapsTo y e m -> MapsTo y e (remove x m).
- Proof. intros m x y e; exact (@Raw.remove_2 elt _ x y e m.(is_bst) m.(is_avl)). Qed.
+ Proof. intros m x y e; exact (@remove_2 elt _ x y e m.(is_bst)). Qed.
Lemma remove_3 : forall m x y e, MapsTo y e (remove x m) -> MapsTo y e m.
- Proof. intros m x y e; exact (@Raw.remove_3 elt _ x y e m.(is_bst) m.(is_avl)). Qed.
+ Proof. intros m x y e; exact (@remove_3 elt _ x y e m.(is_bst)). Qed.
Lemma find_1 : forall m x e, MapsTo x e m -> find x m = Some e.
- Proof. intros m x e; exact (@Raw.find_1 elt _ x e m.(is_bst)). Qed.
+ Proof. intros m x e; exact (@find_1 elt _ x e m.(is_bst)). Qed.
Lemma find_2 : forall m x e, find x m = Some e -> MapsTo x e m.
- Proof. intros m; exact (@Raw.find_2 elt m.(this)). Qed.
+ Proof. intros m; exact (@find_2 elt m.(this)). Qed.
- Lemma fold_1 : forall m (A : Set) (i : A) (f : key -> elt -> A -> A),
+ Lemma fold_1 : forall m (A : Type) (i : A) (f : key -> elt -> A -> A),
fold f m i = fold_left (fun a p => f (fst p) (snd p) a) (elements m) i.
- Proof. intros m; exact (@Raw.fold_1 elt m.(this) m.(is_bst)). Qed.
+ Proof. intros m; exact (@fold_1 elt m.(this) m.(is_bst)). Qed.
Lemma elements_1 : forall m x e,
MapsTo x e m -> InA eq_key_elt (x,e) (elements m).
Proof.
- intros; unfold elements, MapsTo, eq_key_elt; rewrite Raw.elements_mapsto; auto.
+ intros; unfold elements, MapsTo, eq_key_elt; rewrite elements_mapsto; auto.
Qed.
Lemma elements_2 : forall m x e,
InA eq_key_elt (x,e) (elements m) -> MapsTo x e m.
Proof.
- intros; unfold elements, MapsTo, eq_key_elt; rewrite <- Raw.elements_mapsto; auto.
+ intros; unfold elements, MapsTo, eq_key_elt; rewrite <- elements_mapsto; auto.
Qed.
Lemma elements_3 : forall m, sort lt_key (elements m).
- Proof. intros m; exact (@Raw.elements_sort elt m.(this) m.(is_bst)). Qed.
+ Proof. intros m; exact (@elements_sort elt m.(this) m.(is_bst)). Qed.
- Definition Equal cmp m m' :=
+ Lemma elements_3w : forall m, NoDupA eq_key (elements m).
+ Proof. intros m; exact (@elements_nodup elt m.(this) m.(is_bst)). Qed.
+
+ Lemma cardinal_1 : forall m, cardinal m = length (elements m).
+ Proof. intro m; exact (@elements_cardinal elt m.(this)). Qed.
+
+ Definition Equal m m' := forall y, find y m = find y m'.
+ Definition Equiv (eq_elt:elt->elt->Prop) m m' :=
(forall k, In k m <-> In k m') /\
- (forall k e e', MapsTo k e m -> MapsTo k e' m' -> cmp e e' = true).
+ (forall k e e', MapsTo k e m -> MapsTo k e' m' -> eq_elt e e').
+ Definition Equivb cmp := Equiv (Cmp cmp).
- Lemma Equal_Equal : forall cmp m m', Equal cmp m m' <-> Raw.Equal cmp m m'.
+ Lemma Equivb_Equivb : forall cmp m m',
+ Equivb cmp m m' <-> Raw.Proofs.Equivb cmp m m'.
Proof.
- intros; unfold Equal, Raw.Equal, In; intuition.
- generalize (H0 k); do 2 rewrite Raw.In_alt; intuition.
- generalize (H0 k); do 2 rewrite Raw.In_alt; intuition.
- generalize (H0 k); do 2 rewrite <- Raw.In_alt; intuition.
- generalize (H0 k); do 2 rewrite <- Raw.In_alt; intuition.
- Qed.
+ intros; unfold Equivb, Equiv, Raw.Proofs.Equivb, In; intuition.
+ generalize (H0 k); do 2 rewrite In_alt; intuition.
+ generalize (H0 k); do 2 rewrite In_alt; intuition.
+ generalize (H0 k); do 2 rewrite <- In_alt; intuition.
+ generalize (H0 k); do 2 rewrite <- In_alt; intuition.
+ Qed.
Lemma equal_1 : forall m m' cmp,
- Equal cmp m m' -> equal cmp m m' = true.
+ Equivb cmp m m' -> equal cmp m m' = true.
Proof.
- unfold equal; intros m m' cmp; rewrite Equal_Equal.
- destruct (@Raw.equal _ cmp m m'); auto.
+ unfold equal; intros (m,b) (m',b') cmp; rewrite Equivb_Equivb;
+ intros; simpl in *; rewrite equal_Equivb; auto.
Qed.
Lemma equal_2 : forall m m' cmp,
- equal cmp m m' = true -> Equal cmp m m'.
+ equal cmp m m' = true -> Equivb cmp m m'.
Proof.
- unfold equal; intros; rewrite Equal_Equal.
- destruct (@Raw.equal _ cmp m m'); auto; try discriminate.
- Qed.
+ unfold equal; intros (m,b) (m',b') cmp; rewrite Equivb_Equivb;
+ intros; simpl in *; rewrite <-equal_Equivb; auto.
+ Qed.
End Elt.
- Lemma map_1 : forall (elt elt':Set)(m: t elt)(x:key)(e:elt)(f:elt->elt'),
+ Lemma map_1 : forall (elt elt':Type)(m: t elt)(x:key)(e:elt)(f:elt->elt'),
MapsTo x e m -> MapsTo x (f e) (map f m).
- Proof. intros elt elt' m x e f; exact (@Raw.map_1 elt elt' f m.(this) x e). Qed.
+ Proof. intros elt elt' m x e f; exact (@map_1 elt elt' f m.(this) x e). Qed.
- Lemma map_2 : forall (elt elt':Set)(m:t elt)(x:key)(f:elt->elt'), In x (map f m) -> In x m.
+ Lemma map_2 : forall (elt elt':Type)(m:t elt)(x:key)(f:elt->elt'), In x (map f m) -> In x m.
Proof.
- intros elt elt' m x f; do 2 unfold In in *; do 2 rewrite Raw.In_alt; simpl.
- apply Raw.map_2; auto.
+ intros elt elt' m x f; do 2 unfold In in *; do 2 rewrite In_alt; simpl.
+ apply map_2; auto.
Qed.
- Lemma mapi_1 : forall (elt elt':Set)(m: t elt)(x:key)(e:elt)
+ Lemma mapi_1 : forall (elt elt':Type)(m: t elt)(x:key)(e:elt)
(f:key->elt->elt'), MapsTo x e m ->
exists y, E.eq y x /\ MapsTo x (f y e) (mapi f m).
- Proof. intros elt elt' m x e f; exact (@Raw.mapi_1 elt elt' f m.(this) x e). Qed.
- Lemma mapi_2 : forall (elt elt':Set)(m: t elt)(x:key)
+ Proof. intros elt elt' m x e f; exact (@mapi_1 elt elt' f m.(this) x e). Qed.
+ Lemma mapi_2 : forall (elt elt':Type)(m: t elt)(x:key)
(f:key->elt->elt'), In x (mapi f m) -> In x m.
Proof.
- intros elt elt' m x f; unfold In in *; do 2 rewrite Raw.In_alt; simpl; apply Raw.mapi_2; auto.
- Qed.
+ intros elt elt' m x f; unfold In in *; do 2 rewrite In_alt; simpl; apply mapi_2; auto.
+ Qed.
- Lemma map2_1 : forall (elt elt' elt'':Set)(m: t elt)(m': t elt')
+ Lemma map2_1 : forall (elt elt' elt'':Type)(m: t elt)(m': t elt')
(x:key)(f:option elt->option elt'->option elt''),
In x m \/ In x m' ->
find x (map2 f m m') = f (find x m) (find x m').
Proof.
unfold find, map2, In; intros elt elt' elt'' m m' x f.
- do 2 rewrite Raw.In_alt; intros; simpl; apply Raw.map2_1; auto.
+ do 2 rewrite In_alt; intros; simpl; apply map2_1; auto.
apply m.(is_bst).
apply m'.(is_bst).
Qed.
- Lemma map2_2 : forall (elt elt' elt'':Set)(m: t elt)(m': t elt')
+ Lemma map2_2 : forall (elt elt' elt'':Type)(m: t elt)(m': t elt')
(x:key)(f:option elt->option elt'->option elt''),
In x (map2 f m m') -> In x m \/ In x m'.
Proof.
unfold In, map2; intros elt elt' elt'' m m' x f.
- do 3 rewrite Raw.In_alt; intros; simpl in *; eapply Raw.map2_2; eauto.
+ do 3 rewrite In_alt; intros; simpl in *; eapply map2_2; eauto.
apply m.(is_bst).
apply m'.(is_bst).
Qed.
@@ -1891,158 +2000,185 @@ End IntMake.
Module IntMake_ord (I:Int)(X: OrderedType)(D : OrderedType) <:
Sord with Module Data := D
- with Module MapS.E := X.
+ with Module MapS.E := X.
Module Data := D.
- Module MapS := IntMake(I)(X).
- Import MapS.
+ Module Import MapS := IntMake(I)(X).
+ Module LO := FMapList.Make_ord(X)(D).
+ Module R := Raw.
+ Module P := Raw.Proofs.
+
+ Definition t := MapS.t D.t.
+
+ Definition cmp e e' :=
+ match D.compare e e' with EQ _ => true | _ => false end.
+
+ (** One step of comparison of elements *)
+
+ Definition compare_more x1 d1 (cont:R.enumeration D.t -> comparison) e2 :=
+ match e2 with
+ | R.End => Gt
+ | R.More x2 d2 r2 e2 =>
+ match X.compare x1 x2 with
+ | EQ _ => match D.compare d1 d2 with
+ | EQ _ => cont (R.cons r2 e2)
+ | LT _ => Lt
+ | GT _ => Gt
+ end
+ | LT _ => Lt
+ | GT _ => Gt
+ end
+ end.
- Module MD := OrderedTypeFacts(D).
- Import MD.
+ (** Comparison of left tree, middle element, then right tree *)
- Module LO := FMapList.Make_ord(X)(D).
+ Fixpoint compare_cont s1 (cont:R.enumeration D.t -> comparison) e2 :=
+ match s1 with
+ | R.Leaf => cont e2
+ | R.Node l1 x1 d1 r1 _ =>
+ compare_cont l1 (compare_more x1 d1 (compare_cont r1 cont)) e2
+ end.
+
+ (** Initial continuation *)
+
+ Definition compare_end (e2:R.enumeration D.t) :=
+ match e2 with R.End => Eq | _ => Lt end.
+
+ (** The complete comparison *)
- Definition t := MapS.t D.t.
+ Definition compare_pure s1 s2 :=
+ compare_cont s1 compare_end (R.cons s2 (Raw.End _)).
- Definition cmp e e' := match D.compare e e' with EQ _ => true | _ => false end.
+ (** Correctness of this comparison *)
- Definition elements (m:t) :=
- LO.MapS.Build_slist (Raw.elements_sort m.(is_bst)).
+ Definition Cmp c :=
+ match c with
+ | Eq => LO.eq_list
+ | Lt => LO.lt_list
+ | Gt => (fun l1 l2 => LO.lt_list l2 l1)
+ end.
+
+ Lemma cons_Cmp : forall c x1 x2 d1 d2 l1 l2,
+ X.eq x1 x2 -> D.eq d1 d2 ->
+ Cmp c l1 l2 -> Cmp c ((x1,d1)::l1) ((x2,d2)::l2).
+ Proof.
+ destruct c; simpl; intros; P.MX.elim_comp; auto.
+ Qed.
+ Hint Resolve cons_Cmp.
+
+ Lemma compare_end_Cmp :
+ forall e2, Cmp (compare_end e2) nil (P.flatten_e e2).
+ Proof.
+ destruct e2; simpl; auto.
+ Qed.
+
+ Lemma compare_more_Cmp : forall x1 d1 cont x2 d2 r2 e2 l,
+ Cmp (cont (R.cons r2 e2)) l (R.elements r2 ++ P.flatten_e e2) ->
+ Cmp (compare_more x1 d1 cont (R.More x2 d2 r2 e2)) ((x1,d1)::l)
+ (P.flatten_e (R.More x2 d2 r2 e2)).
+ Proof.
+ simpl; intros; destruct X.compare; simpl;
+ try destruct D.compare; simpl; auto; P.MX.elim_comp; auto.
+ Qed.
+
+ Lemma compare_cont_Cmp : forall s1 cont e2 l,
+ (forall e, Cmp (cont e) l (P.flatten_e e)) ->
+ Cmp (compare_cont s1 cont e2) (R.elements s1 ++ l) (P.flatten_e e2).
+ Proof.
+ induction s1 as [|l1 Hl1 x1 d1 r1 Hr1 h1]; simpl; intros; auto.
+ rewrite <- P.elements_node; simpl.
+ apply Hl1; auto. clear e2. intros [|x2 d2 r2 e2].
+ simpl; auto.
+ apply compare_more_Cmp.
+ rewrite <- P.cons_1; auto.
+ Qed.
+
+ Lemma compare_Cmp : forall s1 s2,
+ Cmp (compare_pure s1 s2) (R.elements s1) (R.elements s2).
+ Proof.
+ intros; unfold compare_pure.
+ rewrite (app_nil_end (R.elements s1)).
+ replace (R.elements s2) with (P.flatten_e (R.cons s2 (R.End _))) by
+ (rewrite P.cons_1; simpl; rewrite <- app_nil_end; auto).
+ auto using compare_cont_Cmp, compare_end_Cmp.
+ Qed.
- Definition eq : t -> t -> Prop :=
- fun m1 m2 => LO.eq (elements m1) (elements m2).
+ (** The dependent-style [compare] *)
- Definition lt : t -> t -> Prop :=
- fun m1 m2 => LO.lt (elements m1) (elements m2).
+ Definition eq (m1 m2 : t) := LO.eq_list (elements m1) (elements m2).
+ Definition lt (m1 m2 : t) := LO.lt_list (elements m1) (elements m2).
- Lemma eq_1 : forall m m', Equal cmp m m' -> eq m m'.
+ Definition compare (s s':t) : Compare lt eq s s'.
+ Proof.
+ intros (s,b) (s',b').
+ generalize (compare_Cmp s s').
+ destruct compare_pure; intros; [apply EQ|apply LT|apply GT]; red; auto.
+ Defined.
+
+ (* Proofs about [eq] and [lt] *)
+
+ Definition selements (m1 : t) :=
+ LO.MapS.Build_slist (P.elements_sort m1.(is_bst)).
+
+ Definition seq (m1 m2 : t) := LO.eq (selements m1) (selements m2).
+ Definition slt (m1 m2 : t) := LO.lt (selements m1) (selements m2).
+
+ Lemma eq_seq : forall m1 m2, eq m1 m2 <-> seq m1 m2.
+ Proof.
+ unfold eq, seq, selements, elements, LO.eq; intuition.
+ Qed.
+
+ Lemma lt_slt : forall m1 m2, lt m1 m2 <-> slt m1 m2.
+ Proof.
+ unfold lt, slt, selements, elements, LO.lt; intuition.
+ Qed.
+
+ Lemma eq_1 : forall (m m' : t), Equivb cmp m m' -> eq m m'.
Proof.
intros m m'.
- unfold eq.
- rewrite Equal_Equal.
- rewrite Raw.Equal_elements.
- intros.
- apply LO.eq_1.
- auto.
+ rewrite eq_seq; unfold seq.
+ rewrite Equivb_Equivb.
+ rewrite P.Equivb_elements.
+ auto using LO.eq_1.
Qed.
- Lemma eq_2 : forall m m', eq m m' -> Equal cmp m m'.
+ Lemma eq_2 : forall m m', eq m m' -> Equivb cmp m m'.
Proof.
intros m m'.
- unfold eq.
- rewrite Equal_Equal.
- rewrite Raw.Equal_elements.
+ rewrite eq_seq; unfold seq.
+ rewrite Equivb_Equivb.
+ rewrite P.Equivb_elements.
intros.
generalize (LO.eq_2 H).
auto.
Qed.
-
+
Lemma eq_refl : forall m : t, eq m m.
Proof.
- unfold eq; intros; apply LO.eq_refl.
+ intros; rewrite eq_seq; unfold seq; intros; apply LO.eq_refl.
Qed.
Lemma eq_sym : forall m1 m2 : t, eq m1 m2 -> eq m2 m1.
Proof.
- unfold eq; intros; apply LO.eq_sym; auto.
+ intros m1 m2; rewrite 2 eq_seq; unfold seq; intros; apply LO.eq_sym; auto.
Qed.
Lemma eq_trans : forall m1 m2 m3 : t, eq m1 m2 -> eq m2 m3 -> eq m1 m3.
Proof.
- unfold eq; intros; eapply LO.eq_trans; eauto.
+ intros m1 m2 M3; rewrite 3 eq_seq; unfold seq.
+ intros; eapply LO.eq_trans; eauto.
Qed.
Lemma lt_trans : forall m1 m2 m3 : t, lt m1 m2 -> lt m2 m3 -> lt m1 m3.
Proof.
- unfold lt; intros; eapply LO.lt_trans; eauto.
+ intros m1 m2 m3; rewrite 3 lt_slt; unfold slt;
+ intros; eapply LO.lt_trans; eauto.
Qed.
Lemma lt_not_eq : forall m1 m2 : t, lt m1 m2 -> ~ eq m1 m2.
Proof.
- unfold lt, eq; intros; apply LO.lt_not_eq; auto.
- Qed.
-
- Import Raw.
-
- Definition flatten_slist (e:enumeration D.t)(He:sorted_e e) :=
- LO.MapS.Build_slist (sorted_flatten_e He).
-
- Open Local Scope Z_scope.
-
- Definition compare_aux :
- forall (e1 e2:enumeration D.t)(He1:sorted_e e1)(He2: sorted_e e2),
- Compare LO.lt LO.eq (flatten_slist He1) (flatten_slist He2).
- Proof.
- intros e1 e2; pattern e1, e2 in |- *; apply compare_rec2.
- simple destruct x; simple destruct x'; intuition.
- (* x = x' = End *)
- constructor 2.
- compute; auto.
- (* x = End x' = More *)
- constructor 1.
- compute; auto.
- (* x = More x' = End *)
- constructor 3.
- compute; auto.
- (* x = More k t0 t1 e, x' = More k0 t2 t3 e0 *)
- case (X.compare k k0); intro.
- (* k < k0 *)
- constructor 1.
- compute; MX.elim_comp; auto.
- (* k = k0 *)
- destruct (D.compare t t1).
- constructor 1.
- compute; MX.elim_comp; auto.
- destruct (@cons _ t0 e) as [c1 (H2,(H3,H4))]; try inversion_clear He1; auto.
- destruct (@cons _ t2 e0) as [c2 (H5,(H6,H7))]; try inversion_clear He2; auto.
- assert (measure_e c1 + measure_e c2 <
- measure_e (More k t t0 e) +
- measure_e (More k0 t1 t2 e0)).
- unfold measure_e in *; fold measure_e in *; omega.
- destruct (H c1 c2 H0 H2 H5); clear H.
- constructor 1.
- unfold flatten_slist, LO.lt in *; simpl; simpl in l.
- MX.elim_comp.
- right; split; auto.
- rewrite <- H7; rewrite <- H4; auto.
- constructor 2.
- unfold flatten_slist, LO.eq in *; simpl; simpl in e5.
- MX.elim_comp.
- split; auto.
- rewrite <- H7; rewrite <- H4; auto.
- constructor 3.
- unfold flatten_slist, LO.lt in *; simpl; simpl in l.
- MX.elim_comp.
- right; split; auto.
- rewrite <- H7; rewrite <- H4; auto.
- constructor 3.
- compute; MX.elim_comp; auto.
- (* k > k0 *)
- constructor 3.
- compute; MX.elim_comp; auto.
- Qed.
-
- Definition compare : forall m1 m2, Compare lt eq m1 m2.
- Proof.
- intros (m1,m1_bst,m1_avl) (m2,m2_bst,m2_avl); simpl.
- destruct (@cons _ m1 (End _)) as [x1 (H1,H11)]; auto.
- apply SortedEEnd.
- inversion_clear 2.
- destruct (@cons _ m2 (End _)) as [x2 (H2,H22)]; auto.
- apply SortedEEnd.
- inversion_clear 2.
- simpl in H11; rewrite <- app_nil_end in H11.
- simpl in H22; rewrite <- app_nil_end in H22.
- destruct (compare_aux H1 H2); intuition.
- constructor 1.
- unfold lt, LO.lt, IntMake_ord.elements, flatten_slist in *; simpl in *.
- rewrite <- H0; rewrite <- H4; auto.
- constructor 2.
- unfold eq, LO.eq, IntMake_ord.elements, flatten_slist in *; simpl in *.
- rewrite <- H0; rewrite <- H4; auto.
- constructor 3.
- unfold lt, LO.lt, IntMake_ord.elements, flatten_slist in *; simpl in *.
- rewrite <- H0; rewrite <- H4; auto.
+ intros m1 m2; rewrite lt_slt, eq_seq; unfold slt, seq;
+ intros; apply LO.lt_not_eq; auto.
Qed.
End IntMake_ord.
@@ -2056,3 +2192,4 @@ Module Make_ord (X: OrderedType)(D: OrderedType)
<: Sord with Module Data := D
with Module MapS.E := X
:=IntMake_ord(Z_as_Int)(X)(D).
+
diff --git a/theories/FSets/FMapFacts.v b/theories/FSets/FMapFacts.v
index 0105095a..b307efe3 100644
--- a/theories/FSets/FMapFacts.v
+++ b/theories/FSets/FMapFacts.v
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: FMapFacts.v 8882 2006-05-31 21:55:30Z letouzey $ *)
+(* $Id: FMapFacts.v 10782 2008-04-12 16:08:04Z msozeau $ *)
(** * Finite maps library *)
@@ -15,20 +15,24 @@
different styles: equivalence and boolean equalities.
*)
-Require Import Bool.
-Require Import OrderedType.
+Require Import Bool DecidableType DecidableTypeEx OrderedType.
Require Export FMapInterface.
Set Implicit Arguments.
Unset Strict Implicit.
-Module Facts (M: S).
-Module ME := OrderedTypeFacts M.E.
-Import ME.
-Import M.
-Import Logic. (* to unmask [eq] *)
-Import Peano. (* to unmask [lt] *)
+(** * Facts about weak maps *)
-Lemma MapsTo_fun : forall (elt:Set) m x (e e':elt),
+Module WFacts (E:DecidableType)(Import M:WSfun E).
+
+Notation eq_dec := E.eq_dec.
+Definition eqb x y := if eq_dec x y then true else false.
+
+Lemma eq_bool_alt : forall b b', b=b' <-> (b=true <-> b'=true).
+Proof.
+ destruct b; destruct b'; intuition.
+Qed.
+
+Lemma MapsTo_fun : forall (elt:Type) m x (e e':elt),
MapsTo x e m -> MapsTo x e' m -> e=e'.
Proof.
intros.
@@ -36,19 +40,14 @@ generalize (find_1 H) (find_1 H0); clear H H0.
intros; rewrite H in H0; injection H0; auto.
Qed.
-(** * Specifications written using equivalences *)
+(** ** Specifications written using equivalences *)
Section IffSpec.
-Variable elt elt' elt'': Set.
+Variable elt elt' elt'': Type.
Implicit Type m: t elt.
Implicit Type x y z: key.
Implicit Type e: elt.
-Lemma MapsTo_iff : forall m x y e, E.eq x y -> (MapsTo x e m <-> MapsTo y e m).
-Proof.
-split; apply MapsTo_1; auto.
-Qed.
-
Lemma In_iff : forall m x y, E.eq x y -> (In x m <-> In y m).
Proof.
unfold In.
@@ -57,12 +56,34 @@ apply (MapsTo_1 H H0); auto.
apply (MapsTo_1 (E.eq_sym H) H0); auto.
Qed.
+Lemma MapsTo_iff : forall m x y e, E.eq x y -> (MapsTo x e m <-> MapsTo y e m).
+Proof.
+split; apply MapsTo_1; auto.
+Qed.
+
+Lemma mem_in_iff : forall m x, In x m <-> mem x m = true.
+Proof.
+split; [apply mem_1|apply mem_2].
+Qed.
+
+Lemma not_mem_in_iff : forall m x, ~In x m <-> mem x m = false.
+Proof.
+intros; rewrite mem_in_iff; destruct (mem x m); intuition.
+Qed.
+
+Lemma In_dec : forall m x, { In x m } + { ~ In x m }.
+Proof.
+ intros.
+ generalize (mem_in_iff m x).
+ destruct (mem x m); [left|right]; intuition.
+Qed.
+
Lemma find_mapsto_iff : forall m x e, MapsTo x e m <-> find x m = Some e.
Proof.
split; [apply find_1|apply find_2].
Qed.
-Lemma not_find_mapsto_iff : forall m x, ~In x m <-> find x m = None.
+Lemma not_find_in_iff : forall m x, ~In x m <-> find x m = None.
Proof.
intros.
generalize (find_mapsto_iff m x); destruct (find x m).
@@ -74,17 +95,13 @@ intros; intros (e,H1).
rewrite H in H1; discriminate.
Qed.
-Lemma mem_in_iff : forall m x, In x m <-> mem x m = true.
-Proof.
-split; [apply mem_1|apply mem_2].
-Qed.
-
-Lemma not_mem_in_iff : forall m x, ~In x m <-> mem x m = false.
+Lemma in_find_iff : forall m x, In x m <-> find x m <> None.
Proof.
-intros; rewrite mem_in_iff; destruct (mem x m); intuition.
+intros; rewrite <- not_find_in_iff, mem_in_iff.
+destruct mem; intuition.
Qed.
-Lemma equal_iff : forall m m' cmp, Equal cmp m m' <-> equal cmp m m' = true.
+Lemma equal_iff : forall m m' cmp, Equivb cmp m m' <-> equal cmp m m' = true.
Proof.
split; [apply equal_1|apply equal_2].
Qed.
@@ -114,9 +131,9 @@ intros.
intuition.
destruct (eq_dec x y); [left|right].
split; auto.
-symmetry; apply (MapsTo_fun (e':=e) H); auto.
+symmetry; apply (MapsTo_fun (e':=e) H); auto with map.
split; auto; apply add_3 with x e; auto.
-subst; auto.
+subst; auto with map.
Qed.
Lemma add_in_iff : forall m x y e, In y (add x e m) <-> E.eq x y \/ In y m.
@@ -204,33 +221,33 @@ split.
case_eq (find x m); intros.
exists e.
split.
-apply (MapsTo_fun (m:=map f m) (x:=x)); auto.
-apply find_2; auto.
+apply (MapsTo_fun (m:=map f m) (x:=x)); auto with map.
+apply find_2; auto with map.
assert (In x (map f m)) by (exists b; auto).
destruct (map_2 H1) as (a,H2).
rewrite (find_1 H2) in H; discriminate.
intros (a,(H,H0)).
-subst b; auto.
+subst b; auto with map.
Qed.
Lemma map_in_iff : forall m x (f : elt -> elt'),
In x (map f m) <-> In x m.
Proof.
-split; intros; eauto.
+split; intros; eauto with map.
destruct H as (a,H).
-exists (f a); auto.
+exists (f a); auto with map.
Qed.
Lemma mapi_in_iff : forall m x (f:key->elt->elt'),
In x (mapi f m) <-> In x m.
Proof.
-split; intros; eauto.
+split; intros; eauto with map.
destruct H as (a,H).
destruct (mapi_1 f H) as (y,(H0,H1)).
exists (f y a); auto.
Qed.
-(* Unfortunately, we don't have simple equivalences for [mapi]
+(** Unfortunately, we don't have simple equivalences for [mapi]
and [MapsTo]. The only correct one needs compatibility of [f]. *)
Lemma mapi_inv : forall m x b (f : key -> elt -> elt'),
@@ -240,9 +257,9 @@ Proof.
intros; case_eq (find x m); intros.
exists e.
destruct (@mapi_1 _ _ m x e f) as (y,(H1,H2)).
-apply find_2; auto.
-exists y; repeat split; auto.
-apply (MapsTo_fun (m:=mapi f m) (x:=x)); auto.
+apply find_2; auto with map.
+exists y; repeat split; auto with map.
+apply (MapsTo_fun (m:=mapi f m) (x:=x)); auto with map.
assert (In x (mapi f m)) by (exists b; auto).
destruct (mapi_2 H1) as (a,H2).
rewrite (find_1 H2) in H0; discriminate.
@@ -287,11 +304,11 @@ Ltac map_iff :=
rewrite map_mapsto_iff || rewrite map_in_iff ||
rewrite mapi_in_iff)).
-(** * Specifications written using boolean predicates *)
+(** ** Specifications written using boolean predicates *)
Section BoolSpec.
-Lemma mem_find_b : forall (elt:Set)(m:t elt)(x:key), mem x m = if find x m then true else false.
+Lemma mem_find_b : forall (elt:Type)(m:t elt)(x:key), mem x m = if find x m then true else false.
Proof.
intros.
generalize (find_mapsto_iff m x)(mem_in_iff m x); unfold In.
@@ -303,7 +320,7 @@ destruct H0 as (e,H0).
destruct (H e); intuition discriminate.
Qed.
-Variable elt elt' elt'' : Set.
+Variable elt elt' elt'' : Type.
Implicit Types m : t elt.
Implicit Types x y z : key.
Implicit Types e : elt.
@@ -345,24 +362,24 @@ Qed.
Lemma add_eq_o : forall m x y e,
E.eq x y -> find y (add x e m) = Some e.
Proof.
-auto.
+auto with map.
Qed.
Lemma add_neq_o : forall m x y e,
~ E.eq x y -> find y (add x e m) = find y m.
Proof.
intros.
-case_eq (find y m); intros; auto.
-case_eq (find y (add x e m)); intros; auto.
+case_eq (find y m); intros; auto with map.
+case_eq (find y (add x e m)); intros; auto with map.
rewrite <- H0; symmetry.
-apply find_1; apply add_3 with x e; auto.
+apply find_1; apply add_3 with x e; auto with map.
Qed.
-Hint Resolve add_neq_o.
+Hint Resolve add_neq_o : map.
Lemma add_o : forall m x y e,
find y (add x e m) = if eq_dec x y then Some e else find y m.
Proof.
-intros; destruct (eq_dec x y); auto.
+intros; destruct (eq_dec x y); auto with map.
Qed.
Lemma add_eq_b : forall m x y e,
@@ -394,23 +411,23 @@ destruct (find y (remove x m)); auto.
destruct 2.
exists e; rewrite H0; auto.
Qed.
-Hint Resolve remove_eq_o.
+Hint Resolve remove_eq_o : map.
Lemma remove_neq_o : forall m x y,
~ E.eq x y -> find y (remove x m) = find y m.
Proof.
intros.
-case_eq (find y m); intros; auto.
-case_eq (find y (remove x m)); intros; auto.
+case_eq (find y m); intros; auto with map.
+case_eq (find y (remove x m)); intros; auto with map.
rewrite <- H0; symmetry.
-apply find_1; apply remove_3 with x; auto.
+apply find_1; apply remove_3 with x; auto with map.
Qed.
-Hint Resolve remove_neq_o.
+Hint Resolve remove_neq_o : map.
Lemma remove_o : forall m x y,
find y (remove x m) = if eq_dec x y then None else find y m.
Proof.
-intros; destruct (eq_dec x y); auto.
+intros; destruct (eq_dec x y); auto with map.
Qed.
Lemma remove_eq_b : forall m x y,
@@ -432,7 +449,7 @@ intros; do 2 rewrite mem_find_b; rewrite remove_o; unfold eqb.
destruct (eq_dec x y); auto.
Qed.
-Definition option_map (A:Set)(B:Set)(f:A->B)(o:option A) : option B :=
+Definition option_map (A B:Type)(f:A->B)(o:option A) : option B :=
match o with
| Some a => Some (f a)
| None => None
@@ -494,15 +511,15 @@ Proof.
intros.
case_eq (find x m); intros.
rewrite <- H0.
-apply map2_1; auto.
-left; exists e; auto.
+apply map2_1; auto with map.
+left; exists e; auto with map.
case_eq (find x m'); intros.
rewrite <- H0; rewrite <- H1.
apply map2_1; auto.
-right; exists e; auto.
+right; exists e; auto with map.
rewrite H.
-case_eq (find x (map2 f m m')); intros; auto.
-assert (In x (map2 f m m')) by (exists e; auto).
+case_eq (find x (map2 f m m')); intros; auto with map.
+assert (In x (map2 f m m')) by (exists e; auto with map).
destruct (map2_2 H3) as [(e0,H4)|(e0,H4)].
rewrite (find_1 H4) in H0; discriminate.
rewrite (find_1 H4) in H1; discriminate.
@@ -514,21 +531,18 @@ Proof.
intros.
assert (forall e, find x m = Some e <-> InA (eq_key_elt (elt:=elt)) (x,e) (elements m)).
intros; rewrite <- find_mapsto_iff; apply elements_mapsto_iff.
-assert (NoDupA (eq_key (elt:=elt)) (elements m)).
- apply SortA_NoDupA with (lt_key (elt:=elt)); unfold eq_key, lt_key; intuition eauto.
- destruct y; simpl in *.
- apply (E.lt_not_eq H0 H1).
- exact (elements_3 m).
+assert (H0:=elements_3w m).
generalize (fun e => @findA_NoDupA _ _ _ E.eq_sym E.eq_trans eq_dec (elements m) x e H0).
-unfold eqb.
-destruct (find x m); destruct (findA (fun y : E.t => if eq_dec x y then true else false) (elements m));
+fold (eqb x).
+destruct (find x m); destruct (findA (eqb x) (elements m));
simpl; auto; intros.
symmetry; rewrite <- H1; rewrite <- H; auto.
symmetry; rewrite <- H1; rewrite <- H; auto.
rewrite H; rewrite H1; auto.
Qed.
-Lemma elements_b : forall m x, mem x m = existsb (fun p => eqb x (fst p)) (elements m).
+Lemma elements_b : forall m x,
+ mem x m = existsb (fun p => eqb x (fst p)) (elements m).
Proof.
intros.
generalize (mem_in_iff m x)(elements_in_iff m x)
@@ -554,4 +568,1026 @@ Qed.
End BoolSpec.
+Section Equalities.
+
+Variable elt:Type.
+
+(** * Relations between [Equal], [Equiv] and [Equivb]. *)
+
+(** First, [Equal] is [Equiv] with Leibniz on elements. *)
+
+Lemma Equal_Equiv : forall (m m' : t elt),
+ Equal m m' <-> Equiv (@Logic.eq elt) m m'.
+Proof.
+ unfold Equal, Equiv; split; intros.
+ split; intros.
+ rewrite in_find_iff, in_find_iff, H; intuition.
+ rewrite find_mapsto_iff in H0,H1; congruence.
+ destruct H.
+ specialize (H y).
+ specialize (H0 y).
+ do 2 rewrite in_find_iff in H.
+ generalize (find_mapsto_iff m y)(find_mapsto_iff m' y).
+ do 2 destruct find; auto; intros.
+ f_equal; apply H0; [rewrite H1|rewrite H2]; auto.
+ destruct H as [H _]; now elim H.
+ destruct H as [_ H]; now elim H.
+Qed.
+
+(** [Equivb] and [Equiv] and equivalent when [eq_elt] and [cmp]
+ are related. *)
+
+Section Cmp.
+Variable eq_elt : elt->elt->Prop.
+Variable cmp : elt->elt->bool.
+
+Definition compat_cmp :=
+ forall e e', cmp e e' = true <-> eq_elt e e'.
+
+Lemma Equiv_Equivb : compat_cmp ->
+ forall m m', Equiv eq_elt m m' <-> Equivb cmp m m'.
+Proof.
+ unfold Equivb, Equiv, Cmp; intuition.
+ red in H; rewrite H; eauto.
+ red in H; rewrite <-H; eauto.
+Qed.
+End Cmp.
+
+(** Composition of the two last results: relation between [Equal]
+ and [Equivb]. *)
+
+Lemma Equal_Equivb : forall cmp,
+ (forall e e', cmp e e' = true <-> e = e') ->
+ forall (m m':t elt), Equal m m' <-> Equivb cmp m m'.
+Proof.
+ intros; rewrite Equal_Equiv.
+ apply Equiv_Equivb; auto.
+Qed.
+
+Lemma Equal_Equivb_eqdec :
+ forall eq_elt_dec : (forall e e', { e = e' } + { e <> e' }),
+ let cmp := fun e e' => if eq_elt_dec e e' then true else false in
+ forall (m m':t elt), Equal m m' <-> Equivb cmp m m'.
+Proof.
+intros; apply Equal_Equivb.
+unfold cmp; clear cmp; intros.
+destruct eq_elt_dec; now intuition.
+Qed.
+
+End Equalities.
+
+(** * [Equal] is a setoid equality. *)
+
+Lemma Equal_refl : forall (elt:Type)(m : t elt), Equal m m.
+Proof. red; reflexivity. Qed.
+
+Lemma Equal_sym : forall (elt:Type)(m m' : t elt),
+ Equal m m' -> Equal m' m.
+Proof. unfold Equal; auto. Qed.
+
+Lemma Equal_trans : forall (elt:Type)(m m' m'' : t elt),
+ Equal m m' -> Equal m' m'' -> Equal m m''.
+Proof. unfold Equal; congruence. Qed.
+
+Definition Equal_ST : forall elt:Type, Setoid_Theory (t elt) (@Equal _).
+Proof.
+constructor; red; [apply Equal_refl | apply Equal_sym | apply Equal_trans].
+Qed.
+
+Add Relation key E.eq
+ reflexivity proved by E.eq_refl
+ symmetry proved by E.eq_sym
+ transitivity proved by E.eq_trans
+ as KeySetoid.
+
+Implicit Arguments Equal [[elt]].
+
+Add Parametric Relation (elt : Type) : (t elt) Equal
+ reflexivity proved by (@Equal_refl elt)
+ symmetry proved by (@Equal_sym elt)
+ transitivity proved by (@Equal_trans elt)
+ as EqualSetoid.
+
+Add Parametric Morphism elt : (@In elt) with signature E.eq ==> Equal ==> iff as In_m.
+Proof.
+unfold Equal; intros k k' Hk m m' Hm.
+rewrite (In_iff m Hk), in_find_iff, in_find_iff, Hm; intuition.
+Qed.
+
+Add Parametric Morphism elt : (@MapsTo elt)
+ with signature E.eq ==> @Logic.eq _ ==> Equal ==> iff as MapsTo_m.
+Proof.
+unfold Equal; intros k k' Hk e m m' Hm.
+rewrite (MapsTo_iff m e Hk), find_mapsto_iff, find_mapsto_iff, Hm;
+ intuition.
+Qed.
+
+Add Parametric Morphism elt : (@Empty elt) with signature Equal ==> iff as Empty_m.
+Proof.
+unfold Empty; intros m m' Hm; intuition.
+rewrite <-Hm in H0; eauto.
+rewrite Hm in H0; eauto.
+Qed.
+
+Add Parametric Morphism elt : (@is_empty elt) with signature Equal ==> @Logic.eq _ as is_empty_m.
+Proof.
+intros m m' Hm.
+rewrite eq_bool_alt, <-is_empty_iff, <-is_empty_iff, Hm; intuition.
+Qed.
+
+Add Parametric Morphism elt : (@mem elt) with signature E.eq ==> Equal ==> @Logic.eq _ as mem_m.
+Proof.
+intros k k' Hk m m' Hm.
+rewrite eq_bool_alt, <- mem_in_iff, <-mem_in_iff, Hk, Hm; intuition.
+Qed.
+
+Add Parametric Morphism elt : (@find elt) with signature E.eq ==> Equal ==> @Logic.eq _ as find_m.
+Proof.
+intros k k' Hk m m' Hm.
+generalize (find_mapsto_iff m k)(find_mapsto_iff m' k')
+ (not_find_in_iff m k)(not_find_in_iff m' k');
+do 2 destruct find; auto; intros.
+rewrite <- H, Hk, Hm, H0; auto.
+rewrite <- H1, Hk, Hm, H2; auto.
+symmetry; rewrite <- H2, <-Hk, <-Hm, H1; auto.
+Qed.
+
+Add Parametric Morphism elt : (@add elt) with signature
+ E.eq ==> @Logic.eq _ ==> Equal ==> Equal as add_m.
+Proof.
+intros k k' Hk e m m' Hm y.
+rewrite add_o, add_o; do 2 destruct eq_dec; auto.
+elim n; rewrite <-Hk; auto.
+elim n; rewrite Hk; auto.
+Qed.
+
+Add Parametric Morphism elt : (@remove elt) with signature
+ E.eq ==> Equal ==> Equal as remove_m.
+Proof.
+intros k k' Hk m m' Hm y.
+rewrite remove_o, remove_o; do 2 destruct eq_dec; auto.
+elim n; rewrite <-Hk; auto.
+elim n; rewrite Hk; auto.
+Qed.
+
+Add Parametric Morphism elt elt' : (@map elt elt') with signature @Logic.eq _ ==> Equal ==> Equal as map_m.
+Proof.
+intros f m m' Hm y.
+rewrite map_o, map_o, Hm; auto.
+Qed.
+
+(* Later: Add Morphism cardinal *)
+
+(* old name: *)
+Notation not_find_mapsto_iff := not_find_in_iff.
+
+End WFacts.
+
+(** * Same facts for full maps *)
+
+Module Facts (M:S).
+ Module D := OT_as_DT M.E.
+ Include WFacts D M.
End Facts.
+
+(** * Additional Properties for weak maps
+
+ Results about [fold], [elements], induction principles...
+*)
+
+Module WProperties (E:DecidableType)(M:WSfun E).
+ Module Import F:=WFacts E M.
+ Import M.
+
+ Section Elt.
+ Variable elt:Type.
+
+ Definition Add x (e:elt) m m' := forall y, find y m' = find y (add x e m).
+
+ Notation eqke := (@eq_key_elt elt).
+ Notation eqk := (@eq_key elt).
+
+ Lemma elements_Empty : forall m:t elt, Empty m <-> elements m = nil.
+ Proof.
+ intros.
+ unfold Empty.
+ split; intros.
+ assert (forall a, ~ List.In a (elements m)).
+ red; intros.
+ apply (H (fst a) (snd a)).
+ rewrite elements_mapsto_iff.
+ rewrite InA_alt; exists a; auto.
+ split; auto; split; auto.
+ destruct (elements m); auto.
+ elim (H0 p); simpl; auto.
+ red; intros.
+ rewrite elements_mapsto_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 elt) = nil.
+ Proof.
+ rewrite <-elements_Empty; apply empty_1.
+ Qed.
+
+ Lemma fold_Empty : forall m (A:Type)(f:key->elt->A->A)(i:A),
+ Empty m -> fold f m i = i.
+ Proof.
+ intros.
+ rewrite fold_1.
+ rewrite elements_Empty in H; rewrite H; simpl; auto.
+ Qed.
+
+ Lemma NoDupA_eqk_eqke : forall l, NoDupA eqk l -> NoDupA eqke l.
+ Proof.
+ induction 1; auto.
+ constructor; auto.
+ contradict H.
+ destruct x as (x,y).
+ rewrite InA_alt in *; destruct H as ((a,b),((H1,H2),H3)); simpl in *.
+ exists (a,b); auto.
+ Qed.
+
+ Lemma fold_Equal : forall m1 m2 (A:Type)(eqA:A->A->Prop)(st:Setoid_Theory A eqA)
+ (f:key->elt->A->A)(i:A),
+ compat_op eqke eqA (fun y =>f (fst y) (snd y)) ->
+ transpose eqA (fun y => f (fst y) (snd y)) ->
+ Equal m1 m2 ->
+ eqA (fold f m1 i) (fold f m2 i).
+ Proof.
+ assert (eqke_refl : forall p, eqke p p).
+ red; auto.
+ assert (eqke_sym : forall p p', eqke p p' -> eqke p' p).
+ intros (x1,x2) (y1,y2); unfold eq_key_elt; simpl; intuition.
+ assert (eqke_trans : forall p p' p'', eqke p p' -> eqke p' p'' -> eqke p p'').
+ intros (x1,x2) (y1,y2) (z1,z2); unfold eq_key_elt; simpl.
+ intuition; eauto; congruence.
+ intros; do 2 rewrite fold_1; do 2 rewrite <- fold_left_rev_right.
+ apply fold_right_equivlistA with (eqA:=eqke) (eqB:=eqA); auto.
+ apply NoDupA_rev; auto; apply NoDupA_eqk_eqke; apply elements_3w.
+ apply NoDupA_rev; auto; apply NoDupA_eqk_eqke; apply elements_3w.
+ red; intros.
+ do 2 rewrite InA_rev.
+ destruct x; do 2 rewrite <- elements_mapsto_iff.
+ do 2 rewrite find_mapsto_iff.
+ rewrite H1; split; auto.
+ Qed.
+
+ Lemma fold_Add : forall m1 m2 x e (A:Type)(eqA:A->A->Prop)(st:Setoid_Theory A eqA)
+ (f:key->elt->A->A)(i:A),
+ compat_op eqke eqA (fun y =>f (fst y) (snd y)) ->
+ transpose eqA (fun y =>f (fst y) (snd y)) ->
+ ~In x m1 -> Add x e m1 m2 ->
+ eqA (fold f m2 i) (f x e (fold f m1 i)).
+ Proof.
+ assert (eqke_refl : forall p, eqke p p).
+ red; auto.
+ assert (eqke_sym : forall p p', eqke p p' -> eqke p' p).
+ intros (x1,x2) (y1,y2); unfold eq_key_elt; simpl; intuition.
+ assert (eqke_trans : forall p p' p'', eqke p p' -> eqke p' p'' -> eqke p p'').
+ intros (x1,x2) (y1,y2) (z1,z2); unfold eq_key_elt; simpl.
+ intuition; eauto; congruence.
+ intros; do 2 rewrite fold_1; do 2 rewrite <- fold_left_rev_right.
+ set (f':=fun y x0 => f (fst y) (snd y) x0) in *.
+ change (f x e (fold_right f' i (rev (elements m1))))
+ with (f' (x,e) (fold_right f' i (rev (elements m1)))).
+ apply fold_right_add with (eqA:=eqke)(eqB:=eqA); auto.
+ apply NoDupA_rev; auto; apply NoDupA_eqk_eqke; apply elements_3w.
+ apply NoDupA_rev; auto; apply NoDupA_eqk_eqke; apply elements_3w.
+ rewrite InA_rev.
+ contradict H1.
+ exists e.
+ rewrite elements_mapsto_iff; auto.
+ intros a.
+ rewrite InA_cons; do 2 rewrite InA_rev;
+ destruct a as (a,b); do 2 rewrite <- elements_mapsto_iff.
+ do 2 rewrite find_mapsto_iff; unfold eq_key_elt; simpl.
+ rewrite H2.
+ rewrite add_o.
+ destruct (eq_dec x a); intuition.
+ inversion H3; auto.
+ f_equal; auto.
+ elim H1.
+ exists b; apply MapsTo_1 with a; auto with map.
+ elim n; auto.
+ Qed.
+
+ Lemma cardinal_fold : forall m : t elt,
+ cardinal m = fold (fun _ _ => S) m 0.
+ Proof.
+ intros; rewrite cardinal_1, fold_1.
+ symmetry; apply fold_left_length; auto.
+ Qed.
+
+ Lemma cardinal_Empty : forall m : t elt,
+ Empty m <-> cardinal m = 0.
+ Proof.
+ intros.
+ rewrite cardinal_1, elements_Empty.
+ destruct (elements m); intuition; discriminate.
+ Qed.
+
+ Lemma Equal_cardinal : forall m m' : t elt,
+ Equal m m' -> cardinal m = cardinal m'.
+ Proof.
+ intros; do 2 rewrite cardinal_fold.
+ apply fold_Equal with (eqA:=@eq _); auto.
+ constructor; auto; congruence.
+ red; auto.
+ red; auto.
+ Qed.
+
+ Lemma cardinal_1 : forall m : t elt, Empty m -> cardinal m = 0.
+ Proof.
+ intros; rewrite <- cardinal_Empty; auto.
+ Qed.
+
+ Lemma cardinal_2 :
+ forall m m' x e, ~ In x m -> Add x e m m' -> cardinal m' = S (cardinal m).
+ Proof.
+ intros; do 2 rewrite cardinal_fold.
+ change S with ((fun _ _ => S) x e).
+ apply fold_Add; auto.
+ constructor; intros; auto; congruence.
+ red; simpl; auto.
+ red; simpl; auto.
+ Qed.
+
+ Lemma cardinal_inv_1 : forall m : t elt,
+ cardinal m = 0 -> Empty m.
+ Proof.
+ intros; rewrite cardinal_Empty; auto.
+ Qed.
+ Hint Resolve cardinal_inv_1 : map.
+
+ Lemma cardinal_inv_2 :
+ forall m n, cardinal m = S n -> { p : key*elt | MapsTo (fst p) (snd p) m }.
+ Proof.
+ intros; rewrite M.cardinal_1 in *.
+ generalize (elements_mapsto_iff m).
+ destruct (elements m); try discriminate.
+ exists p; auto.
+ rewrite H0; destruct p; simpl; auto.
+ constructor; red; auto.
+ Qed.
+
+ Lemma cardinal_inv_2b :
+ forall m, cardinal m <> 0 -> { p : key*elt | MapsTo (fst p) (snd p) m }.
+ Proof.
+ intros.
+ generalize (@cardinal_inv_2 m); destruct cardinal.
+ elim H;auto.
+ eauto.
+ Qed.
+
+ Lemma map_induction :
+ forall P : t elt -> Type,
+ (forall m, Empty m -> P m) ->
+ (forall m m', P m -> forall x e, ~In x m -> Add x e m m' -> P m') ->
+ forall m, P m.
+ Proof.
+ intros; remember (cardinal m) as n; revert m Heqn; induction n; intros.
+ apply X; apply cardinal_inv_1; auto.
+
+ destruct (cardinal_inv_2 (sym_eq Heqn)) as ((x,e),H0); simpl in *.
+ assert (Add x e (remove x m) m).
+ red; intros.
+ rewrite add_o; rewrite remove_o; destruct (eq_dec x y); eauto with map.
+ apply X0 with (remove x m) x e; auto with map.
+ apply IHn; auto with map.
+ assert (S n = S (cardinal (remove x m))).
+ rewrite Heqn; eapply cardinal_2; eauto with map.
+ inversion H1; auto with map.
+ Qed.
+
+ (** * Let's emulate some functions not present in the interface *)
+
+ Definition filter (f : key -> elt -> bool)(m : t elt) :=
+ fold (fun k e m => if f k e then add k e m else m) m (empty _).
+
+ Definition for_all (f : key -> elt -> bool)(m : t elt) :=
+ fold (fun k e b => if f k e then b else false) m true.
+
+ Definition exists_ (f : key -> elt -> bool)(m : t elt) :=
+ fold (fun k e b => if f k e then true else b) m false.
+
+ Definition partition (f : key -> elt -> bool)(m : t elt) :=
+ (filter f m, filter (fun k e => negb (f k e))).
+
+ Section Specs.
+ Variable f : key -> elt -> bool.
+ Hypothesis Hf : forall e, compat_bool E.eq (fun k => f k e).
+
+ Lemma filter_iff : forall m k e,
+ MapsTo k e (filter f m) <-> MapsTo k e m /\ f k e = true.
+ Proof.
+ unfold filter; intros.
+ rewrite fold_1.
+ rewrite <- fold_left_rev_right.
+ rewrite (elements_mapsto_iff m).
+ rewrite <- (InA_rev eqke (k,e) (elements m)).
+ assert (NoDupA eqk (rev (elements m))).
+ apply NoDupA_rev; auto; try apply elements_3w; auto.
+ intros (k1,e1); compute; auto.
+ intros (k1,e1)(k2,e2); compute; auto.
+ intros (k1,e1)(k2,e2)(k3,e3); compute; eauto.
+ induction (rev (elements m)); simpl; auto.
+
+ rewrite empty_mapsto_iff.
+ intuition.
+ inversion H1.
+
+ destruct a as (k',e'); simpl.
+ inversion_clear H.
+ case_eq (f k' e'); intros; simpl;
+ try rewrite add_mapsto_iff; rewrite IHl; clear IHl; intuition.
+ constructor; red; auto.
+ rewrite (Hf e' H2),H4 in H; auto.
+ inversion_clear H3.
+ compute in H2; destruct H2; auto.
+ destruct (E.eq_dec k' k); auto.
+ elim H0.
+ rewrite InA_alt in *; destruct H2 as (w,Hw); exists w; intuition.
+ red in H2; red; simpl in *; intuition.
+ rewrite e0; auto.
+ inversion_clear H3; auto.
+ compute in H2; destruct H2.
+ rewrite (Hf e H2),H3,H in H4; discriminate.
+ Qed.
+
+ Lemma for_all_iff : forall m,
+ for_all f m = true <-> (forall k e, MapsTo k e m -> f k e = true).
+ Proof.
+ cut (forall m : t elt,
+ for_all f m = true <->
+ (forall k e, InA eqke (k,e) (rev (elements m)) -> f k e = true)).
+ intros; rewrite H; split; intros.
+ apply H0; rewrite InA_rev, <- elements_mapsto_iff; auto.
+ apply H0; rewrite InA_rev, <- elements_mapsto_iff in H1; auto.
+
+ unfold for_all; intros.
+ rewrite fold_1.
+ rewrite <- fold_left_rev_right.
+ assert (NoDupA eqk (rev (elements m))).
+ apply NoDupA_rev; auto; try apply elements_3w; auto.
+ intros (k1,e1); compute; auto.
+ intros (k1,e1)(k2,e2); compute; auto.
+ intros (k1,e1)(k2,e2)(k3,e3); compute; eauto.
+ induction (rev (elements m)); simpl; auto.
+
+ intuition.
+ inversion H1.
+
+ destruct a as (k,e); simpl.
+ inversion_clear H.
+ case_eq (f k e); intros; simpl;
+ try rewrite IHl; clear IHl; intuition.
+ inversion_clear H3; auto.
+ compute in H4; destruct H4.
+ rewrite (Hf e0 H3), H4; auto.
+ rewrite <-H, <-(H2 k e); auto.
+ constructor; red; auto.
+ Qed.
+
+ Lemma exists_iff : forall m,
+ exists_ f m = true <->
+ (exists p, MapsTo (fst p) (snd p) m /\ f (fst p) (snd p) = true).
+ Proof.
+ cut (forall m : t elt,
+ exists_ f m = true <->
+ (exists p, InA eqke p (rev (elements m))
+ /\ f (fst p) (snd p) = true)).
+ intros; rewrite H; split; intros.
+ destruct H0 as ((k,e),Hke); exists (k,e).
+ rewrite InA_rev, <-elements_mapsto_iff in Hke; auto.
+ destruct H0 as ((k,e),Hke); exists (k,e).
+ rewrite InA_rev, <-elements_mapsto_iff; auto.
+ unfold exists_; intros.
+ rewrite fold_1.
+ rewrite <- fold_left_rev_right.
+ assert (NoDupA eqk (rev (elements m))).
+ apply NoDupA_rev; auto; try apply elements_3w; auto.
+ intros (k1,e1); compute; auto.
+ intros (k1,e1)(k2,e2); compute; auto.
+ intros (k1,e1)(k2,e2)(k3,e3); compute; eauto.
+ induction (rev (elements m)); simpl; auto.
+
+ intuition; try discriminate.
+ destruct H0 as ((k,e),(Hke,_)); inversion Hke.
+
+ destruct a as (k,e); simpl.
+ inversion_clear H.
+ case_eq (f k e); intros; simpl;
+ try rewrite IHl; clear IHl; intuition.
+ exists (k,e); simpl; split; auto.
+ constructor; red; auto.
+ destruct H2 as ((k',e'),(Hke',Hf')); exists (k',e'); simpl; auto.
+ destruct H2 as ((k',e'),(Hke',Hf')); simpl in *.
+ inversion_clear Hke'.
+ compute in H2; destruct H2.
+ rewrite (Hf e' H2), H3,H in Hf'; discriminate.
+ exists (k',e'); auto.
+ Qed.
+ End Specs.
+
+ (** specialized versions analyzing only keys (resp. elements) *)
+
+ Definition filter_dom (f : key -> bool) := filter (fun k _ => f k).
+ Definition filter_range (f : elt -> bool) := filter (fun _ => f).
+ Definition for_all_dom (f : key -> bool) := for_all (fun k _ => f k).
+ Definition for_all_range (f : elt -> bool) := for_all (fun _ => f).
+ Definition exists_dom (f : key -> bool) := exists_ (fun k _ => f k).
+ Definition exists_range (f : elt -> bool) := exists_ (fun _ => f).
+ Definition partition_dom (f : key -> bool) := partition (fun k _ => f k).
+ Definition partition_range (f : elt -> bool) := partition (fun _ => f).
+
+ End Elt.
+
+ Add Parametric Morphism elt : (@cardinal elt) with signature Equal ==> @Logic.eq _ as cardinal_m.
+ Proof. intros; apply Equal_cardinal; auto. Qed.
+
+End WProperties.
+
+(** * Same Properties for full maps *)
+
+Module Properties (M:S).
+ Module D := OT_as_DT M.E.
+ Include WProperties D M.
+End Properties.
+
+(** * Properties specific to maps with ordered keys *)
+
+Module OrdProperties (M:S).
+ Module Import ME := OrderedTypeFacts M.E.
+ Module Import O:=KeyOrderedType M.E.
+ Module Import P:=Properties M.
+ Import F.
+ Import M.
+
+ Section Elt.
+ Variable elt:Type.
+
+ Notation eqke := (@eqke elt).
+ Notation eqk := (@eqk elt).
+ Notation ltk := (@ltk elt).
+ Notation cardinal := (@cardinal elt).
+ Notation Equal := (@Equal elt).
+ Notation Add := (@Add elt).
+
+ Definition Above x (m:t elt) := forall y, In y m -> E.lt y x.
+ Definition Below x (m:t elt) := forall y, In y m -> E.lt x y.
+
+ Section Elements.
+
+ Lemma sort_equivlistA_eqlistA : forall l l' : list (key*elt),
+ sort ltk l -> sort ltk l' -> equivlistA eqke l l' -> eqlistA eqke l l'.
+ Proof.
+ apply SortA_equivlistA_eqlistA; eauto;
+ unfold O.eqke, O.ltk; simpl; intuition; eauto.
+ Qed.
+
+ Ltac clean_eauto := unfold O.eqke, O.ltk; simpl; intuition; eauto.
+
+ Definition gtb (p p':key*elt) := match E.compare (fst p) (fst p') with GT _ => true | _ => false end.
+ Definition leb p := fun p' => negb (gtb p p').
+
+ Definition elements_lt p m := List.filter (gtb p) (elements m).
+ Definition elements_ge p m := List.filter (leb p) (elements m).
+
+ Lemma gtb_1 : forall p p', gtb p p' = true <-> ltk p' p.
+ Proof.
+ intros (x,e) (y,e'); unfold gtb, O.ltk; simpl.
+ destruct (E.compare x y); intuition; try discriminate; ME.order.
+ Qed.
+
+ Lemma leb_1 : forall p p', leb p p' = true <-> ~ltk p' p.
+ Proof.
+ intros (x,e) (y,e'); unfold leb, gtb, O.ltk; simpl.
+ destruct (E.compare x y); intuition; try discriminate; ME.order.
+ Qed.
+
+ Lemma gtb_compat : forall p, compat_bool eqke (gtb p).
+ Proof.
+ red; intros (x,e) (a,e') (b,e'') H; red in H; simpl in *; destruct H.
+ generalize (gtb_1 (x,e) (a,e'))(gtb_1 (x,e) (b,e''));
+ destruct (gtb (x,e) (a,e')); destruct (gtb (x,e) (b,e'')); auto.
+ unfold O.ltk in *; simpl in *; intros.
+ symmetry; rewrite H2.
+ apply ME.eq_lt with a; auto.
+ rewrite <- H1; auto.
+ unfold O.ltk in *; simpl in *; intros.
+ rewrite H1.
+ apply ME.eq_lt with b; auto.
+ rewrite <- H2; auto.
+ Qed.
+
+ Lemma leb_compat : forall p, compat_bool eqke (leb p).
+ Proof.
+ red; intros x a b H.
+ unfold leb; f_equal; apply gtb_compat; auto.
+ Qed.
+
+ Hint Resolve gtb_compat leb_compat elements_3 : map.
+
+ Lemma elements_split : forall p m,
+ elements m = elements_lt p m ++ elements_ge p m.
+ Proof.
+ unfold elements_lt, elements_ge, leb; intros.
+ apply filter_split with (eqA:=eqk) (ltA:=ltk); eauto with map.
+ intros; destruct x; destruct y; destruct p.
+ rewrite gtb_1 in H; unfold O.ltk in H; simpl in *.
+ assert (~ltk (t1,e0) (k,e1)).
+ unfold gtb, O.ltk in *; simpl in *.
+ destruct (E.compare k t1); intuition; try discriminate; ME.order.
+ unfold O.ltk in *; simpl in *; ME.order.
+ Qed.
+
+ Lemma elements_Add : forall m m' x e, ~In x m -> Add x e m m' ->
+ eqlistA eqke (elements m')
+ (elements_lt (x,e) m ++ (x,e):: elements_ge (x,e) m).
+ Proof.
+ intros; unfold elements_lt, elements_ge.
+ apply sort_equivlistA_eqlistA; auto with map.
+ apply (@SortA_app _ eqke); auto with map.
+ apply (@filter_sort _ eqke); auto with map; clean_eauto.
+ constructor; auto with map.
+ apply (@filter_sort _ eqke); auto with map; clean_eauto.
+ rewrite (@InfA_alt _ eqke); auto with map; try (clean_eauto; fail).
+ intros.
+ rewrite filter_InA in H1; auto with map; destruct H1.
+ rewrite leb_1 in H2.
+ destruct y; unfold O.ltk in *; simpl in *.
+ rewrite <- elements_mapsto_iff in H1.
+ assert (~E.eq x t0).
+ contradict H.
+ exists e0; apply MapsTo_1 with t0; auto.
+ ME.order.
+ apply (@filter_sort _ eqke); auto with map; clean_eauto.
+ intros.
+ rewrite filter_InA in H1; auto with map; destruct H1.
+ rewrite gtb_1 in H3.
+ destruct y; destruct x0; unfold O.ltk in *; simpl in *.
+ inversion_clear H2.
+ red in H4; simpl in *; destruct H4.
+ ME.order.
+ rewrite filter_InA in H4; auto with map; destruct H4.
+ rewrite leb_1 in H4.
+ unfold O.ltk in *; simpl in *; ME.order.
+ red; intros a; destruct a.
+ rewrite InA_app_iff; rewrite InA_cons.
+ do 2 (rewrite filter_InA; auto with map).
+ do 2 rewrite <- elements_mapsto_iff.
+ rewrite leb_1; rewrite gtb_1.
+ rewrite find_mapsto_iff; rewrite (H0 t0); rewrite <- find_mapsto_iff.
+ rewrite add_mapsto_iff.
+ unfold O.eqke, O.ltk; simpl.
+ destruct (E.compare t0 x); intuition.
+ right; split; auto; ME.order.
+ ME.order.
+ elim H.
+ exists e0; apply MapsTo_1 with t0; auto.
+ right; right; split; auto; ME.order.
+ ME.order.
+ right; split; auto; ME.order.
+ Qed.
+
+ Lemma elements_Add_Above : forall m m' x e,
+ Above x m -> Add x e m m' ->
+ eqlistA eqke (elements m') (elements m ++ (x,e)::nil).
+ Proof.
+ intros.
+ apply sort_equivlistA_eqlistA; auto with map.
+ apply (@SortA_app _ eqke); auto with map.
+ intros.
+ inversion_clear H2.
+ destruct x0; destruct y.
+ rewrite <- elements_mapsto_iff in H1.
+ unfold O.eqke, O.ltk in *; simpl in *; destruct H3.
+ apply ME.lt_eq with x; auto.
+ apply H; firstorder.
+ inversion H3.
+ red; intros a; destruct a.
+ rewrite InA_app_iff; rewrite InA_cons; rewrite InA_nil.
+ do 2 rewrite <- elements_mapsto_iff.
+ rewrite find_mapsto_iff; rewrite (H0 t0); rewrite <- find_mapsto_iff.
+ rewrite add_mapsto_iff; unfold O.eqke; simpl.
+ intuition.
+ destruct (ME.eq_dec x t0); auto.
+ elimtype False.
+ assert (In t0 m).
+ exists e0; auto.
+ generalize (H t0 H1).
+ ME.order.
+ Qed.
+
+ Lemma elements_Add_Below : forall m m' x e,
+ Below x m -> Add x e m m' ->
+ eqlistA eqke (elements m') ((x,e)::elements m).
+ Proof.
+ intros.
+ apply sort_equivlistA_eqlistA; auto with map.
+ change (sort ltk (((x,e)::nil) ++ elements m)).
+ apply (@SortA_app _ eqke); auto with map.
+ intros.
+ inversion_clear H1.
+ destruct y; destruct x0.
+ rewrite <- elements_mapsto_iff in H2.
+ unfold O.eqke, O.ltk in *; simpl in *; destruct H3.
+ apply ME.eq_lt with x; auto.
+ apply H; firstorder.
+ inversion H3.
+ red; intros a; destruct a.
+ rewrite InA_cons.
+ do 2 rewrite <- elements_mapsto_iff.
+ rewrite find_mapsto_iff; rewrite (H0 t0); rewrite <- find_mapsto_iff.
+ rewrite add_mapsto_iff; unfold O.eqke; simpl.
+ intuition.
+ destruct (ME.eq_dec x t0); auto.
+ elimtype False.
+ assert (In t0 m).
+ exists e0; auto.
+ generalize (H t0 H1).
+ ME.order.
+ Qed.
+
+ Lemma elements_Equal_eqlistA : forall (m m': t elt),
+ Equal m m' -> eqlistA eqke (elements m) (elements m').
+ Proof.
+ intros.
+ apply sort_equivlistA_eqlistA; auto with map.
+ red; intros.
+ destruct x; do 2 rewrite <- elements_mapsto_iff.
+ do 2 rewrite find_mapsto_iff; rewrite H; split; auto.
+ Qed.
+
+ End Elements.
+
+ Section Min_Max_Elt.
+
+ (** We emulate two [max_elt] and [min_elt] functions. *)
+
+ Fixpoint max_elt_aux (l:list (key*elt)) := match l with
+ | nil => None
+ | (x,e)::nil => Some (x,e)
+ | (x,e)::l => max_elt_aux l
+ end.
+ Definition max_elt m := max_elt_aux (elements m).
+
+ Lemma max_elt_Above :
+ forall m x e, max_elt m = Some (x,e) -> Above x (remove x m).
+ Proof.
+ red; intros.
+ rewrite remove_in_iff in H0.
+ destruct H0.
+ rewrite elements_in_iff in H1.
+ destruct H1.
+ unfold max_elt in *.
+ generalize (elements_3 m).
+ revert x e H y x0 H0 H1.
+ induction (elements m).
+ simpl; intros; try discriminate.
+ intros.
+ destruct a; destruct l; simpl in *.
+ injection H; clear H; intros; subst.
+ inversion_clear H1.
+ red in H; simpl in *; intuition.
+ elim H0; eauto.
+ inversion H.
+ change (max_elt_aux (p::l) = Some (x,e)) in H.
+ generalize (IHl x e H); clear IHl; intros IHl.
+ inversion_clear H1; [ | inversion_clear H2; eauto ].
+ red in H3; simpl in H3; destruct H3.
+ destruct p as (p1,p2).
+ destruct (ME.eq_dec p1 x).
+ apply ME.lt_eq with p1; auto.
+ inversion_clear H2.
+ inversion_clear H5.
+ red in H2; simpl in H2; ME.order.
+ apply E.lt_trans with p1; auto.
+ inversion_clear H2.
+ inversion_clear H5.
+ red in H2; simpl in H2; ME.order.
+ eapply IHl; eauto.
+ econstructor; eauto.
+ red; eauto.
+ inversion H2; auto.
+ Qed.
+
+ Lemma max_elt_MapsTo :
+ forall m x e, max_elt m = Some (x,e) -> MapsTo x e m.
+ Proof.
+ intros.
+ unfold max_elt in *.
+ rewrite elements_mapsto_iff.
+ induction (elements m).
+ simpl; try discriminate.
+ destruct a; destruct l; simpl in *.
+ injection H; intros; subst; constructor; red; auto.
+ constructor 2; auto.
+ Qed.
+
+ Lemma max_elt_Empty :
+ forall m, max_elt m = None -> Empty m.
+ Proof.
+ intros.
+ unfold max_elt in *.
+ rewrite elements_Empty.
+ induction (elements m); auto.
+ destruct a; destruct l; simpl in *; try discriminate.
+ assert (H':=IHl H); discriminate.
+ Qed.
+
+ Definition min_elt m : option (key*elt) := match elements m with
+ | nil => None
+ | (x,e)::_ => Some (x,e)
+ end.
+
+ Lemma min_elt_Below :
+ forall m x e, min_elt m = Some (x,e) -> Below x (remove x m).
+ Proof.
+ unfold min_elt, Below; intros.
+ rewrite remove_in_iff in H0; destruct H0.
+ rewrite elements_in_iff in H1.
+ destruct H1.
+ generalize (elements_3 m).
+ destruct (elements m).
+ try discriminate.
+ destruct p; injection H; intros; subst.
+ inversion_clear H1.
+ red in H2; destruct H2; simpl in *; ME.order.
+ inversion_clear H4.
+ rewrite (@InfA_alt _ eqke) in H3; eauto.
+ apply (H3 (y,x0)); auto.
+ unfold lt_key; simpl; intuition; eauto.
+ intros (x1,x2) (y1,y2) (z1,z2); compute; intuition; eauto.
+ intros (x1,x2) (y1,y2) (z1,z2); compute; intuition; eauto.
+ Qed.
+
+ Lemma min_elt_MapsTo :
+ forall m x e, min_elt m = Some (x,e) -> MapsTo x e m.
+ Proof.
+ intros.
+ unfold min_elt in *.
+ rewrite elements_mapsto_iff.
+ destruct (elements m).
+ simpl; try discriminate.
+ destruct p; simpl in *.
+ injection H; intros; subst; constructor; red; auto.
+ Qed.
+
+ Lemma min_elt_Empty :
+ forall m, min_elt m = None -> Empty m.
+ Proof.
+ intros.
+ unfold min_elt in *.
+ rewrite elements_Empty.
+ destruct (elements m); auto.
+ destruct p; simpl in *; discriminate.
+ Qed.
+
+ End Min_Max_Elt.
+
+ Section Induction_Principles.
+
+ Lemma map_induction_max :
+ forall P : t elt -> Type,
+ (forall m, Empty m -> P m) ->
+ (forall m m', P m -> forall x e, Above x m -> Add x e m m' -> P m') ->
+ forall m, P m.
+ Proof.
+ intros; remember (cardinal m) as n; revert m Heqn; induction n; intros.
+ apply X; apply cardinal_inv_1; auto.
+
+ case_eq (max_elt m); intros.
+ destruct p.
+ assert (Add k e (remove k m) m).
+ red; intros.
+ rewrite add_o; rewrite remove_o; destruct (eq_dec k y); eauto.
+ apply find_1; apply MapsTo_1 with k; auto.
+ apply max_elt_MapsTo; auto.
+ apply X0 with (remove k m) k e; auto with map.
+ apply IHn.
+ assert (S n = S (cardinal (remove k m))).
+ rewrite Heqn.
+ eapply cardinal_2; eauto with map.
+ inversion H1; auto.
+ eapply max_elt_Above; eauto.
+
+ apply X; apply max_elt_Empty; auto.
+ Qed.
+
+ Lemma map_induction_min :
+ forall P : t elt -> Type,
+ (forall m, Empty m -> P m) ->
+ (forall m m', P m -> forall x e, Below x m -> Add x e m m' -> P m') ->
+ forall m, P m.
+ Proof.
+ intros; remember (cardinal m) as n; revert m Heqn; induction n; intros.
+ apply X; apply cardinal_inv_1; auto.
+
+ case_eq (min_elt m); intros.
+ destruct p.
+ assert (Add k e (remove k m) m).
+ red; intros.
+ rewrite add_o; rewrite remove_o; destruct (eq_dec k y); eauto.
+ apply find_1; apply MapsTo_1 with k; auto.
+ apply min_elt_MapsTo; auto.
+ apply X0 with (remove k m) k e; auto.
+ apply IHn.
+ assert (S n = S (cardinal (remove k m))).
+ rewrite Heqn.
+ eapply cardinal_2; eauto with map.
+ inversion H1; auto.
+ eapply min_elt_Below; eauto.
+
+ apply X; apply min_elt_Empty; auto.
+ Qed.
+
+ End Induction_Principles.
+
+ Section Fold_properties.
+
+ (** The following lemma has already been proved on Weak Maps,
+ but with one additionnal hypothesis (some [transpose] fact). *)
+
+ Lemma fold_Equal : forall s1 s2 (A:Type)(eqA:A->A->Prop)(st:Setoid_Theory A eqA)
+ (f:key->elt->A->A)(i:A),
+ compat_op eqke eqA (fun y =>f (fst y) (snd y)) ->
+ Equal s1 s2 ->
+ eqA (fold f s1 i) (fold f s2 i).
+ Proof.
+ intros.
+ do 2 rewrite fold_1.
+ do 2 rewrite <- fold_left_rev_right.
+ apply fold_right_eqlistA with (eqA:=eqke) (eqB:=eqA); auto.
+ apply eqlistA_rev.
+ apply elements_Equal_eqlistA; auto.
+ Qed.
+
+ Lemma fold_Add : forall s1 s2 x e (A:Type)(eqA:A->A->Prop)(st:Setoid_Theory A eqA)
+ (f:key->elt->A->A)(i:A),
+ compat_op eqke eqA (fun y =>f (fst y) (snd y)) ->
+ transpose eqA (fun y =>f (fst y) (snd y)) ->
+ ~In x s1 -> Add x e s1 s2 ->
+ eqA (fold f s2 i) (f x e (fold f s1 i)).
+ Proof.
+ intros; do 2 rewrite fold_1; do 2 rewrite <- fold_left_rev_right.
+ set (f':=fun y x0 => f (fst y) (snd y) x0) in *.
+ change (f x e (fold_right f' i (rev (elements s1))))
+ with (f' (x,e) (fold_right f' i (rev (elements s1)))).
+ trans_st (fold_right f' i
+ (rev (elements_lt (x, e) s1 ++ (x,e) :: elements_ge (x, e) s1))).
+ apply fold_right_eqlistA with (eqA:=eqke) (eqB:=eqA); auto.
+ apply eqlistA_rev.
+ apply elements_Add; auto.
+ rewrite distr_rev; simpl.
+ rewrite app_ass; simpl.
+ rewrite (elements_split (x,e) s1).
+ rewrite distr_rev; simpl.
+ apply fold_right_commutes with (eqA:=eqke) (eqB:=eqA); auto.
+ Qed.
+
+ Lemma fold_Add_Above : forall s1 s2 x e (A:Type)(eqA:A->A->Prop)(st:Setoid_Theory A eqA)
+ (f:key->elt->A->A)(i:A),
+ compat_op eqke eqA (fun y =>f (fst y) (snd y)) ->
+ Above x s1 -> Add x e s1 s2 ->
+ eqA (fold f s2 i) (f x e (fold f s1 i)).
+ Proof.
+ intros; do 2 rewrite fold_1; do 2 rewrite <- fold_left_rev_right.
+ set (f':=fun y x0 => f (fst y) (snd y) x0) in *.
+ trans_st (fold_right f' i (rev (elements s1 ++ (x,e)::nil))).
+ apply fold_right_eqlistA with (eqA:=eqke) (eqB:=eqA); auto.
+ apply eqlistA_rev.
+ apply elements_Add_Above; auto.
+ rewrite distr_rev; simpl.
+ refl_st.
+ Qed.
+
+ Lemma fold_Add_Below : forall s1 s2 x e (A:Type)(eqA:A->A->Prop)(st:Setoid_Theory A eqA)
+ (f:key->elt->A->A)(i:A),
+ compat_op eqke eqA (fun y =>f (fst y) (snd y)) ->
+ Below x s1 -> Add x e s1 s2 ->
+ eqA (fold f s2 i) (fold f s1 (f x e i)).
+ Proof.
+ intros; do 2 rewrite fold_1; do 2 rewrite <- fold_left_rev_right.
+ set (f':=fun y x0 => f (fst y) (snd y) x0) in *.
+ trans_st (fold_right f' i (rev (((x,e)::nil)++elements s1))).
+ apply fold_right_eqlistA with (eqA:=eqke) (eqB:=eqA); auto.
+ apply eqlistA_rev.
+ simpl; apply elements_Add_Below; auto.
+ rewrite distr_rev; simpl.
+ rewrite fold_right_app.
+ refl_st.
+ Qed.
+
+ End Fold_properties.
+
+ End Elt.
+
+End OrdProperties.
+
+
+
+
+
diff --git a/theories/FSets/FMapFullAVL.v b/theories/FSets/FMapFullAVL.v
new file mode 100644
index 00000000..57cbbcc4
--- /dev/null
+++ b/theories/FSets/FMapFullAVL.v
@@ -0,0 +1,823 @@
+
+(***********************************************************************)
+(* 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 *)
+(***********************************************************************)
+
+(* Finite map library. *)
+
+(* $Id: FMapFullAVL.v 10748 2008-04-03 18:28:26Z letouzey $ *)
+
+(** * FMapFullAVL
+
+ This file contains some complements to [FMapAVL].
+
+ - Functor [AvlProofs] proves that trees of [FMapAVL] are not only
+ binary search trees, but moreover well-balanced ones. This is done
+ by proving that all operations preserve the balancing.
+
+ - We then pack the previous elements in a [IntMake] functor
+ similar to the one of [FMapAVL], but richer.
+
+ - In final [IntMake_ord] functor, the [compare] function is
+ different from the one in [FMapAVL]: this non-structural
+ version is closer to the original Ocaml code.
+
+*)
+
+Require Import Recdef FMapInterface FMapList ZArith Int FMapAVL.
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+
+Module AvlProofs (Import I:Int)(X: OrderedType).
+Module Import Raw := Raw I X.
+Module Import II:=MoreInt(I).
+Import Raw.Proofs.
+Open Local Scope pair_scope.
+Open Local Scope Int_scope.
+
+Section Elt.
+Variable elt : Type.
+Implicit Types m r : t elt.
+
+(** * AVL trees *)
+
+(** [avl s] : [s] is a properly balanced AVL tree,
+ i.e. for any node the heights of the two children
+ differ by at most 2 *)
+
+Inductive avl : t elt -> Prop :=
+ | RBLeaf : avl (Leaf _)
+ | RBNode : forall x e l r h,
+ avl l ->
+ avl r ->
+ -(2) <= height l - height r <= 2 ->
+ h = max (height l) (height r) + 1 ->
+ avl (Node l x e r h).
+
+
+(** * Automation and dedicated tactics about [avl]. *)
+
+Hint Constructors avl.
+
+Lemma height_non_negative : forall (s : t elt), avl s ->
+ height s >= 0.
+Proof.
+ induction s; simpl; intros; auto with zarith.
+ inv avl; intuition; omega_max.
+Qed.
+
+Ltac avl_nn_hyp H :=
+ let nz := fresh "nz" in assert (nz := height_non_negative H).
+
+Ltac avl_nn h :=
+ let t := type of h in
+ match type of t with
+ | Prop => avl_nn_hyp h
+ | _ => match goal with H : avl h |- _ => avl_nn_hyp H end
+ end.
+
+(* Repeat the previous tactic.
+ Drawback: need to clear the [avl _] hyps ... Thank you Ltac *)
+
+Ltac avl_nns :=
+ match goal with
+ | H:avl _ |- _ => avl_nn_hyp H; clear H; avl_nns
+ | _ => idtac
+ end.
+
+
+(** * Basic results about [avl], [height] *)
+
+Lemma avl_node : forall x e l r, avl l -> avl r ->
+ -(2) <= height l - height r <= 2 ->
+ avl (Node l x e r (max (height l) (height r) + 1)).
+Proof.
+ intros; auto.
+Qed.
+Hint Resolve avl_node.
+
+(** Results about [height] *)
+
+Lemma height_0 : forall l, avl l -> height l = 0 ->
+ l = Leaf _.
+Proof.
+ destruct 1; intuition; simpl in *.
+ avl_nns; simpl in *; elimtype False; omega_max.
+Qed.
+
+
+(** * Empty map *)
+
+Lemma empty_avl : avl (empty elt).
+Proof.
+ unfold empty; auto.
+Qed.
+
+
+(** * Helper functions *)
+
+Lemma create_avl :
+ forall l x e r, avl l -> avl r -> -(2) <= height l - height r <= 2 ->
+ avl (create l x e r).
+Proof.
+ unfold create; auto.
+Qed.
+
+Lemma create_height :
+ forall l x e r, avl l -> avl r -> -(2) <= height l - height r <= 2 ->
+ height (create l x e r) = max (height l) (height r) + 1.
+Proof.
+ unfold create; intros; auto.
+Qed.
+
+Lemma bal_avl : forall l x e r, avl l -> avl r ->
+ -(3) <= height l - height r <= 3 -> avl (bal l x e r).
+Proof.
+ intros l x e r; functional induction (bal l x e r); intros; clearf;
+ inv avl; simpl in *;
+ match goal with |- avl (assert_false _ _ _ _) => avl_nns
+ | _ => repeat apply create_avl; simpl in *; auto
+ end; omega_max.
+Qed.
+
+Lemma bal_height_1 : forall l x e r, avl l -> avl r ->
+ -(3) <= height l - height r <= 3 ->
+ 0 <= height (bal l x e r) - max (height l) (height r) <= 1.
+Proof.
+ intros l x e r; functional induction (bal l x e r); intros; clearf;
+ inv avl; avl_nns; simpl in *; omega_max.
+Qed.
+
+Lemma bal_height_2 :
+ forall l x e r, avl l -> avl r -> -(2) <= height l - height r <= 2 ->
+ height (bal l x e r) == max (height l) (height r) +1.
+Proof.
+ intros l x e r; functional induction (bal l x e r); intros; clearf;
+ inv avl; avl_nns; simpl in *; omega_max.
+Qed.
+
+Ltac omega_bal := match goal with
+ | H:avl ?l, H':avl ?r |- context [ bal ?l ?x ?e ?r ] =>
+ generalize (bal_height_1 x e H H') (bal_height_2 x e H H');
+ omega_max
+ end.
+
+(** * Insertion *)
+
+Lemma add_avl_1 : forall m x e, avl m ->
+ avl (add x e m) /\ 0 <= height (add x e m) - height m <= 1.
+Proof.
+ intros m x e; functional induction (add x e m); intros; inv avl; simpl in *.
+ intuition; try constructor; simpl; auto; try omega_max.
+ (* LT *)
+ destruct IHt; auto.
+ split.
+ apply bal_avl; auto; omega_max.
+ omega_bal.
+ (* EQ *)
+ intuition; omega_max.
+ (* GT *)
+ destruct IHt; auto.
+ split.
+ apply bal_avl; auto; omega_max.
+ omega_bal.
+Qed.
+
+Lemma add_avl : forall m x e, avl m -> avl (add x e m).
+Proof.
+ intros; generalize (add_avl_1 x e H); intuition.
+Qed.
+Hint Resolve add_avl.
+
+(** * Extraction of minimum binding *)
+
+Lemma remove_min_avl_1 : forall l x e r h, avl (Node l x e r h) ->
+ avl (remove_min l x e r)#1 /\
+ 0 <= height (Node l x e r h) - height (remove_min l x e r)#1 <= 1.
+Proof.
+ intros l x e r; functional induction (remove_min l x e r); simpl in *; intros.
+ inv avl; simpl in *; split; auto.
+ avl_nns; omega_max.
+ inversion_clear H.
+ rewrite e0 in IHp;simpl in IHp;destruct (IHp _x); auto.
+ split; simpl in *.
+ apply bal_avl; auto; omega_max.
+ omega_bal.
+Qed.
+
+Lemma remove_min_avl : forall l x e r h, avl (Node l x e r h) ->
+ avl (remove_min l x e r)#1.
+Proof.
+ intros; generalize (remove_min_avl_1 H); intuition.
+Qed.
+
+(** * Merging two trees *)
+
+Lemma merge_avl_1 : forall m1 m2, avl m1 -> avl m2 ->
+ -(2) <= height m1 - height m2 <= 2 ->
+ avl (merge m1 m2) /\
+ 0<= height (merge m1 m2) - max (height m1) (height m2) <=1.
+Proof.
+ intros m1 m2; functional induction (merge m1 m2); intros;
+ try factornode _x _x0 _x1 _x2 _x3 as m1.
+ simpl; split; auto; avl_nns; omega_max.
+ simpl; split; auto; avl_nns; omega_max.
+ generalize (remove_min_avl_1 H0).
+ rewrite e1; destruct 1.
+ split.
+ apply bal_avl; auto.
+ omega_max.
+ omega_bal.
+Qed.
+
+Lemma merge_avl : forall m1 m2, avl m1 -> avl m2 ->
+ -(2) <= height m1 - height m2 <= 2 -> avl (merge m1 m2).
+Proof.
+ intros; generalize (merge_avl_1 H H0 H1); intuition.
+Qed.
+
+
+(** * Deletion *)
+
+Lemma remove_avl_1 : forall m x, avl m ->
+ avl (remove x m) /\ 0 <= height m - height (remove x m) <= 1.
+Proof.
+ intros m x; functional induction (remove x m); intros.
+ split; auto; omega_max.
+ (* LT *)
+ inv avl.
+ destruct (IHt H0).
+ split.
+ apply bal_avl; auto.
+ omega_max.
+ omega_bal.
+ (* EQ *)
+ inv avl.
+ generalize (merge_avl_1 H0 H1 H2).
+ intuition omega_max.
+ (* GT *)
+ inv avl.
+ destruct (IHt H1).
+ split.
+ apply bal_avl; auto.
+ omega_max.
+ omega_bal.
+Qed.
+
+Lemma remove_avl : forall m x, avl m -> avl (remove x m).
+Proof.
+ intros; generalize (remove_avl_1 x H); intuition.
+Qed.
+Hint Resolve remove_avl.
+
+
+(** * Join *)
+
+Lemma join_avl_1 : forall l x d r, avl l -> avl r ->
+ avl (join l x d r) /\
+ 0<= height (join l x d r) - max (height l) (height r) <= 1.
+Proof.
+ join_tac.
+
+ split; simpl; auto.
+ destruct (add_avl_1 x d H0).
+ avl_nns; omega_max.
+ set (l:=Node ll lx ld lr lh) in *.
+ split; auto.
+ destruct (add_avl_1 x d H).
+ simpl (height (Leaf elt)).
+ avl_nns; omega_max.
+
+ inversion_clear H.
+ assert (height (Node rl rx rd rr rh) = rh); auto.
+ set (r := Node rl rx rd rr rh) in *; clearbody r.
+ destruct (Hlr x d r H2 H0); clear Hrl Hlr.
+ set (j := join lr x d r) in *; clearbody j.
+ simpl.
+ assert (-(3) <= height ll - height j <= 3) by omega_max.
+ split.
+ apply bal_avl; auto.
+ omega_bal.
+
+ inversion_clear H0.
+ assert (height (Node ll lx ld lr lh) = lh); auto.
+ set (l := Node ll lx ld lr lh) in *; clearbody l.
+ destruct (Hrl H H1); clear Hrl Hlr.
+ set (j := join l x d rl) in *; clearbody j.
+ simpl.
+ assert (-(3) <= height j - height rr <= 3) by omega_max.
+ split.
+ apply bal_avl; auto.
+ omega_bal.
+
+ clear Hrl Hlr.
+ assert (height (Node ll lx ld lr lh) = lh); auto.
+ assert (height (Node rl rx rd rr rh) = rh); auto.
+ set (l := Node ll lx ld lr lh) in *; clearbody l.
+ set (r := Node rl rx rd rr rh) in *; clearbody r.
+ assert (-(2) <= height l - height r <= 2) by omega_max.
+ split.
+ apply create_avl; auto.
+ rewrite create_height; auto; omega_max.
+Qed.
+
+Lemma join_avl : forall l x d r, avl l -> avl r -> avl (join l x d r).
+Proof.
+ intros; destruct (join_avl_1 x d H H0); auto.
+Qed.
+Hint Resolve join_avl.
+
+(** concat *)
+
+Lemma concat_avl : forall m1 m2, avl m1 -> avl m2 -> avl (concat m1 m2).
+Proof.
+ intros m1 m2; functional induction (concat m1 m2); auto.
+ intros; apply join_avl; auto.
+ generalize (remove_min_avl H0); rewrite e1; simpl; auto.
+Qed.
+Hint Resolve concat_avl.
+
+(** split *)
+
+Lemma split_avl : forall m x, avl m ->
+ avl (split x m)#l /\ avl (split x m)#r.
+Proof.
+ intros m x; functional induction (split x m); simpl; auto.
+ rewrite e1 in IHt;simpl in IHt;inversion_clear 1; intuition.
+ simpl; inversion_clear 1; auto.
+ rewrite e1 in IHt;simpl in IHt;inversion_clear 1; intuition.
+Qed.
+
+End Elt.
+Hint Constructors avl.
+
+Section Map.
+Variable elt elt' : Type.
+Variable f : elt -> elt'.
+
+Lemma map_height : forall m, height (map f m) = height m.
+Proof.
+destruct m; simpl; auto.
+Qed.
+
+Lemma map_avl : forall m, avl m -> avl (map f m).
+Proof.
+induction m; simpl; auto.
+inversion_clear 1; constructor; auto; do 2 rewrite map_height; auto.
+Qed.
+
+End Map.
+
+Section Mapi.
+Variable elt elt' : Type.
+Variable f : key -> elt -> elt'.
+
+Lemma mapi_height : forall m, height (mapi f m) = height m.
+Proof.
+destruct m; simpl; auto.
+Qed.
+
+Lemma mapi_avl : forall m, avl m -> avl (mapi f m).
+Proof.
+induction m; simpl; auto.
+inversion_clear 1; constructor; auto; do 2 rewrite mapi_height; auto.
+Qed.
+
+End Mapi.
+
+Section Map_option.
+Variable elt elt' : Type.
+Variable f : key -> elt -> option elt'.
+
+Lemma map_option_avl : forall m, avl m -> avl (map_option f m).
+Proof.
+induction m; simpl; auto; intros.
+inv avl; destruct (f k e); auto using join_avl, concat_avl.
+Qed.
+
+End Map_option.
+
+Section Map2_opt.
+Variable elt elt' elt'' : Type.
+Variable f : key -> elt -> option elt' -> option elt''.
+Variable mapl : t elt -> t elt''.
+Variable mapr : t elt' -> t elt''.
+Hypothesis mapl_avl : forall m, avl m -> avl (mapl m).
+Hypothesis mapr_avl : forall m', avl m' -> avl (mapr m').
+
+Notation map2_opt := (map2_opt f mapl mapr).
+
+Lemma map2_opt_avl : forall m1 m2, avl m1 -> avl m2 ->
+ avl (map2_opt m1 m2).
+Proof.
+intros m1 m2; functional induction (map2_opt m1 m2); auto;
+factornode _x0 _x1 _x2 _x3 _x4 as r2; intros;
+destruct (split_avl x1 H0); rewrite e1 in *; simpl in *; inv avl;
+auto using join_avl, concat_avl.
+Qed.
+
+End Map2_opt.
+
+Section Map2.
+Variable elt elt' elt'' : Type.
+Variable f : option elt -> option elt' -> option elt''.
+
+Lemma map2_avl : forall m1 m2, avl m1 -> avl m2 -> avl (map2 f m1 m2).
+Proof.
+unfold map2; auto using map2_opt_avl, map_option_avl.
+Qed.
+
+End Map2.
+End AvlProofs.
+
+(** * Encapsulation
+
+ We can implement [S] with balanced binary search trees.
+ When compared to [FMapAVL], we maintain here two invariants
+ (bst and avl) instead of only bst, which is enough for fulfilling
+ the FMap interface.
+*)
+
+Module IntMake (I:Int)(X: OrderedType) <: S with Module E := X.
+
+ Module E := X.
+ Module Import AvlProofs := AvlProofs I X.
+ Import Raw.
+ Import Raw.Proofs.
+
+ Record bbst (elt:Type) :=
+ Bbst {this :> tree elt; is_bst : bst this; is_avl: avl this}.
+
+ Definition t := bbst.
+ Definition key := E.t.
+
+ Section Elt.
+ Variable elt elt' elt'': Type.
+
+ Implicit Types m : t elt.
+ Implicit Types x y : key.
+ Implicit Types e : elt.
+
+ Definition empty : t elt := Bbst (empty_bst elt) (empty_avl elt).
+ Definition is_empty m : bool := is_empty m.(this).
+ Definition add x e m : t elt :=
+ Bbst (add_bst x e m.(is_bst)) (add_avl x e m.(is_avl)).
+ Definition remove x m : t elt :=
+ Bbst (remove_bst x m.(is_bst)) (remove_avl x m.(is_avl)).
+ Definition mem x m : bool := mem x m.(this).
+ Definition find x m : option elt := find x m.(this).
+ Definition map f m : t elt' :=
+ Bbst (map_bst f m.(is_bst)) (map_avl f m.(is_avl)).
+ Definition mapi (f:key->elt->elt') m : t elt' :=
+ Bbst (mapi_bst f m.(is_bst)) (mapi_avl f m.(is_avl)).
+ Definition map2 f m (m':t elt') : t elt'' :=
+ Bbst (map2_bst f m.(is_bst) m'.(is_bst)) (map2_avl f m.(is_avl) m'.(is_avl)).
+ Definition elements m : list (key*elt) := elements m.(this).
+ Definition cardinal m := cardinal m.(this).
+ Definition fold (A:Type) (f:key->elt->A->A) m i := fold (A:=A) f m.(this) i.
+ Definition equal cmp m m' : bool := equal cmp m.(this) m'.(this).
+
+ Definition MapsTo x e m : Prop := MapsTo x e m.(this).
+ Definition In x m : Prop := In0 x m.(this).
+ Definition Empty m : Prop := Empty m.(this).
+
+ Definition eq_key : (key*elt) -> (key*elt) -> Prop := @PX.eqk elt.
+ Definition eq_key_elt : (key*elt) -> (key*elt) -> Prop := @PX.eqke elt.
+ Definition lt_key : (key*elt) -> (key*elt) -> Prop := @PX.ltk elt.
+
+ Lemma MapsTo_1 : forall m x y e, E.eq x y -> MapsTo x e m -> MapsTo y e m.
+ Proof. intros m; exact (@MapsTo_1 _ m.(this)). Qed.
+
+ Lemma mem_1 : forall m x, In x m -> mem x m = true.
+ Proof.
+ unfold In, mem; intros m x; rewrite In_alt; simpl; apply mem_1; auto.
+ apply m.(is_bst).
+ Qed.
+
+ Lemma mem_2 : forall m x, mem x m = true -> In x m.
+ Proof.
+ unfold In, mem; intros m x; rewrite In_alt; simpl; apply mem_2; auto.
+ Qed.
+
+ Lemma empty_1 : Empty empty.
+ Proof. exact (@empty_1 elt). Qed.
+
+ Lemma is_empty_1 : forall m, Empty m -> is_empty m = true.
+ Proof. intros m; exact (@is_empty_1 _ m.(this)). Qed.
+ Lemma is_empty_2 : forall m, is_empty m = true -> Empty m.
+ Proof. intros m; exact (@is_empty_2 _ m.(this)). Qed.
+
+ Lemma add_1 : forall m x y e, E.eq x y -> MapsTo y e (add x e m).
+ Proof. intros m x y e; exact (@add_1 elt _ x y e). Qed.
+ Lemma add_2 : forall m x y e e', ~ E.eq x y -> MapsTo y e m -> MapsTo y e (add x e' m).
+ Proof. intros m x y e e'; exact (@add_2 elt _ x y e e'). Qed.
+ Lemma add_3 : forall m x y e e', ~ E.eq x y -> MapsTo y e (add x e' m) -> MapsTo y e m.
+ Proof. intros m x y e e'; exact (@add_3 elt _ x y e e'). Qed.
+
+ Lemma remove_1 : forall m x y, E.eq x y -> ~ In y (remove x m).
+ Proof.
+ unfold In, remove; intros m x y; rewrite In_alt; simpl; apply remove_1; auto.
+ apply m.(is_bst).
+ Qed.
+ Lemma remove_2 : forall m x y e, ~ E.eq x y -> MapsTo y e m -> MapsTo y e (remove x m).
+ Proof. intros m x y e; exact (@remove_2 elt _ x y e m.(is_bst)). Qed.
+ Lemma remove_3 : forall m x y e, MapsTo y e (remove x m) -> MapsTo y e m.
+ Proof. intros m x y e; exact (@remove_3 elt _ x y e m.(is_bst)). Qed.
+
+
+ Lemma find_1 : forall m x e, MapsTo x e m -> find x m = Some e.
+ Proof. intros m x e; exact (@find_1 elt _ x e m.(is_bst)). Qed.
+ Lemma find_2 : forall m x e, find x m = Some e -> MapsTo x e m.
+ Proof. intros m; exact (@find_2 elt m.(this)). Qed.
+
+ Lemma fold_1 : forall m (A : Type) (i : A) (f : key -> elt -> A -> A),
+ fold f m i = fold_left (fun a p => f (fst p) (snd p) a) (elements m) i.
+ Proof. intros m; exact (@fold_1 elt m.(this) m.(is_bst)). Qed.
+
+ Lemma elements_1 : forall m x e,
+ MapsTo x e m -> InA eq_key_elt (x,e) (elements m).
+ Proof.
+ intros; unfold elements, MapsTo, eq_key_elt; rewrite elements_mapsto; auto.
+ Qed.
+
+ Lemma elements_2 : forall m x e,
+ InA eq_key_elt (x,e) (elements m) -> MapsTo x e m.
+ Proof.
+ intros; unfold elements, MapsTo, eq_key_elt; rewrite <- elements_mapsto; auto.
+ Qed.
+
+ Lemma elements_3 : forall m, sort lt_key (elements m).
+ Proof. intros m; exact (@elements_sort elt m.(this) m.(is_bst)). Qed.
+
+ Lemma elements_3w : forall m, NoDupA eq_key (elements m).
+ Proof. intros m; exact (@elements_nodup elt m.(this) m.(is_bst)). Qed.
+
+ Lemma cardinal_1 : forall m, cardinal m = length (elements m).
+ Proof. intro m; exact (@elements_cardinal elt m.(this)). Qed.
+
+ Definition Equal m m' := forall y, find y m = find y m'.
+ Definition Equiv (eq_elt:elt->elt->Prop) m m' :=
+ (forall k, In k m <-> In k m') /\
+ (forall k e e', MapsTo k e m -> MapsTo k e' m' -> eq_elt e e').
+ Definition Equivb cmp := Equiv (Cmp cmp).
+
+ Lemma Equivb_Equivb : forall cmp m m',
+ Equivb cmp m m' <-> Raw.Proofs.Equivb cmp m m'.
+ Proof.
+ intros; unfold Equivb, Equiv, Raw.Proofs.Equivb, In; intuition.
+ generalize (H0 k); do 2 rewrite In_alt; intuition.
+ generalize (H0 k); do 2 rewrite In_alt; intuition.
+ generalize (H0 k); do 2 rewrite <- In_alt; intuition.
+ generalize (H0 k); do 2 rewrite <- In_alt; intuition.
+ Qed.
+
+ Lemma equal_1 : forall m m' cmp,
+ Equivb cmp m m' -> equal cmp m m' = true.
+ Proof.
+ unfold equal; intros (m,b,a) (m',b',a') cmp; rewrite Equivb_Equivb;
+ intros; simpl in *; rewrite equal_Equivb; auto.
+ Qed.
+
+ Lemma equal_2 : forall m m' cmp,
+ equal cmp m m' = true -> Equivb cmp m m'.
+ Proof.
+ unfold equal; intros (m,b,a) (m',b',a') cmp; rewrite Equivb_Equivb;
+ intros; simpl in *; rewrite <-equal_Equivb; auto.
+ Qed.
+
+ End Elt.
+
+ Lemma map_1 : forall (elt elt':Type)(m: t elt)(x:key)(e:elt)(f:elt->elt'),
+ MapsTo x e m -> MapsTo x (f e) (map f m).
+ Proof. intros elt elt' m x e f; exact (@map_1 elt elt' f m.(this) x e). Qed.
+
+ Lemma map_2 : forall (elt elt':Type)(m:t elt)(x:key)(f:elt->elt'), In x (map f m) -> In x m.
+ Proof.
+ intros elt elt' m x f; do 2 unfold In in *; do 2 rewrite In_alt; simpl.
+ apply map_2; auto.
+ Qed.
+
+ Lemma mapi_1 : forall (elt elt':Type)(m: t elt)(x:key)(e:elt)
+ (f:key->elt->elt'), MapsTo x e m ->
+ exists y, E.eq y x /\ MapsTo x (f y e) (mapi f m).
+ Proof. intros elt elt' m x e f; exact (@mapi_1 elt elt' f m.(this) x e). Qed.
+ Lemma mapi_2 : forall (elt elt':Type)(m: t elt)(x:key)
+ (f:key->elt->elt'), In x (mapi f m) -> In x m.
+ Proof.
+ intros elt elt' m x f; unfold In in *; do 2 rewrite In_alt; simpl; apply mapi_2; auto.
+ Qed.
+
+ Lemma map2_1 : forall (elt elt' elt'':Type)(m: t elt)(m': t elt')
+ (x:key)(f:option elt->option elt'->option elt''),
+ In x m \/ In x m' ->
+ find x (map2 f m m') = f (find x m) (find x m').
+ Proof.
+ unfold find, map2, In; intros elt elt' elt'' m m' x f.
+ do 2 rewrite In_alt; intros; simpl; apply map2_1; auto.
+ apply m.(is_bst).
+ apply m'.(is_bst).
+ Qed.
+
+ Lemma map2_2 : forall (elt elt' elt'':Type)(m: t elt)(m': t elt')
+ (x:key)(f:option elt->option elt'->option elt''),
+ In x (map2 f m m') -> In x m \/ In x m'.
+ Proof.
+ unfold In, map2; intros elt elt' elt'' m m' x f.
+ do 3 rewrite In_alt; intros; simpl in *; eapply map2_2; eauto.
+ apply m.(is_bst).
+ apply m'.(is_bst).
+ Qed.
+
+End IntMake.
+
+
+Module IntMake_ord (I:Int)(X: OrderedType)(D : OrderedType) <:
+ Sord with Module Data := D
+ with Module MapS.E := X.
+
+ Module Data := D.
+ Module Import MapS := IntMake(I)(X).
+ Import AvlProofs.
+ Import Raw.Proofs.
+ Module Import MD := OrderedTypeFacts(D).
+ Module LO := FMapList.Make_ord(X)(D).
+
+ Definition t := MapS.t D.t.
+
+ Definition cmp e e' :=
+ match D.compare e e' with EQ _ => true | _ => false end.
+
+ Definition elements (m:t) :=
+ LO.MapS.Build_slist (Raw.Proofs.elements_sort m.(is_bst)).
+
+ (** * As comparison function, we propose here a non-structural
+ version faithful to the code of Ocaml's Map library, instead of
+ the structural version of FMapAVL *)
+
+ Fixpoint cardinal_e (e:Raw.enumeration D.t) :=
+ match e with
+ | Raw.End => 0%nat
+ | Raw.More _ _ r e => S (Raw.cardinal r + cardinal_e e)
+ end.
+
+ Lemma cons_cardinal_e : forall m e,
+ cardinal_e (Raw.cons m e) = (Raw.cardinal m + cardinal_e e)%nat.
+ Proof.
+ induction m; simpl; intros; auto.
+ rewrite IHm1; simpl; rewrite <- plus_n_Sm; auto with arith.
+ Qed.
+
+ Definition cardinal_e_2 ee :=
+ (cardinal_e (fst ee) + cardinal_e (snd ee))%nat.
+
+ Function compare_aux (ee:Raw.enumeration D.t * Raw.enumeration D.t)
+ { measure cardinal_e_2 ee } : comparison :=
+ match ee with
+ | (Raw.End, Raw.End) => Eq
+ | (Raw.End, Raw.More _ _ _ _) => Lt
+ | (Raw.More _ _ _ _, Raw.End) => Gt
+ | (Raw.More x1 d1 r1 e1, Raw.More x2 d2 r2 e2) =>
+ match X.compare x1 x2 with
+ | EQ _ => match D.compare d1 d2 with
+ | EQ _ => compare_aux (Raw.cons r1 e1, Raw.cons r2 e2)
+ | LT _ => Lt
+ | GT _ => Gt
+ end
+ | LT _ => Lt
+ | GT _ => Gt
+ end
+ end.
+ Proof.
+ intros; unfold cardinal_e_2; simpl;
+ abstract (do 2 rewrite cons_cardinal_e; romega with * ).
+ Defined.
+
+ Definition Cmp c :=
+ match c with
+ | Eq => LO.eq_list
+ | Lt => LO.lt_list
+ | Gt => (fun l1 l2 => LO.lt_list l2 l1)
+ end.
+
+ Lemma cons_Cmp : forall c x1 x2 d1 d2 l1 l2,
+ X.eq x1 x2 -> D.eq d1 d2 ->
+ Cmp c l1 l2 -> Cmp c ((x1,d1)::l1) ((x2,d2)::l2).
+ Proof.
+ destruct c; simpl; intros; MX.elim_comp; auto.
+ Qed.
+ Hint Resolve cons_Cmp.
+
+ Lemma compare_aux_Cmp : forall e,
+ Cmp (compare_aux e) (flatten_e (fst e)) (flatten_e (snd e)).
+ Proof.
+ intros e; functional induction (compare_aux e); simpl in *;
+ auto; intros; try clear e0; try clear e3; try MX.elim_comp; auto.
+ rewrite 2 cons_1 in IHc; auto.
+ Qed.
+
+ Lemma compare_Cmp : forall m1 m2,
+ Cmp (compare_aux (Raw.cons m1 (Raw.End _), Raw.cons m2 (Raw.End _)))
+ (Raw.elements m1) (Raw.elements m2).
+ Proof.
+ intros.
+ assert (H1:=cons_1 m1 (Raw.End _)).
+ assert (H2:=cons_1 m2 (Raw.End _)).
+ simpl in *; rewrite <- app_nil_end in *; rewrite <-H1,<-H2.
+ apply (@compare_aux_Cmp (Raw.cons m1 (Raw.End _),
+ Raw.cons m2 (Raw.End _))).
+ Qed.
+
+ Definition eq (m1 m2 : t) := LO.eq_list (Raw.elements m1) (Raw.elements m2).
+ Definition lt (m1 m2 : t) := LO.lt_list (Raw.elements m1) (Raw.elements m2).
+
+ Definition compare (s s':t) : Compare lt eq s s'.
+ Proof.
+ intros (s,b,a) (s',b',a').
+ generalize (compare_Cmp s s').
+ destruct compare_aux; intros; [apply EQ|apply LT|apply GT]; red; auto.
+ Defined.
+
+
+ (* Proofs about [eq] and [lt] *)
+
+ Definition selements (m1 : t) :=
+ LO.MapS.Build_slist (elements_sort m1.(is_bst)).
+
+ Definition seq (m1 m2 : t) := LO.eq (selements m1) (selements m2).
+ Definition slt (m1 m2 : t) := LO.lt (selements m1) (selements m2).
+
+ Lemma eq_seq : forall m1 m2, eq m1 m2 <-> seq m1 m2.
+ Proof.
+ unfold eq, seq, selements, elements, LO.eq; intuition.
+ Qed.
+
+ Lemma lt_slt : forall m1 m2, lt m1 m2 <-> slt m1 m2.
+ Proof.
+ unfold lt, slt, selements, elements, LO.lt; intuition.
+ Qed.
+
+ Lemma eq_1 : forall (m m' : t), MapS.Equivb cmp m m' -> eq m m'.
+ Proof.
+ intros m m'.
+ rewrite eq_seq; unfold seq.
+ rewrite Equivb_Equivb.
+ rewrite Equivb_elements.
+ auto using LO.eq_1.
+ Qed.
+
+ Lemma eq_2 : forall m m', eq m m' -> MapS.Equivb cmp m m'.
+ Proof.
+ intros m m'.
+ rewrite eq_seq; unfold seq.
+ rewrite Equivb_Equivb.
+ rewrite Equivb_elements.
+ intros.
+ generalize (LO.eq_2 H).
+ auto.
+ Qed.
+
+ Lemma eq_refl : forall m : t, eq m m.
+ Proof.
+ intros; rewrite eq_seq; unfold seq; intros; apply LO.eq_refl.
+ Qed.
+
+ Lemma eq_sym : forall m1 m2 : t, eq m1 m2 -> eq m2 m1.
+ Proof.
+ intros m1 m2; rewrite 2 eq_seq; unfold seq; intros; apply LO.eq_sym; auto.
+ Qed.
+
+ Lemma eq_trans : forall m1 m2 m3 : t, eq m1 m2 -> eq m2 m3 -> eq m1 m3.
+ Proof.
+ intros m1 m2 M3; rewrite 3 eq_seq; unfold seq.
+ intros; eapply LO.eq_trans; eauto.
+ Qed.
+
+ Lemma lt_trans : forall m1 m2 m3 : t, lt m1 m2 -> lt m2 m3 -> lt m1 m3.
+ Proof.
+ intros m1 m2 m3; rewrite 3 lt_slt; unfold slt;
+ intros; eapply LO.lt_trans; eauto.
+ Qed.
+
+ Lemma lt_not_eq : forall m1 m2 : t, lt m1 m2 -> ~ eq m1 m2.
+ Proof.
+ intros m1 m2; rewrite lt_slt, eq_seq; unfold slt, seq;
+ intros; apply LO.lt_not_eq; auto.
+ Qed.
+
+End IntMake_ord.
+
+(* 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).
+
+Module Make_ord (X: OrderedType)(D: OrderedType)
+ <: Sord with Module Data := D
+ with Module MapS.E := X
+ :=IntMake_ord(Z_as_Int)(X)(D).
+
diff --git a/theories/FSets/FMapIntMap.v b/theories/FSets/FMapIntMap.v
deleted file mode 100644
index c7681bd4..00000000
--- a/theories/FSets/FMapIntMap.v
+++ /dev/null
@@ -1,622 +0,0 @@
-(***********************************************************************)
-(* 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 *)
-(***********************************************************************)
-
-(* Finite sets library.
- * Authors: Pierre Letouzey and Jean-Christophe Filliâtre
- * Institution: LRI, CNRS UMR 8623 - Université Paris Sud
- * 91405 Orsay, France *)
-
-(* $Id: FMapIntMap.v 8876 2006-05-30 13:43:15Z letouzey $ *)
-
-Require Import Bool.
-Require Import NArith Ndigits Ndec Nnat.
-Require Import Allmaps.
-Require Import OrderedType.
-Require Import OrderedTypeEx.
-Require Import FMapInterface FMapList.
-
-
-Set Implicit Arguments.
-
-(** * An implementation of [FMapInterface.S] based on [IntMap] *)
-
-(** Keys are of type [N]. The main functions are directly taken from
- [IntMap]. Since they have no exact counterpart in [IntMap], functions
- [fold], [map2] and [equal] are for now obtained by translation
- to sorted lists. *)
-
-(** [N] is an ordered type, using not the usual order on numbers,
- but lexicographic ordering on bits (lower bit considered first). *)
-
-Module NUsualOrderedType <: UsualOrderedType.
- Definition t:=N.
- Definition eq:=@eq N.
- Definition eq_refl := @refl_equal t.
- Definition eq_sym := @sym_eq t.
- Definition eq_trans := @trans_eq t.
-
- Definition lt p q:= Nless p q = true.
-
- Definition lt_trans := Nless_trans.
-
- Lemma lt_not_eq : forall x y : t, lt x y -> ~ eq x y.
- Proof.
- intros; intro.
- rewrite H0 in H.
- red in H.
- rewrite Nless_not_refl in H; discriminate.
- Qed.
-
- Definition compare : forall x y : t, Compare lt eq x y.
- Proof.
- intros x y.
- destruct (Nless_total x y) as [[H|H]|H].
- apply LT; unfold lt; auto.
- apply GT; unfold lt; auto.
- apply EQ; auto.
- Qed.
-
-End NUsualOrderedType.
-
-
-(** The module of maps over [N] keys based on [IntMap] *)
-
-Module MapIntMap <: S with Module E:=NUsualOrderedType.
-
- Module E:=NUsualOrderedType.
- Module ME:=OrderedTypeFacts(E).
- Module PE:=KeyOrderedType(E).
-
- Definition key := N.
-
- Definition t := Map.
-
- Section A.
- Variable A:Set.
-
- Definition empty : t A := M0 A.
-
- Definition is_empty (m : t A) : bool :=
- MapEmptyp _ (MapCanonicalize _ m).
-
- Definition find (x:key)(m: t A) : option A := MapGet _ m x.
-
- Definition mem (x:key)(m: t A) : bool :=
- match find x m with
- | Some _ => true
- | None => false
- end.
-
- Definition add (x:key)(v:A)(m:t A) : t A := MapPut _ m x v.
-
- Definition remove (x:key)(m:t A) : t A := MapRemove _ m x.
-
- Definition elements (m : t A) : list (N*A) := alist_of_Map _ m.
-
- Definition MapsTo (x:key)(v:A)(m:t A) := find x m = Some v.
-
- Definition In (x:key)(m:t A) := exists e:A, MapsTo x e m.
-
- Definition Empty m := forall (a : key)(e:A) , ~ MapsTo a e m.
-
- Definition eq_key (p p':key*A) := E.eq (fst p) (fst p').
-
- Definition eq_key_elt (p p':key*A) :=
- E.eq (fst p) (fst p') /\ (snd p) = (snd p').
-
- Definition lt_key (p p':key*A) := E.lt (fst p) (fst p').
-
- Lemma Empty_alt : forall m, Empty m <-> forall a, find a m = None.
- Proof.
- unfold Empty, MapsTo.
- intuition.
- generalize (H a).
- destruct (find a m); intuition.
- elim (H0 a0); auto.
- rewrite H in H0; discriminate.
- Qed.
-
- Section Spec.
- Variable m m' m'' : t A.
- Variable x y z : key.
- Variable e e' : A.
-
- Lemma MapsTo_1 : E.eq x y -> MapsTo x e m -> MapsTo y e m.
- Proof. intros; rewrite <- H; auto. Qed.
-
- Lemma find_1 : MapsTo x e m -> find x m = Some e.
- Proof. unfold MapsTo; auto. Qed.
-
- Lemma find_2 : find x m = Some e -> MapsTo x e m.
- Proof. red; auto. Qed.
-
- Lemma empty_1 : Empty empty.
- Proof.
- rewrite Empty_alt; intros; unfold empty, find; simpl; auto.
- Qed.
-
- Lemma is_empty_1 : Empty m -> is_empty m = true.
- Proof.
- unfold Empty, is_empty, find; intros.
- cut (MapCanonicalize _ m = M0 _).
- intros; rewrite H0; simpl; auto.
- apply mapcanon_unique.
- apply mapcanon_exists_2.
- constructor.
- red; red; simpl; intros.
- rewrite <- (mapcanon_exists_1 _ m).
- unfold MapsTo, find in *.
- generalize (H a).
- destruct (MapGet _ m a); auto.
- intros; generalize (H0 a0); destruct 1; auto.
- Qed.
-
- Lemma is_empty_2 : is_empty m = true -> Empty m.
- Proof.
- unfold Empty, is_empty, MapsTo, find; intros.
- generalize (MapEmptyp_complete _ _ H); clear H; intros.
- rewrite (mapcanon_exists_1 _ m).
- rewrite H; simpl; auto.
- discriminate.
- Qed.
-
- Lemma mem_1 : In x m -> mem x m = true.
- Proof.
- unfold In, MapsTo, mem.
- destruct (find x m); auto.
- destruct 1; discriminate.
- Qed.
-
- Lemma mem_2 : forall m x, mem x m = true -> In x m.
- Proof.
- unfold In, MapsTo, mem.
- intros.
- destruct (find x0 m0); auto; try discriminate.
- exists a; auto.
- Qed.
-
- Lemma add_1 : E.eq x y -> MapsTo y e (add x e m).
- Proof.
- unfold MapsTo, find, add.
- intro H; rewrite H; clear H.
- rewrite MapPut_semantics.
- rewrite Neqb_correct; auto.
- Qed.
-
- Lemma add_2 : ~ E.eq x y -> MapsTo y e m -> MapsTo y e (add x e' m).
- Proof.
- unfold MapsTo, find, add.
- intros.
- rewrite MapPut_semantics.
- rewrite H0.
- generalize (Neqb_complete x y).
- destruct (Neqb x y); auto.
- intros.
- elim H; auto.
- apply H1; auto.
- Qed.
-
- Lemma add_3 : ~ E.eq x y -> MapsTo y e (add x e' m) -> MapsTo y e m.
- Proof.
- unfold MapsTo, find, add.
- rewrite MapPut_semantics.
- intro H.
- generalize (Neqb_complete x y).
- destruct (Neqb x y); auto.
- intros; elim H; auto.
- apply H0; auto.
- Qed.
-
- Lemma remove_1 : E.eq x y -> ~ In y (remove x m).
- Proof.
- unfold In, MapsTo, find, remove.
- rewrite MapRemove_semantics.
- intro H.
- rewrite H; rewrite Neqb_correct.
- red; destruct 1; discriminate.
- Qed.
-
- Lemma remove_2 : ~ E.eq x y -> MapsTo y e m -> MapsTo y e (remove x m).
- Proof.
- unfold MapsTo, find, remove.
- rewrite MapRemove_semantics.
- intros.
- rewrite H0.
- generalize (Neqb_complete x y).
- destruct (Neqb x y); auto.
- intros; elim H; apply H1; auto.
- Qed.
-
- Lemma remove_3 : MapsTo y e (remove x m) -> MapsTo y e m.
- Proof.
- unfold MapsTo, find, remove.
- rewrite MapRemove_semantics.
- destruct (Neqb x y); intros; auto.
- discriminate.
- Qed.
-
- Lemma alist_sorted_sort : forall l, alist_sorted A l=true -> sort lt_key l.
- Proof.
- induction l.
- auto.
- simpl.
- destruct a.
- destruct l.
- auto.
- destruct p.
- intros; destruct (andb_prop _ _ H); auto.
- Qed.
-
- Lemma elements_3 : sort lt_key (elements m).
- Proof.
- unfold elements.
- apply alist_sorted_sort.
- apply alist_of_Map_sorts.
- Qed.
-
- Lemma elements_1 :
- MapsTo x e m -> InA eq_key_elt (x,e) (elements m).
- Proof.
- unfold MapsTo, find, elements.
- rewrite InA_alt.
- intro H.
- exists (x,e).
- split.
- red; simpl; unfold E.eq; auto.
- rewrite alist_of_Map_semantics in H.
- generalize H.
- set (l:=alist_of_Map A m); clearbody l; clear.
- induction l; simpl; auto.
- intro; discriminate.
- destruct a; simpl; auto.
- generalize (Neqb_complete a x).
- destruct (Neqb a x); auto.
- left.
- injection H0; auto.
- intros; f_equal; auto.
- Qed.
-
- Lemma elements_2 :
- InA eq_key_elt (x,e) (elements m) -> MapsTo x e m.
- Proof.
- generalize elements_3.
- unfold MapsTo, find, elements.
- rewrite InA_alt.
- intros H ((e0,a),(H0,H1)).
- red in H0; simpl in H0; unfold E.eq in H0; destruct H0; subst.
- rewrite alist_of_Map_semantics.
- generalize H H1; clear H H1.
- set (l:=alist_of_Map A m); clearbody l; clear.
- induction l; simpl; auto.
- intro; contradiction.
- intros.
- destruct a0; simpl.
- inversion H1.
- injection H0; intros; subst.
- rewrite Neqb_correct; auto.
- assert (InA eq_key (e0,a) l).
- rewrite InA_alt.
- exists (e0,a); split; auto.
- red; simpl; auto; red; auto.
- generalize (PE.Sort_In_cons_1 H H2).
- unfold PE.ltk; simpl.
- intros H3; generalize (E.lt_not_eq H3).
- generalize (Neqb_complete a0 e0).
- destruct (Neqb a0 e0); auto.
- destruct 2.
- apply H4; auto.
- inversion H; auto.
- Qed.
-
- Definition Equal cmp m m' :=
- (forall k, In k m <-> In k m') /\
- (forall k e e', MapsTo k e m -> MapsTo k e' m' -> cmp e e' = true).
-
- (** unfortunately, the [MapFold] of [IntMap] isn't compatible with
- the FMap interface. We use a naive version for now : *)
-
- Definition fold (B:Set)(f:key -> A -> B -> B)(m:t A)(i:B) : B :=
- fold_left (fun a p => f (fst p) (snd p) a) (elements m) i.
-
- Lemma fold_1 :
- forall (B:Set) (i : B) (f : key -> A -> B -> B),
- fold f m i = fold_left (fun a p => f (fst p) (snd p) a) (elements m) i.
- Proof. auto. Qed.
-
- End Spec.
-
- Variable B : Set.
-
- Fixpoint mapi_aux (pf:N->N)(f : N -> A -> B)(m:t A) { struct m }: t B :=
- match m with
- | M0 => M0 _
- | M1 x y => M1 _ x (f (pf x) y)
- | M2 m0 m1 => M2 _ (mapi_aux (fun n => pf (Ndouble n)) f m0)
- (mapi_aux (fun n => pf (Ndouble_plus_one n)) f m1)
- end.
-
- Definition mapi := mapi_aux (fun n => n).
-
- Definition map (f:A->B) := mapi (fun _ => f).
-
- End A.
-
- Lemma mapi_aux_1 : forall (elt elt':Set)(m: t elt)(pf:N->N)(x:key)(e:elt)
- (f:key->elt->elt'), MapsTo x e m ->
- exists y, E.eq y x /\ MapsTo x (f (pf y) e) (mapi_aux pf f m).
- Proof.
- unfold MapsTo; induction m; simpl; auto.
- inversion 1.
-
- intros.
- exists x; split; [red; auto|].
- generalize (Neqb_complete a x).
- destruct (Neqb a x); try discriminate.
- injection H; intros; subst; auto.
- rewrite H1; auto.
-
- intros.
- exists x; split; [red;auto|].
- destruct x; simpl in *.
- destruct (IHm1 (fun n : N => pf (Ndouble n)) _ _ f H) as (y,(Hy,Hy')).
- rewrite Hy in Hy'; simpl in Hy'; auto.
- destruct p; simpl in *.
- destruct (IHm2 (fun n : N => pf (Ndouble_plus_one n)) _ _ f H) as (y,(Hy,Hy')).
- rewrite Hy in Hy'; simpl in Hy'; auto.
- destruct (IHm1 (fun n : N => pf (Ndouble n)) _ _ f H) as (y,(Hy,Hy')).
- rewrite Hy in Hy'; simpl in Hy'; auto.
- destruct (IHm2 (fun n : N => pf (Ndouble_plus_one n)) _ _ f H) as (y,(Hy,Hy')).
- rewrite Hy in Hy'; simpl in Hy'; auto.
- Qed.
-
- Lemma mapi_1 : forall (elt elt':Set)(m: t elt)(x:key)(e:elt)
- (f:key->elt->elt'), MapsTo x e m ->
- exists y, E.eq y x /\ MapsTo x (f y e) (mapi f m).
- Proof.
- intros elt elt' m; exact (mapi_aux_1 (fun n => n)).
- Qed.
-
- Lemma mapi_aux_2 : forall (elt elt':Set)(m: t elt)(pf:N->N)(x:key)
- (f:key->elt->elt'), In x (mapi_aux pf f m) -> In x m.
- Proof.
- unfold In, MapsTo.
- induction m; simpl in *.
- intros pf x f (e,He); inversion He.
- intros pf x f (e,He).
- exists a0.
- destruct (Neqb a x); try discriminate; auto.
- intros pf x f (e,He).
- destruct x; [|destruct p]; eauto.
- Qed.
-
- Lemma mapi_2 : forall (elt elt':Set)(m: t elt)(x:key)
- (f:key->elt->elt'), In x (mapi f m) -> In x m.
- Proof.
- intros elt elt' m; exact (mapi_aux_2 m (fun n => n)).
- Qed.
-
- Lemma map_1 : forall (elt elt':Set)(m: t elt)(x:key)(e:elt)(f:elt->elt'),
- MapsTo x e m -> MapsTo x (f e) (map f m).
- Proof.
- unfold map; intros.
- destruct (@mapi_1 _ _ m x e (fun _ => f)) as (e',(_,H0)); auto.
- Qed.
-
- Lemma map_2 : forall (elt elt':Set)(m: t elt)(x:key)(f:elt->elt'),
- In x (map f m) -> In x m.
- Proof.
- unfold map; intros.
- eapply mapi_2; eauto.
- Qed.
-
- Module L := FMapList.Raw E.
-
- (** Not exactly pretty nor perfect, but should suffice as a first naive implem.
- Anyway, map2 isn't in Ocaml...
- *)
-
- Definition anti_elements (A:Set)(l:list (key*A)) := L.fold (@add _) l (empty _).
-
- Definition map2 (A B C:Set)(f:option A->option B -> option C)(m:t A)(m':t B) : t C :=
- anti_elements (L.map2 f (elements m) (elements m')).
-
- Lemma add_spec : forall (A:Set)(m:t A) x y e,
- find x (add y e m) = if ME.eq_dec x y then Some e else find x m.
- Proof.
- intros.
- destruct (ME.eq_dec x y).
- apply find_1.
- eapply MapsTo_1 with y; eauto.
- red; auto.
- apply add_1; auto.
- red; auto.
- case_eq (find x m); intros.
- apply find_1.
- apply add_2; unfold E.eq in *; auto.
- case_eq (find x (add y e m)); auto; intros.
- rewrite <- H; symmetry.
- apply find_1; auto.
- apply (@add_3 _ m y x a e); unfold E.eq in *; auto.
- Qed.
-
- Lemma anti_elements_mapsto_aux : forall (A:Set)(l:list (key*A)) m k e,
- NoDupA (eq_key (A:=A)) l ->
- (forall x, L.PX.In x l -> In x m -> False) ->
- (MapsTo k e (L.fold (@add _) l m) <-> L.PX.MapsTo k e l \/ MapsTo k e m).
- Proof.
- induction l.
- simpl; auto.
- intuition.
- inversion H2.
- simpl; destruct a; intros.
- rewrite IHl; clear IHl.
- inversion H; auto.
- intros.
- inversion_clear H.
- assert (~E.eq x k).
- swap H3.
- destruct H1.
- apply InA_eqA with (x,x0); eauto.
- unfold eq_key, E.eq; eauto.
- unfold eq_key, E.eq; congruence.
- apply (H0 x).
- destruct H1; exists x0; auto.
- revert H2.
- unfold In.
- intros (e',He').
- exists e'; apply (@add_3 _ m k x e' a); unfold E.eq; auto.
- intuition.
- red in H2.
- rewrite add_spec in H2; auto.
- destruct (ME.eq_dec k0 k).
- inversion_clear H2; subst; auto.
- right; apply find_2; auto.
- inversion_clear H2; auto.
- compute in H1; destruct H1.
- subst; right; apply add_1; auto.
- red; auto.
- inversion_clear H.
- destruct (ME.eq_dec k0 k).
- unfold E.eq in *; subst.
- destruct (H0 k); eauto.
- red; eauto.
- right; apply add_2; unfold E.eq in *; auto.
- Qed.
-
- Lemma anti_elements_mapsto : forall (A:Set) l k e, NoDupA (eq_key (A:=A)) l ->
- (MapsTo k e (anti_elements l) <-> L.PX.MapsTo k e l).
- Proof.
- intros.
- unfold anti_elements.
- rewrite anti_elements_mapsto_aux; auto; unfold empty; auto.
- inversion 2.
- inversion H2.
- intuition.
- inversion H1.
- Qed.
-
- Lemma find_anti_elements : forall (A:Set)(l: list (key*A)) x, sort (@lt_key _) l ->
- find x (anti_elements l) = L.find x l.
- Proof.
- intros.
- case_eq (L.find x l); intros.
- apply find_1.
- rewrite anti_elements_mapsto.
- apply L.PX.Sort_NoDupA; auto.
- apply L.find_2; auto.
- case_eq (find x (anti_elements l)); auto; intros.
- rewrite <- H0; symmetry.
- apply L.find_1; auto.
- rewrite <- anti_elements_mapsto.
- apply L.PX.Sort_NoDupA; auto.
- apply find_2; auto.
- Qed.
-
- Lemma find_elements : forall (A:Set)(m: t A) x,
- L.find x (elements m) = find x m.
- Proof.
- intros.
- case_eq (find x m); intros.
- apply L.find_1.
- apply elements_3; auto.
- red; apply elements_1.
- apply find_2; auto.
- case_eq (L.find x (elements m)); auto; intros.
- rewrite <- H; symmetry.
- apply find_1; auto.
- apply elements_2.
- apply L.find_2; auto.
- Qed.
-
- Lemma elements_in : forall (A:Set)(s:t A) x, L.PX.In x (elements s) <-> In x s.
- Proof.
- intros.
- unfold L.PX.In, In.
- firstorder.
- exists x0.
- red; rewrite <- find_elements; auto.
- apply L.find_1; auto.
- apply elements_3.
- exists x0.
- apply L.find_2.
- rewrite find_elements; auto.
- Qed.
-
- Lemma map2_1 : forall (A B C:Set)(m: t A)(m': t B)(x:key)
- (f:option A->option B ->option C),
- In x m \/ In x m' -> find x (map2 f m m') = f (find x m) (find x m').
- Proof.
- unfold map2; intros.
- rewrite find_anti_elements; auto.
- rewrite <- find_elements; auto.
- rewrite <- find_elements; auto.
- apply L.map2_1; auto.
- apply elements_3; auto.
- apply elements_3; auto.
- do 2 rewrite elements_in; auto.
- apply L.map2_sorted; auto.
- apply elements_3; auto.
- apply elements_3; auto.
- Qed.
-
- Lemma map2_2 : forall (A B C: Set)(m: t A)(m': t B)(x:key)
- (f:option A->option B ->option C),
- In x (map2 f m m') -> In x m \/ In x m'.
- Proof.
- unfold map2; intros.
- do 2 rewrite <- elements_in.
- apply L.map2_2 with (f:=f); auto.
- apply elements_3; auto.
- apply elements_3; auto.
- destruct H.
- exists x0.
- rewrite <- anti_elements_mapsto; auto.
- apply L.PX.Sort_NoDupA; auto.
- apply L.map2_sorted; auto.
- apply elements_3; auto.
- apply elements_3; auto.
- Qed.
-
- (** same trick for [equal] *)
-
- Definition equal (A:Set)(cmp:A -> A -> bool)(m m' : t A) : bool :=
- L.equal cmp (elements m) (elements m').
-
- Lemma equal_1 :
- forall (A:Set)(m: t A)(m': t A)(cmp: A -> A -> bool),
- Equal cmp m m' -> equal cmp m m' = true.
- Proof.
- unfold equal, Equal.
- intros.
- apply L.equal_1.
- apply elements_3.
- apply elements_3.
- unfold L.Equal.
- destruct H.
- split; intros.
- do 2 rewrite elements_in; auto.
- apply (H0 k);
- red; rewrite <- find_elements; apply L.find_1; auto;
- apply elements_3.
- Qed.
-
- Lemma equal_2 :
- forall (A:Set)(m: t A)(m': t A)(cmp: A -> A -> bool),
- equal cmp m m' = true -> Equal cmp m m'.
- Proof.
- unfold equal, Equal.
- intros.
- destruct (L.equal_2 (elements_3 m) (elements_3 m') H); clear H.
- split.
- intros; do 2 rewrite <- elements_in; auto.
- intros; apply (H1 k);
- apply L.find_2; rewrite find_elements;auto.
- Qed.
-
-End MapIntMap.
-
diff --git a/theories/FSets/FMapInterface.v b/theories/FSets/FMapInterface.v
index dde74a0a..1e475887 100644
--- a/theories/FSets/FMapInterface.v
+++ b/theories/FSets/FMapInterface.v
@@ -6,42 +6,72 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: FMapInterface.v 8671 2006-03-29 08:31:28Z letouzey $ *)
+(* $Id: FMapInterface.v 10616 2008-03-04 17:33:35Z letouzey $ *)
(** * Finite map library *)
-(** This file proposes an interface for finite maps *)
+(** This file proposes interfaces for finite maps *)
-(* begin hide *)
+Require Export Bool DecidableType OrderedType.
Set Implicit Arguments.
Unset Strict Implicit.
-Require Import FSetInterface.
-(* end hide *)
-
-(** When compared with Ocaml Map, this signature has been split in two:
- - The first part [S] contains the usual operators (add, find, ...)
- It only requires a ordered key type, the data type can be arbitrary.
- The only function that asks more is [equal], whose first argument should
- be an equality on data.
- - Then, [Sord] extends [S] with a complete comparison function. For
- that, the data type should have a decidable total ordering.
+
+(** When compared with Ocaml Map, this signature has been split in
+ several parts :
+
+ - The first parts [WSfun] and [WS] propose signatures for weak
+ maps, which are maps with no ordering on the key type nor the
+ data type. [WSfun] and [WS] are almost identical, apart from the
+ fact that [WSfun] is expressed in a functorial way whereas [WS]
+ is self-contained. For obtaining an instance of such signatures,
+ a decidable equality on keys in enough (see for example
+ [FMapWeakList]). These signatures contain the usual operators
+ (add, find, ...). The only function that asks for more is
+ [equal], whose first argument should be a comparison on data.
+
+ - Then comes [Sfun] and [S], that extend [WSfun] and [WS] to the
+ case where the key type is ordered. The main novelty is that
+ [elements] is required to produce sorted lists.
+
+ - Finally, [Sord] extends [S] with a complete comparison function. For
+ that, the data type should have a decidable total ordering as well.
+
+ If unsure, what you're looking for is probably [S]: apart from [Sord],
+ all other signatures are subsets of [S].
+
+ Some additional differences with Ocaml:
+
+ - no [iter] function, useless since Coq is purely functional
+ - [option] types are used instead of [Not_found] exceptions
+ - more functions are provided: [elements] and [cardinal] and [map2]
+
*)
-Module Type S.
+Definition Cmp (elt:Type)(cmp:elt->elt->bool) e1 e2 := cmp e1 e2 = true.
- Declare Module E : OrderedType.
+(** ** Weak signature for maps
+
+ No requirements for an ordering on keys nor elements, only decidability
+ of equality on keys. First, a functorial signature: *)
+
+Module Type WSfun (E : EqualityType).
+
+ (** The module E of base objects is meant to be a [DecidableType]
+ (and used to be so). But requiring only an [EqualityType] here
+ allows subtyping between weak and ordered maps. *)
Definition key := E.t.
- Parameter t : Set -> Set. (** the abstract type of maps *)
+ Parameter t : Type -> Type.
+ (** the abstract type of maps *)
Section Types.
- Variable elt:Set.
+ Variable elt:Type.
Parameter empty : t elt.
- (** The empty map. *)
+ (** The empty map. *)
Parameter is_empty : t elt -> bool.
(** Test whether a map is empty or not. *)
@@ -53,8 +83,7 @@ Module Type S.
Parameter find : key -> t elt -> option elt.
(** [find x m] returns the current binding of [x] in [m],
- or raises [Not_found] if no such binding exists.
- NB: in Coq, the exception mechanism becomes a option type. *)
+ or [None] if no such binding exists. *)
Parameter remove : key -> t elt -> t elt.
(** [remove x m] returns a map containing the same bindings as [m],
@@ -64,45 +93,36 @@ Module Type S.
(** [mem x m] returns [true] if [m] contains a binding for [x],
and [false] otherwise. *)
- (** Coq comment: [iter] is useless in a purely functional world *)
- (** val iter : (key -> 'a -> unit) -> 'a t -> unit *)
- (** iter f m applies f to all bindings in map m. f receives the key as
- first argument, and the associated value as second argument.
- The bindings are passed to f in increasing order with respect to the
- ordering over the type of the keys. Only current bindings are
- presented to f: bindings hidden by more recent bindings are not
- passed to f. *)
-
- Variable elt' : Set.
- Variable elt'': Set.
+ Variable elt' elt'' : Type.
Parameter map : (elt -> elt') -> t elt -> t elt'.
(** [map f m] returns a map with same domain as [m], where the associated
value a of all bindings of [m] has been replaced by the result of the
- application of [f] to [a]. The bindings are passed to [f] in
- increasing order with respect to the ordering over the type of the
- keys. *)
+ application of [f] to [a]. Since Coq is purely functional, the order
+ in which the bindings are passed to [f] is irrelevant. *)
Parameter mapi : (key -> elt -> elt') -> t elt -> t elt'.
- (** Same as [S.map], but the function receives as arguments both the
+ (** Same as [map], but the function receives as arguments both the
key and the associated value for each binding of the map. *)
- Parameter map2 : (option elt -> option elt' -> option elt'') -> t elt -> t elt' -> t elt''.
- (** Not present in Ocaml.
- [map f m m'] creates a new map whose bindings belong to the ones of either
- [m] or [m']. The presence and value for a key [k] is determined by [f e e']
- where [e] and [e'] are the (optional) bindings of [k] in [m] and [m']. *)
+ Parameter map2 :
+ (option elt -> option elt' -> option elt'') -> t elt -> t elt' -> t elt''.
+ (** [map2 f m m'] creates a new map whose bindings belong to the ones
+ of either [m] or [m']. The presence and value for a key [k] is
+ determined by [f e e'] where [e] and [e'] are the (optional) bindings
+ of [k] in [m] and [m']. *)
Parameter elements : t elt -> list (key*elt).
- (** Not present in Ocaml.
- [elements m] returns an assoc list corresponding to the bindings of [m].
- Elements of this list are sorted with respect to their first components.
- Useful to specify [fold] ... *)
+ (** [elements m] returns an assoc list corresponding to the bindings
+ of [m], in any order. *)
- Parameter fold : forall A: Set, (key -> elt -> A -> A) -> t elt -> A -> A.
+ Parameter cardinal : t elt -> nat.
+ (** [cardinal m] returns the number of bindings in [m]. *)
+
+ Parameter fold : forall A: Type, (key -> elt -> A -> A) -> t elt -> A -> A.
(** [fold f m a] computes [(f kN dN ... (f k1 d1 a)...)],
where [k1] ... [kN] are the keys of all bindings in [m]
- (in increasing order), and [d1] ... [dN] are the associated data. *)
+ (in any order), and [d1] ... [dN] are the associated data. *)
Parameter equal : (elt -> elt -> bool) -> t elt -> t elt -> bool.
(** [equal cmp m1 m2] tests whether the maps [m1] and [m2] are equal,
@@ -127,8 +147,6 @@ Module Type S.
Definition eq_key_elt (p p':key*elt) :=
E.eq (fst p) (fst p') /\ (snd p) = (snd p').
- Definition lt_key (p p':key*elt) := E.lt (fst p) (fst p').
-
(** Specification of [MapsTo] *)
Parameter MapsTo_1 : E.eq x y -> MapsTo x e m -> MapsTo y e m.
@@ -162,61 +180,123 @@ Module Type S.
MapsTo x e m -> InA eq_key_elt (x,e) (elements m).
Parameter elements_2 :
InA eq_key_elt (x,e) (elements m) -> MapsTo x e m.
- Parameter elements_3 : sort lt_key (elements m).
+ (** When compared with ordered maps, here comes the only
+ property that is really weaker: *)
+ Parameter elements_3w : NoDupA eq_key (elements m).
+
+ (** Specification of [cardinal] *)
+ Parameter cardinal_1 : cardinal m = length (elements m).
(** Specification of [fold] *)
Parameter fold_1 :
- forall (A : Set) (i : A) (f : key -> elt -> A -> A),
+ forall (A : Type) (i : A) (f : key -> elt -> A -> A),
fold f m i = fold_left (fun a p => f (fst p) (snd p) a) (elements m) i.
+
+ (** Equality of maps *)
- Definition Equal cmp m m' :=
+ (** Caveat: there are at least three distinct equality predicates on maps.
+ - The simpliest (and maybe most natural) way is to consider keys up to
+ their equivalence [E.eq], but elements up to Leibniz equality, in
+ the spirit of [eq_key_elt] above. This leads to predicate [Equal].
+ - Unfortunately, this [Equal] predicate can't be used to describe
+ the [equal] function, since this function (for compatibility with
+ ocaml) expects a boolean comparison [cmp] that may identify more
+ elements than Leibniz. So logical specification of [equal] is done
+ via another predicate [Equivb]
+ - This predicate [Equivb] is quite ad-hoc with its boolean [cmp],
+ it can be generalized in a [Equiv] expecting a more general
+ (possibly non-decidable) equality predicate on elements *)
+
+ Definition Equal m m' := forall y, find y m = find y m'.
+ Definition Equiv (eq_elt:elt->elt->Prop) m m' :=
(forall k, In k m <-> In k m') /\
- (forall k e e', MapsTo k e m -> MapsTo k e' m' -> cmp e e' = true).
+ (forall k e e', MapsTo k e m -> MapsTo k e' m' -> eq_elt e e').
+ Definition Equivb (cmp: elt->elt->bool) := Equiv (Cmp cmp).
- Variable cmp : elt -> elt -> bool.
+ (** Specification of [equal] *)
- (** Specification of [equal] *)
- Parameter equal_1 : Equal cmp m m' -> equal cmp m m' = true.
- Parameter equal_2 : equal cmp m m' = true -> Equal cmp m m'.
+ Variable cmp : elt -> elt -> bool.
- End Spec.
- End Types.
+ Parameter equal_1 : Equivb cmp m m' -> equal cmp m m' = true.
+ Parameter equal_2 : equal cmp m m' = true -> Equivb cmp m m'.
+
+ End Spec.
+ End Types.
(** Specification of [map] *)
- Parameter map_1 : forall (elt elt':Set)(m: t elt)(x:key)(e:elt)(f:elt->elt'),
+ Parameter map_1 : forall (elt elt':Type)(m: t elt)(x:key)(e:elt)(f:elt->elt'),
MapsTo x e m -> MapsTo x (f e) (map f m).
- Parameter map_2 : forall (elt elt':Set)(m: t elt)(x:key)(f:elt->elt'),
+ Parameter map_2 : forall (elt elt':Type)(m: t elt)(x:key)(f:elt->elt'),
In x (map f m) -> In x m.
(** Specification of [mapi] *)
- Parameter mapi_1 : forall (elt elt':Set)(m: t elt)(x:key)(e:elt)
+ Parameter mapi_1 : forall (elt elt':Type)(m: t elt)(x:key)(e:elt)
(f:key->elt->elt'), MapsTo x e m ->
exists y, E.eq y x /\ MapsTo x (f y e) (mapi f m).
- Parameter mapi_2 : forall (elt elt':Set)(m: t elt)(x:key)
+ Parameter mapi_2 : forall (elt elt':Type)(m: t elt)(x:key)
(f:key->elt->elt'), In x (mapi f m) -> In x m.
(** Specification of [map2] *)
- Parameter map2_1 : forall (elt elt' elt'':Set)(m: t elt)(m': t elt')
+ Parameter map2_1 : forall (elt elt' elt'':Type)(m: t elt)(m': t elt')
(x:key)(f:option elt->option elt'->option elt''),
In x m \/ In x m' ->
find x (map2 f m m') = f (find x m) (find x m').
- Parameter map2_2 : forall (elt elt' elt'':Set)(m: t elt)(m': t elt')
+ Parameter map2_2 : forall (elt elt' elt'':Type)(m: t elt)(m': t elt')
(x:key)(f:option elt->option elt'->option elt''),
In x (map2 f m m') -> In x m \/ In x m'.
- (* begin hide *)
- Hint Immediate MapsTo_1 mem_2 is_empty_2.
-
- Hint Resolve mem_1 is_empty_1 is_empty_2 add_1 add_2 add_3 remove_1
- remove_2 remove_3 find_1 find_2 fold_1 map_1 map_2 mapi_1 mapi_2.
- (* end hide *)
+ Hint Immediate MapsTo_1 mem_2 is_empty_2
+ map_2 mapi_2 add_3 remove_3 find_2
+ : map.
+ Hint Resolve mem_1 is_empty_1 is_empty_2 add_1 add_2 remove_1
+ remove_2 find_1 fold_1 map_1 mapi_1 mapi_2
+ : map.
+End WSfun.
+
+
+(** ** Static signature for Weak Maps
+
+ Similar to [WSfun] but expressed in a self-contained way. *)
+
+Module Type WS.
+ Declare Module E : EqualityType.
+ Include Type WSfun E.
+End WS.
+
+
+
+(** ** Maps on ordered keys, functorial signature *)
+
+Module Type Sfun (E : OrderedType).
+ Include Type WSfun E.
+ Section elt.
+ Variable elt:Type.
+ Definition lt_key (p p':key*elt) := E.lt (fst p) (fst p').
+ (* Additional specification of [elements] *)
+ Parameter elements_3 : forall m, sort lt_key (elements m).
+ (** 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. *)
+ End elt.
+End Sfun.
+
+
+
+(** ** Maps on ordered keys, self-contained signature *)
+
+Module Type S.
+ Declare Module E : OrderedType.
+ Include Type Sfun E.
End S.
+
+(** ** Maps with ordering both on keys and datas *)
+
Module Type Sord.
-
+
Declare Module Data : OrderedType.
Declare Module MapS : S.
Import MapS.
@@ -234,12 +314,11 @@ Module Type Sord.
Definition cmp e e' := match Data.compare e e' with EQ _ => true | _ => false end.
- Parameter eq_1 : forall m m', Equal cmp m m' -> eq m m'.
- Parameter eq_2 : forall m m', eq m m' -> Equal cmp m m'.
+ Parameter eq_1 : forall m m', Equivb cmp m m' -> eq m m'.
+ Parameter eq_2 : forall m m', eq m m' -> Equivb cmp m m'.
Parameter compare : forall m1 m2, Compare lt eq m1 m2.
- (** Total ordering between maps. The first argument (in Coq: Data.compare)
- is a total ordering used to compare data associated with equal keys
- in the two maps. *)
+ (** Total ordering between maps. [Data.compare] is a total ordering
+ used to compare data associated with equal keys in the two maps. *)
-End Sord. \ No newline at end of file
+End Sord.
diff --git a/theories/FSets/FMapList.v b/theories/FSets/FMapList.v
index 067f5a3e..23bf8196 100644
--- a/theories/FSets/FMapList.v
+++ b/theories/FSets/FMapList.v
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: FMapList.v 9035 2006-07-09 15:42:09Z herbelin $ *)
+(* $Id: FMapList.v 10616 2008-03-04 17:33:35Z letouzey $ *)
(** * Finite map library *)
@@ -14,7 +14,6 @@
[FMapInterface.S] using lists of pairs ordered (increasing) with respect to
left projection. *)
-Require Import FSetInterface.
Require Import FMapInterface.
Set Implicit Arguments.
@@ -22,26 +21,14 @@ Unset Strict Implicit.
Module Raw (X:OrderedType).
-Module E := X.
-Module MX := OrderedTypeFacts X.
-Module PX := KeyOrderedType X.
-Import MX.
-Import PX.
+Module Import MX := OrderedTypeFacts X.
+Module Import PX := KeyOrderedType X.
Definition key := X.t.
-Definition t (elt:Set) := list (X.t * elt).
+Definition t (elt:Type) := list (X.t * elt).
Section Elt.
-Variable elt : Set.
-
-(* Now in KeyOrderedType:
-Definition eqk (p p':key*elt) := X.eq (fst p) (fst p').
-Definition eqke (p p':key*elt) :=
- X.eq (fst p) (fst p') /\ (snd p) = (snd p').
-Definition ltk (p p':key*elt) := X.lt (fst p) (fst p').
-Definition MapsTo (k:key)(e:elt):= InA eqke (k,e).
-Definition In k m := exists e:elt, MapsTo k e m.
-*)
+Variable elt : Type.
Notation eqk := (eqk (elt:=elt)).
Notation eqke := (eqke (elt:=elt)).
@@ -347,15 +334,22 @@ Proof.
auto.
Qed.
+Lemma elements_3w : forall m (Hm:Sort m), NoDupA eqk (elements m).
+Proof.
+ intros.
+ apply Sort_NoDupA.
+ apply elements_3; auto.
+Qed.
+
(** * [fold] *)
-Function fold (A:Set)(f:key->elt->A->A)(m:t elt) (acc:A) {struct m} : A :=
+Function fold (A:Type)(f:key->elt->A->A)(m:t elt) (acc:A) {struct m} : A :=
match m with
| nil => acc
| (k,e)::m' => fold f m' (f k e acc)
end.
-Lemma fold_1 : forall m (A:Set)(i:A)(f:key->elt->A->A),
+Lemma fold_1 : forall m (A:Type)(i:A)(f:key->elt->A->A),
fold f m i = fold_left (fun a p => f (fst p) (snd p) a) (elements m) i.
Proof.
intros; functional induction (fold f m i); auto.
@@ -374,29 +368,24 @@ Function equal (cmp:elt->elt->bool)(m m' : t elt) { struct m } : bool :=
| _, _ => false
end.
-Definition Equal cmp m m' :=
+Definition Equivb cmp m m' :=
(forall k, In k m <-> In k m') /\
(forall k e e', MapsTo k e m -> MapsTo k e' m' -> cmp e e' = true).
Lemma equal_1 : forall m (Hm:Sort m) m' (Hm': Sort m') cmp,
- Equal cmp m m' -> equal cmp m m' = true.
+ Equivb cmp m m' -> equal cmp m m' = true.
Proof.
intros m Hm m' Hm' cmp; generalize Hm Hm'; clear Hm Hm'.
- functional induction (equal cmp m m'); simpl; subst;auto; unfold Equal;
- intuition; subst; match goal with
- | [H: X.compare _ _ = _ |- _ ] => clear H
- | _ => idtac
- end.
-
-
-
+ functional induction (equal cmp m m'); simpl; subst;auto; unfold Equivb;
+ intuition; subst.
+ match goal with H: X.compare _ _ = _ |- _ => clear H end.
assert (cmp_e_e':cmp e e' = true).
apply H1 with x; auto.
rewrite cmp_e_e'; simpl.
apply IHb; auto.
inversion_clear Hm; auto.
inversion_clear Hm'; auto.
- unfold Equal; intuition.
+ unfold Equivb; intuition.
destruct (H0 k).
assert (In k ((x,e) ::l)).
destruct H as (e'', hyp); exists e''; auto.
@@ -459,14 +448,12 @@ Qed.
Lemma equal_2 : forall m (Hm:Sort m) m' (Hm:Sort m') cmp,
- equal cmp m m' = true -> Equal cmp m m'.
+ equal cmp m m' = true -> Equivb cmp m m'.
Proof.
intros m Hm m' Hm' cmp; generalize Hm Hm'; clear Hm Hm'.
- functional induction (equal cmp m m'); simpl; subst;auto; unfold Equal;
- intuition; try discriminate; subst; match goal with
- | [H: X.compare _ _ = _ |- _ ] => clear H
- | _ => idtac
- end.
+ functional induction (equal cmp m m'); simpl; subst;auto; unfold Equivb;
+ intuition; try discriminate; subst;
+ try match goal with H: X.compare _ _ = _ |- _ => clear H end.
inversion H0.
@@ -502,13 +489,13 @@ Proof.
elim (Sort_Inf_NotIn H2 H3).
exists e0; apply MapsTo_eq with k; auto; order.
apply H8 with k; auto.
-Qed.
+Qed.
-(** This lemma isn't part of the spec of [Equal], but is used in [FMapAVL] *)
+(** This lemma isn't part of the spec of [Equivb], but is used in [FMapAVL] *)
Lemma equal_cons : forall cmp l1 l2 x y, Sort (x::l1) -> Sort (y::l2) ->
eqk x y -> cmp (snd x) (snd y) = true ->
- (Equal cmp l1 l2 <-> Equal cmp (x :: l1) (y :: l2)).
+ (Equivb cmp l1 l2 <-> Equivb cmp (x :: l1) (y :: l2)).
Proof.
intros.
inversion H; subst.
@@ -527,7 +514,7 @@ Proof.
rewrite H2; simpl; auto.
Qed.
-Variable elt':Set.
+Variable elt':Type.
(** * [map] and [mapi] *)
@@ -548,7 +535,7 @@ Section Elt2.
(* A new section is necessary for previous definitions to work
with different [elt], especially [MapsTo]... *)
-Variable elt elt' : Set.
+Variable elt elt' : Type.
(** Specification of [map] *)
@@ -684,10 +671,10 @@ Section Elt3.
(** * [map2] *)
-Variable elt elt' elt'' : Set.
+Variable elt elt' elt'' : Type.
Variable f : option elt -> option elt' -> option elt''.
-Definition option_cons (A:Set)(k:key)(o:option A)(l:list (key*A)) :=
+Definition option_cons (A:Type)(k:key)(o:option A)(l:list (key*A)) :=
match o with
| Some e => (k,e)::l
| None => l
@@ -739,7 +726,7 @@ Fixpoint combine (m : t elt) : t elt' -> t oee' :=
end
end.
-Definition fold_right_pair (A B C:Set)(f: A->B->C->C)(l:list (A*B))(i:C) :=
+Definition fold_right_pair (A B C:Type)(f: A->B->C->C)(l:list (A*B))(i:C) :=
List.fold_right (fun p => f (fst p) (snd p)) i l.
Definition map2_alt m m' :=
@@ -1038,12 +1025,12 @@ Module E := X.
Definition key := E.t.
-Record slist (elt:Set) : Set :=
+Record slist (elt:Type) :=
{this :> Raw.t elt; sorted : sort (@Raw.PX.ltk elt) this}.
-Definition t (elt:Set) : Set := slist elt.
+Definition t (elt:Type) : Type := slist elt.
Section Elt.
- Variable elt elt' elt'':Set.
+ Variable elt elt' elt'':Type.
Implicit Types m : t elt.
Implicit Types x y : key.
@@ -1060,13 +1047,19 @@ Section Elt.
Definition map2 f m (m':t elt') : t elt'' :=
Build_slist (Raw.map2_sorted f m.(sorted) m'.(sorted)).
Definition elements m : list (key*elt) := @Raw.elements elt m.(this).
- Definition fold (A:Set)(f:key->elt->A->A) m (i:A) : A := @Raw.fold elt A f m.(this) i.
+ Definition cardinal m := length m.(this).
+ Definition fold (A:Type)(f:key->elt->A->A) m (i:A) : A := @Raw.fold elt A f m.(this) i.
Definition equal cmp m m' : bool := @Raw.equal elt cmp m.(this) m'.(this).
Definition MapsTo x e m : Prop := Raw.PX.MapsTo x e m.(this).
Definition In x m : Prop := Raw.PX.In x m.(this).
Definition Empty m : Prop := Raw.Empty m.(this).
- Definition Equal cmp m m' : Prop := @Raw.Equal elt cmp m.(this) m'.(this).
+
+ Definition Equal m m' := forall y, find y m = find y m'.
+ Definition Equiv (eq_elt:elt->elt->Prop) m m' :=
+ (forall k, In k m <-> In k m') /\
+ (forall k e e', MapsTo k e m -> MapsTo k e' m' -> eq_elt e e').
+ Definition Equivb cmp m m' : Prop := @Raw.Equivb elt cmp m.(this) m'.(this).
Definition eq_key : (key*elt) -> (key*elt) -> Prop := @Raw.PX.eqk elt.
Definition eq_key_elt : (key*elt) -> (key*elt) -> Prop:= @Raw.PX.eqke elt.
@@ -1113,34 +1106,39 @@ Section Elt.
Proof. intros m; exact (@Raw.elements_2 elt m.(this)). Qed.
Lemma elements_3 : forall m, sort lt_key (elements m).
Proof. intros m; exact (@Raw.elements_3 elt m.(this) m.(sorted)). Qed.
+ Lemma elements_3w : forall m, NoDupA eq_key (elements m).
+ Proof. intros m; exact (@Raw.elements_3w elt m.(this) m.(sorted)). Qed.
+
+ Lemma cardinal_1 : forall m, cardinal m = length (elements m).
+ Proof. intros; reflexivity. Qed.
- Lemma fold_1 : forall m (A : Set) (i : A) (f : key -> elt -> A -> A),
+ Lemma fold_1 : forall m (A : Type) (i : A) (f : key -> elt -> A -> A),
fold f m i = fold_left (fun a p => f (fst p) (snd p) a) (elements m) i.
Proof. intros m; exact (@Raw.fold_1 elt m.(this)). Qed.
- Lemma equal_1 : forall m m' cmp, Equal cmp m m' -> equal cmp m m' = true.
+ Lemma equal_1 : forall m m' cmp, Equivb cmp m m' -> equal cmp m m' = true.
Proof. intros m m'; exact (@Raw.equal_1 elt m.(this) m.(sorted) m'.(this) m'.(sorted)). Qed.
- Lemma equal_2 : forall m m' cmp, equal cmp m m' = true -> Equal cmp m m'.
+ Lemma equal_2 : forall m m' cmp, equal cmp m m' = true -> Equivb cmp m m'.
Proof. intros m m'; exact (@Raw.equal_2 elt m.(this) m.(sorted) m'.(this) m'.(sorted)). Qed.
End Elt.
- Lemma map_1 : forall (elt elt':Set)(m: t elt)(x:key)(e:elt)(f:elt->elt'),
+ Lemma map_1 : forall (elt elt':Type)(m: t elt)(x:key)(e:elt)(f:elt->elt'),
MapsTo x e m -> MapsTo x (f e) (map f m).
Proof. intros elt elt' m; exact (@Raw.map_1 elt elt' m.(this)). Qed.
- Lemma map_2 : forall (elt elt':Set)(m: t elt)(x:key)(f:elt->elt'),
+ Lemma map_2 : forall (elt elt':Type)(m: t elt)(x:key)(f:elt->elt'),
In x (map f m) -> In x m.
Proof. intros elt elt' m; exact (@Raw.map_2 elt elt' m.(this)). Qed.
- Lemma mapi_1 : forall (elt elt':Set)(m: t elt)(x:key)(e:elt)
+ Lemma mapi_1 : forall (elt elt':Type)(m: t elt)(x:key)(e:elt)
(f:key->elt->elt'), MapsTo x e m ->
exists y, E.eq y x /\ MapsTo x (f y e) (mapi f m).
Proof. intros elt elt' m; exact (@Raw.mapi_1 elt elt' m.(this)). Qed.
- Lemma mapi_2 : forall (elt elt':Set)(m: t elt)(x:key)
+ Lemma mapi_2 : forall (elt elt':Type)(m: t elt)(x:key)
(f:key->elt->elt'), In x (mapi f m) -> In x m.
Proof. intros elt elt' m; exact (@Raw.mapi_2 elt elt' m.(this)). Qed.
- Lemma map2_1 : forall (elt elt' elt'':Set)(m: t elt)(m': t elt')
+ Lemma map2_1 : forall (elt elt' elt'':Type)(m: t elt)(m': t elt')
(x:key)(f:option elt->option elt'->option elt''),
In x m \/ In x m' ->
find x (map2 f m m') = f (find x m) (find x m').
@@ -1148,7 +1146,7 @@ Section Elt.
intros elt elt' elt'' m m' x f;
exact (@Raw.map2_1 elt elt' elt'' f m.(this) m.(sorted) m'.(this) m'.(sorted) x).
Qed.
- Lemma map2_2 : forall (elt elt' elt'':Set)(m: t elt)(m': t elt')
+ Lemma map2_2 : forall (elt elt' elt'':Type)(m: t elt)(m': t elt')
(x:key)(f:option elt->option elt'->option elt''),
In x (map2 f m m') -> In x m \/ In x m'.
Proof.
@@ -1229,7 +1227,7 @@ Proof.
unfold equal, eq in H6; simpl in H6; auto.
Qed.
-Lemma eq_1 : forall m m', Equal cmp m m' -> eq m m'.
+Lemma eq_1 : forall m m', Equivb cmp m m' -> eq m m'.
Proof.
intros.
generalize (@equal_1 D.t m m' cmp).
@@ -1237,7 +1235,7 @@ Proof.
intuition.
Qed.
-Lemma eq_2 : forall m m', eq m m' -> Equal cmp m m'.
+Lemma eq_2 : forall m m', eq m m' -> Equivb cmp m m'.
Proof.
intros.
generalize (@equal_2 D.t m m' cmp).
diff --git a/theories/FSets/FMapPositive.v b/theories/FSets/FMapPositive.v
index 44724767..9bc2a599 100644
--- a/theories/FSets/FMapPositive.v
+++ b/theories/FSets/FMapPositive.v
@@ -11,11 +11,12 @@
* Institution: LRI, CNRS UMR 8623 - Université Paris Sud
* 91405 Orsay, France *)
-(* $Id: FMapPositive.v 9862 2007-05-25 16:57:06Z letouzey $ *)
+(* $Id: FMapPositive.v 10739 2008-04-01 14:45:20Z herbelin $ *)
Require Import Bool.
Require Import ZArith.
Require Import OrderedType.
+Require Import OrderedTypeEx.
Require Import FMapInterface.
Set Implicit Arguments.
@@ -36,9 +37,12 @@ Open Local Scope positive_scope.
usual order (see [OrderedTypeEx]), we use here a lexicographic order
over bits, which is more natural here (lower bits are considered first). *)
-Module PositiveOrderedTypeBits <: OrderedType.
+Module PositiveOrderedTypeBits <: UsualOrderedType.
Definition t:=positive.
Definition eq:=@eq positive.
+ Definition eq_refl := @refl_equal t.
+ Definition eq_sym := @sym_eq t.
+ Definition eq_trans := @trans_eq t.
Fixpoint bits_lt (p q:positive) { struct p } : Prop :=
match p, q with
@@ -52,15 +56,6 @@ Module PositiveOrderedTypeBits <: OrderedType.
Definition lt:=bits_lt.
- Lemma eq_refl : forall x : t, eq x x.
- Proof. red; auto. Qed.
-
- Lemma eq_sym : forall x y : t, eq x y -> eq y x.
- Proof. red; auto. Qed.
-
- Lemma eq_trans : forall x y z : t, eq x y -> eq y z -> eq x z.
- Proof. red; intros; transitivity y; 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.
@@ -171,17 +166,18 @@ Qed.
Module PositiveMap <: S with Module E:=PositiveOrderedTypeBits.
Module E:=PositiveOrderedTypeBits.
+ Module ME:=KeyOrderedType E.
Definition key := positive.
- Inductive tree (A : Set) : Set :=
+ Inductive tree (A : Type) :=
| Leaf : tree A
| Node : tree A -> option A -> tree A -> tree A.
Definition t := tree.
Section A.
- Variable A:Set.
+ Variable A:Type.
Implicit Arguments Leaf [A].
@@ -280,6 +276,15 @@ Module PositiveMap <: S with Module E:=PositiveOrderedTypeBits.
Definition elements (m : t A) := xelements m xH.
+ (** [cardinal] *)
+
+ Fixpoint cardinal (m : t A) : nat :=
+ match m with
+ | Leaf => 0%nat
+ | Node l None r => (cardinal l + cardinal r)%nat
+ | Node l (Some _) r => S (cardinal l + cardinal r)
+ end.
+
Section CompcertSpec.
Theorem gempty:
@@ -560,6 +565,16 @@ Module PositiveMap <: S with Module E:=PositiveOrderedTypeBits.
exact (xelements_complete i xH m v H).
Qed.
+ Lemma cardinal_1 :
+ forall (m: t A), cardinal m = length (elements m).
+ Proof.
+ unfold elements.
+ intros m; set (p:=1); clearbody p; revert m p.
+ induction m; simpl; auto; intros.
+ rewrite (IHm1 (append p 2)), (IHm2 (append p 3)); auto.
+ destruct o; rewrite app_length; simpl; omega.
+ Qed.
+
End CompcertSpec.
Definition MapsTo (i:positive)(v:A)(m:t A) := find i m = Some v.
@@ -793,11 +808,17 @@ Module PositiveMap <: S with Module E:=PositiveOrderedTypeBits.
apply xelements_sort; auto.
Qed.
+ Lemma elements_3w : NoDupA eq_key (elements m).
+ Proof.
+ change eq_key with (@ME.eqk A).
+ apply ME.Sort_NoDupA; apply elements_3; auto.
+ Qed.
+
End FMapSpec.
(** [map] and [mapi] *)
- Variable B : Set.
+ Variable B : Type.
Fixpoint xmapi (f : positive -> A -> B) (m : t A) (i : positive)
{struct m} : t B :=
@@ -815,7 +836,7 @@ Module PositiveMap <: S with Module E:=PositiveOrderedTypeBits.
End A.
Lemma xgmapi:
- forall (A B: Set) (f: positive -> A -> B) (i j : positive) (m: t A),
+ forall (A B: Type) (f: positive -> A -> B) (i j : positive) (m: t A),
find i (xmapi f m j) = option_map (f (append j i)) (find i m).
Proof.
induction i; intros; destruct m; simpl; auto.
@@ -825,7 +846,7 @@ Module PositiveMap <: S with Module E:=PositiveOrderedTypeBits.
Qed.
Theorem gmapi:
- forall (A B: Set) (f: positive -> A -> B) (i: positive) (m: t A),
+ forall (A B: Type) (f: positive -> A -> B) (i: positive) (m: t A),
find i (mapi f m) = option_map (f i) (find i m).
Proof.
intros.
@@ -836,7 +857,7 @@ Module PositiveMap <: S with Module E:=PositiveOrderedTypeBits.
Qed.
Lemma mapi_1 :
- forall (elt elt':Set)(m: t elt)(x:key)(e:elt)(f:key->elt->elt'),
+ forall (elt elt':Type)(m: t elt)(x:key)(e:elt)(f:key->elt->elt'),
MapsTo x e m ->
exists y, E.eq y x /\ MapsTo x (f y e) (mapi f m).
Proof.
@@ -851,7 +872,7 @@ Module PositiveMap <: S with Module E:=PositiveOrderedTypeBits.
Qed.
Lemma mapi_2 :
- forall (elt elt':Set)(m: t elt)(x:key)(f:key->elt->elt'),
+ forall (elt elt':Type)(m: t elt)(x:key)(f:key->elt->elt'),
In x (mapi f m) -> In x m.
Proof.
intros.
@@ -864,21 +885,21 @@ Module PositiveMap <: S with Module E:=PositiveOrderedTypeBits.
simpl in *; discriminate.
Qed.
- Lemma map_1 : forall (elt elt':Set)(m: t elt)(x:key)(e:elt)(f:elt->elt'),
+ Lemma map_1 : forall (elt elt':Type)(m: t elt)(x:key)(e:elt)(f:elt->elt'),
MapsTo x e m -> MapsTo x (f e) (map f m).
Proof.
intros; unfold map.
destruct (mapi_1 (fun _ => f) H); intuition.
Qed.
- Lemma map_2 : forall (elt elt':Set)(m: t elt)(x:key)(f:elt->elt'),
+ Lemma map_2 : forall (elt elt':Type)(m: t elt)(x:key)(f:elt->elt'),
In x (map f m) -> In x m.
Proof.
intros; unfold map in *; eapply mapi_2; eauto.
Qed.
Section map2.
- Variable A B C : Set.
+ Variable A B C : Type.
Variable f : option A -> option B -> option C.
Implicit Arguments Leaf [A].
@@ -927,10 +948,10 @@ Module PositiveMap <: S with Module E:=PositiveOrderedTypeBits.
End map2.
- Definition map2 (elt elt' elt'':Set)(f:option elt->option elt'->option elt'') :=
+ Definition map2 (elt elt' elt'':Type)(f:option elt->option elt'->option elt'') :=
_map2 (fun o1 o2 => match o1,o2 with None,None => None | _, _ => f o1 o2 end).
- Lemma map2_1 : forall (elt elt' elt'':Set)(m: t elt)(m': t elt')
+ Lemma map2_1 : forall (elt elt' elt'':Type)(m: t elt)(m': t elt')
(x:key)(f:option elt->option elt'->option elt''),
In x m \/ In x m' ->
find x (map2 f m m') = f (find x m) (find x m').
@@ -946,7 +967,7 @@ Module PositiveMap <: S with Module E:=PositiveOrderedTypeBits.
destruct H; intuition; try discriminate.
Qed.
- Lemma map2_2 : forall (elt elt' elt'':Set)(m: t elt)(m': t elt')
+ Lemma map2_2 : forall (elt elt' elt'':Type)(m: t elt)(m': t elt')
(x:key)(f:option elt->option elt'->option elt''),
In x (map2 f m m') -> In x m \/ In x m'.
Proof.
@@ -962,17 +983,17 @@ Module PositiveMap <: S with Module E:=PositiveOrderedTypeBits.
Qed.
- Definition fold (A B : Set) (f: positive -> A -> B -> B) (tr: t A) (v: B) :=
+ Definition fold (A : Type)(B : Type) (f: positive -> A -> B -> B) (tr: t A) (v: B) :=
List.fold_left (fun a p => f (fst p) (snd p) a) (elements tr) v.
Lemma fold_1 :
- forall (A:Set)(m:t A)(B:Set)(i : B) (f : key -> A -> B -> B),
+ forall (A:Type)(m:t A)(B:Type)(i : B) (f : key -> A -> B -> B),
fold f m i = fold_left (fun a p => f (fst p) (snd p) a) (elements m) i.
Proof.
intros; unfold fold; auto.
Qed.
- Fixpoint equal (A:Set)(cmp : A -> A -> bool)(m1 m2 : t A) {struct m1} : bool :=
+ Fixpoint equal (A:Type)(cmp : A -> A -> bool)(m1 m2 : t A) {struct m1} : bool :=
match m1, m2 with
| Leaf, _ => is_empty m2
| _, Leaf => is_empty m1
@@ -985,12 +1006,15 @@ Module PositiveMap <: S with Module E:=PositiveOrderedTypeBits.
&& equal cmp l1 l2 && equal cmp r1 r2
end.
- Definition Equal (A:Set)(cmp:A->A->bool)(m m':t A) :=
+ Definition Equal (A:Type)(m m':t A) :=
+ forall y, find y m = find y m'.
+ Definition Equiv (A:Type)(eq_elt:A->A->Prop) m m' :=
(forall k, In k m <-> In k m') /\
- (forall k e e', MapsTo k e m -> MapsTo k e' m' -> cmp e e' = true).
+ (forall k e e', MapsTo k e m -> MapsTo k e' m' -> eq_elt e e').
+ Definition Equivb (A:Type)(cmp: A->A->bool) := Equiv (Cmp cmp).
- Lemma equal_1 : forall (A:Set)(m m':t A)(cmp:A->A->bool),
- Equal cmp m m' -> equal cmp m m' = true.
+ Lemma equal_1 : forall (A:Type)(m m':t A)(cmp:A->A->bool),
+ Equivb cmp m m' -> equal cmp m m' = true.
Proof.
induction m.
(* m = Leaf *)
@@ -1024,11 +1048,11 @@ Module PositiveMap <: S with Module E:=PositiveOrderedTypeBits.
destruct H2; red in H2; simpl in H2; discriminate.
(* m' = Node *)
destruct 1.
- assert (Equal cmp m1 m'1).
+ assert (Equivb cmp m1 m'1).
split.
intros k; generalize (H (xO k)); unfold In, MapsTo; simpl; auto.
intros k e e'; generalize (H0 (xO k) e e'); unfold In, MapsTo; simpl; auto.
- assert (Equal cmp m2 m'2).
+ assert (Equivb cmp m2 m'2).
split.
intros k; generalize (H (xI k)); unfold In, MapsTo; simpl; auto.
intros k e e'; generalize (H0 (xI k) e e'); unfold In, MapsTo; simpl; auto.
@@ -1043,8 +1067,8 @@ Module PositiveMap <: S with Module E:=PositiveOrderedTypeBits.
apply andb_true_intro; split; auto.
Qed.
- Lemma equal_2 : forall (A:Set)(m m':t A)(cmp:A->A->bool),
- equal cmp m m' = true -> Equal cmp m m'.
+ Lemma equal_2 : forall (A:Type)(m m':t A)(cmp:A->A->bool),
+ equal cmp m m' = true -> Equivb cmp m m'.
Proof.
induction m.
(* m = Leaf *)
@@ -1103,7 +1127,7 @@ Module PositiveMapAdditionalFacts.
(* Derivable from the Map interface *)
Theorem gsspec:
- forall (A:Set)(i j: positive) (x: A) (m: t A),
+ forall (A:Type)(i j: positive) (x: A) (m: t A),
find i (add j x m) = if peq_dec i j then Some x else find i m.
Proof.
intros.
@@ -1112,7 +1136,7 @@ Module PositiveMapAdditionalFacts.
(* Not derivable from the Map interface *)
Theorem gsident:
- forall (A:Set)(i: positive) (m: t A) (v: A),
+ forall (A:Type)(i: positive) (m: t A) (v: A),
find i m = Some v -> add i v m = m.
Proof.
induction i; intros; destruct m; simpl; simpl in H; try congruence.
@@ -1121,7 +1145,7 @@ Module PositiveMapAdditionalFacts.
Qed.
Lemma xmap2_lr :
- forall (A B : Set)(f g: option A -> option A -> option B)(m : t A),
+ forall (A B : Type)(f g: option A -> option A -> option B)(m : t A),
(forall (i j : option A), f i j = g j i) ->
xmap2_l f m = xmap2_r g m.
Proof.
@@ -1132,7 +1156,7 @@ Module PositiveMapAdditionalFacts.
Qed.
Theorem map2_commut:
- forall (A B: Set) (f g: option A -> option A -> option B),
+ forall (A B: Type) (f g: option A -> option A -> option B),
(forall (i j: option A), f i j = g j i) ->
forall (m1 m2: t A),
_map2 f m1 m2 = _map2 g m2 m1.
diff --git a/theories/FSets/FMapWeak.v b/theories/FSets/FMapWeak.v
deleted file mode 100644
index 1ad190a4..00000000
--- a/theories/FSets/FMapWeak.v
+++ /dev/null
@@ -1,15 +0,0 @@
-(***********************************************************************)
-(* 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: FMapWeak.v 8844 2006-05-22 17:22:36Z letouzey $ *)
-
-Require Export DecidableType.
-Require Export DecidableTypeEx.
-Require Export FMapWeakInterface.
-Require Export FMapWeakList.
-Require Export FMapWeakFacts. \ No newline at end of file
diff --git a/theories/FSets/FMapWeakFacts.v b/theories/FSets/FMapWeakFacts.v
deleted file mode 100644
index 18f73a3f..00000000
--- a/theories/FSets/FMapWeakFacts.v
+++ /dev/null
@@ -1,599 +0,0 @@
-(***********************************************************************)
-(* 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: FMapWeakFacts.v 8882 2006-05-31 21:55:30Z letouzey $ *)
-
-(** * Finite maps library *)
-
-(** This functor derives additional facts from [FMapWeakInterface.S]. These
- facts are mainly the specifications of [FMapWeakInterface.S] written using
- different styles: equivalence and boolean equalities.
-*)
-
-Require Import Bool.
-Require Import OrderedType.
-Require Export FMapWeakInterface.
-Set Implicit Arguments.
-Unset Strict Implicit.
-
-Module Facts (M: S).
-Import M.
-Import Logic. (* to unmask [eq] *)
-Import Peano. (* to unmask [lt] *)
-
-Lemma MapsTo_fun : forall (elt:Set) m x (e e':elt),
- MapsTo x e m -> MapsTo x e' m -> e=e'.
-Proof.
-intros.
-generalize (find_1 H) (find_1 H0); clear H H0.
-intros; rewrite H in H0; injection H0; auto.
-Qed.
-
-(** * Specifications written using equivalences *)
-
-Section IffSpec.
-Variable elt elt' elt'': Set.
-Implicit Type m: t elt.
-Implicit Type x y z: key.
-Implicit Type e: elt.
-
-Lemma MapsTo_iff : forall m x y e, E.eq x y -> (MapsTo x e m <-> MapsTo y e m).
-Proof.
-split; apply MapsTo_1; auto.
-Qed.
-
-Lemma In_iff : forall m x y, E.eq x y -> (In x m <-> In y m).
-Proof.
-unfold In.
-split; intros (e0,H0); exists e0.
-apply (MapsTo_1 H H0); auto.
-apply (MapsTo_1 (E.eq_sym H) H0); auto.
-Qed.
-
-Lemma find_mapsto_iff : forall m x e, MapsTo x e m <-> find x m = Some e.
-Proof.
-split; [apply find_1|apply find_2].
-Qed.
-
-Lemma not_find_mapsto_iff : forall m x, ~In x m <-> find x m = None.
-Proof.
-intros.
-generalize (find_mapsto_iff m x); destruct (find x m).
-split; intros; try discriminate.
-destruct H0.
-exists e; rewrite H; auto.
-split; auto.
-intros; intros (e,H1).
-rewrite H in H1; discriminate.
-Qed.
-
-Lemma mem_in_iff : forall m x, In x m <-> mem x m = true.
-Proof.
-split; [apply mem_1|apply mem_2].
-Qed.
-
-Lemma not_mem_in_iff : forall m x, ~In x m <-> mem x m = false.
-Proof.
-intros; rewrite mem_in_iff; destruct (mem x m); intuition.
-Qed.
-
-Lemma equal_iff : forall m m' cmp, Equal cmp m m' <-> equal cmp m m' = true.
-Proof.
-split; [apply equal_1|apply equal_2].
-Qed.
-
-Lemma empty_mapsto_iff : forall x e, MapsTo x e (empty elt) <-> False.
-Proof.
-intuition; apply (empty_1 H).
-Qed.
-
-Lemma empty_in_iff : forall x, In x (empty elt) <-> False.
-Proof.
-unfold In.
-split; [intros (e,H); rewrite empty_mapsto_iff in H|]; intuition.
-Qed.
-
-Lemma is_empty_iff : forall m, Empty m <-> is_empty m = true.
-Proof.
-split; [apply is_empty_1|apply is_empty_2].
-Qed.
-
-Lemma add_mapsto_iff : forall m x y e e',
- MapsTo y e' (add x e m) <->
- (E.eq x y /\ e=e') \/
- (~E.eq x y /\ MapsTo y e' m).
-Proof.
-intros.
-intuition.
-destruct (E.eq_dec x y); [left|right].
-split; auto.
-symmetry; apply (MapsTo_fun (e':=e) H); auto.
-split; auto; apply add_3 with x e; auto.
-subst; auto.
-Qed.
-
-Lemma add_in_iff : forall m x y e, In y (add x e m) <-> E.eq x y \/ In y m.
-Proof.
-unfold In; split.
-intros (e',H).
-destruct (E.eq_dec x y) as [E|E]; auto.
-right; exists e'; auto.
-apply (add_3 E H).
-destruct (E.eq_dec x y) as [E|E]; auto.
-intros.
-exists e; apply add_1; auto.
-intros [H|(e',H)].
-destruct E; auto.
-exists e'; apply add_2; auto.
-Qed.
-
-Lemma add_neq_mapsto_iff : forall m x y e e',
- ~ E.eq x y -> (MapsTo y e' (add x e m) <-> MapsTo y e' m).
-Proof.
-split; [apply add_3|apply add_2]; auto.
-Qed.
-
-Lemma add_neq_in_iff : forall m x y e,
- ~ E.eq x y -> (In y (add x e m) <-> In y m).
-Proof.
-split; intros (e',H0); exists e'.
-apply (add_3 H H0).
-apply add_2; auto.
-Qed.
-
-Lemma remove_mapsto_iff : forall m x y e,
- MapsTo y e (remove x m) <-> ~E.eq x y /\ MapsTo y e m.
-Proof.
-intros.
-split; intros.
-split.
-assert (In y (remove x m)) by (exists e; auto).
-intro H1; apply (remove_1 H1 H0).
-apply remove_3 with x; auto.
-apply remove_2; intuition.
-Qed.
-
-Lemma remove_in_iff : forall m x y, In y (remove x m) <-> ~E.eq x y /\ In y m.
-Proof.
-unfold In; split.
-intros (e,H).
-split.
-assert (In y (remove x m)) by (exists e; auto).
-intro H1; apply (remove_1 H1 H0).
-exists e; apply remove_3 with x; auto.
-intros (H,(e,H0)); exists e; apply remove_2; auto.
-Qed.
-
-Lemma remove_neq_mapsto_iff : forall m x y e,
- ~ E.eq x y -> (MapsTo y e (remove x m) <-> MapsTo y e m).
-Proof.
-split; [apply remove_3|apply remove_2]; auto.
-Qed.
-
-Lemma remove_neq_in_iff : forall m x y,
- ~ E.eq x y -> (In y (remove x m) <-> In y m).
-Proof.
-split; intros (e',H0); exists e'.
-apply (remove_3 H0).
-apply remove_2; auto.
-Qed.
-
-Lemma elements_mapsto_iff : forall m x e,
- MapsTo x e m <-> InA (@eq_key_elt _) (x,e) (elements m).
-Proof.
-split; [apply elements_1 | apply elements_2].
-Qed.
-
-Lemma elements_in_iff : forall m x,
- In x m <-> exists e, InA (@eq_key_elt _) (x,e) (elements m).
-Proof.
-unfold In; split; intros (e,H); exists e; [apply elements_1 | apply elements_2]; auto.
-Qed.
-
-Lemma map_mapsto_iff : forall m x b (f : elt -> elt'),
- MapsTo x b (map f m) <-> exists a, b = f a /\ MapsTo x a m.
-Proof.
-split.
-case_eq (find x m); intros.
-exists e.
-split.
-apply (MapsTo_fun (m:=map f m) (x:=x)); auto.
-apply find_2; auto.
-assert (In x (map f m)) by (exists b; auto).
-destruct (map_2 H1) as (a,H2).
-rewrite (find_1 H2) in H; discriminate.
-intros (a,(H,H0)).
-subst b; auto.
-Qed.
-
-Lemma map_in_iff : forall m x (f : elt -> elt'),
- In x (map f m) <-> In x m.
-Proof.
-split; intros; eauto.
-destruct H as (a,H).
-exists (f a); auto.
-Qed.
-
-Lemma mapi_in_iff : forall m x (f:key->elt->elt'),
- In x (mapi f m) <-> In x m.
-Proof.
-split; intros; eauto.
-destruct H as (a,H).
-destruct (mapi_1 f H) as (y,(H0,H1)).
-exists (f y a); auto.
-Qed.
-
-(* Unfortunately, we don't have simple equivalences for [mapi]
- and [MapsTo]. The only correct one needs compatibility of [f]. *)
-
-Lemma mapi_inv : forall m x b (f : key -> elt -> elt'),
- MapsTo x b (mapi f m) ->
- exists a, exists y, E.eq y x /\ b = f y a /\ MapsTo x a m.
-Proof.
-intros; case_eq (find x m); intros.
-exists e.
-destruct (@mapi_1 _ _ m x e f) as (y,(H1,H2)).
-apply find_2; auto.
-exists y; repeat split; auto.
-apply (MapsTo_fun (m:=mapi f m) (x:=x)); auto.
-assert (In x (mapi f m)) by (exists b; auto).
-destruct (mapi_2 H1) as (a,H2).
-rewrite (find_1 H2) in H0; discriminate.
-Qed.
-
-Lemma mapi_1bis : forall m x e (f:key->elt->elt'),
- (forall x y e, E.eq x y -> f x e = f y e) ->
- MapsTo x e m -> MapsTo x (f x e) (mapi f m).
-Proof.
-intros.
-destruct (mapi_1 f H0) as (y,(H1,H2)).
-replace (f x e) with (f y e) by auto.
-auto.
-Qed.
-
-Lemma mapi_mapsto_iff : forall m x b (f:key->elt->elt'),
- (forall x y e, E.eq x y -> f x e = f y e) ->
- (MapsTo x b (mapi f m) <-> exists a, b = f x a /\ MapsTo x a m).
-Proof.
-split.
-intros.
-destruct (mapi_inv H0) as (a,(y,(H1,(H2,H3)))).
-exists a; split; auto.
-subst b; auto.
-intros (a,(H0,H1)).
-subst b.
-apply mapi_1bis; auto.
-Qed.
-
-(** Things are even worse for [map2] : we don't try to state any
- equivalence, see instead boolean results below. *)
-
-End IffSpec.
-
-(** Useful tactic for simplifying expressions like [In y (add x e (remove z m))] *)
-
-Ltac map_iff :=
- repeat (progress (
- rewrite add_mapsto_iff || rewrite add_in_iff ||
- rewrite remove_mapsto_iff || rewrite remove_in_iff ||
- rewrite empty_mapsto_iff || rewrite empty_in_iff ||
- rewrite map_mapsto_iff || rewrite map_in_iff ||
- rewrite mapi_in_iff)).
-
-(** * Specifications written using boolean predicates *)
-
-Section BoolSpec.
-
-Definition eqb x y := if E.eq_dec x y then true else false.
-
-Lemma mem_find_b : forall (elt:Set)(m:t elt)(x:key), mem x m = if find x m then true else false.
-Proof.
-intros.
-generalize (find_mapsto_iff m x)(mem_in_iff m x); unfold In.
-destruct (find x m); destruct (mem x m); auto.
-intros.
-rewrite <- H0; exists e; rewrite H; auto.
-intuition.
-destruct H0 as (e,H0).
-destruct (H e); intuition discriminate.
-Qed.
-
-Variable elt elt' elt'' : Set.
-Implicit Types m : t elt.
-Implicit Types x y z : key.
-Implicit Types e : elt.
-
-Lemma mem_b : forall m x y, E.eq x y -> mem x m = mem y m.
-Proof.
-intros.
-generalize (mem_in_iff m x) (mem_in_iff m y)(In_iff m H).
-destruct (mem x m); destruct (mem y m); intuition.
-Qed.
-
-Lemma find_o : forall m x y, E.eq x y -> find x m = find y m.
-Proof.
-intros.
-generalize (find_mapsto_iff m x) (find_mapsto_iff m y) (fun e => MapsTo_iff m e H).
-destruct (find x m); destruct (find y m); intros.
-rewrite <- H0; rewrite H2; rewrite H1; auto.
-symmetry; rewrite <- H1; rewrite <- H2; rewrite H0; auto.
-rewrite <- H0; rewrite H2; rewrite H1; auto.
-auto.
-Qed.
-
-Lemma empty_o : forall x, find x (empty elt) = None.
-Proof.
-intros.
-case_eq (find x (empty elt)); intros; auto.
-generalize (find_2 H).
-rewrite empty_mapsto_iff; intuition.
-Qed.
-
-Lemma empty_a : forall x, mem x (empty elt) = false.
-Proof.
-intros.
-case_eq (mem x (empty elt)); intros; auto.
-generalize (mem_2 H).
-rewrite empty_in_iff; intuition.
-Qed.
-
-Lemma add_eq_o : forall m x y e,
- E.eq x y -> find y (add x e m) = Some e.
-Proof.
-auto.
-Qed.
-
-Lemma add_neq_o : forall m x y e,
- ~ E.eq x y -> find y (add x e m) = find y m.
-Proof.
-intros.
-case_eq (find y m); intros; auto.
-case_eq (find y (add x e m)); intros; auto.
-rewrite <- H0; symmetry.
-apply find_1; apply add_3 with x e; auto.
-Qed.
-Hint Resolve add_neq_o.
-
-Lemma add_o : forall m x y e,
- find y (add x e m) = if E.eq_dec x y then Some e else find y m.
-Proof.
-intros; destruct (E.eq_dec x y); auto.
-Qed.
-
-Lemma add_eq_b : forall m x y e,
- E.eq x y -> mem y (add x e m) = true.
-Proof.
-intros; rewrite mem_find_b; rewrite add_eq_o; auto.
-Qed.
-
-Lemma add_neq_b : forall m x y e,
- ~E.eq x y -> mem y (add x e m) = mem y m.
-Proof.
-intros; do 2 rewrite mem_find_b; rewrite add_neq_o; auto.
-Qed.
-
-Lemma add_b : forall m x y e,
- mem y (add x e m) = eqb x y || mem y m.
-Proof.
-intros; do 2 rewrite mem_find_b; rewrite add_o; unfold eqb.
-destruct (E.eq_dec x y); simpl; auto.
-Qed.
-
-Lemma remove_eq_o : forall m x y,
- E.eq x y -> find y (remove x m) = None.
-Proof.
-intros.
-generalize (remove_1 (m:=m) H).
-generalize (find_mapsto_iff (remove x m) y).
-destruct (find y (remove x m)); auto.
-destruct 2.
-exists e; rewrite H0; auto.
-Qed.
-Hint Resolve remove_eq_o.
-
-Lemma remove_neq_o : forall m x y,
- ~ E.eq x y -> find y (remove x m) = find y m.
-Proof.
-intros.
-case_eq (find y m); intros; auto.
-case_eq (find y (remove x m)); intros; auto.
-rewrite <- H0; symmetry.
-apply find_1; apply remove_3 with x; auto.
-Qed.
-Hint Resolve remove_neq_o.
-
-Lemma remove_o : forall m x y,
- find y (remove x m) = if E.eq_dec x y then None else find y m.
-Proof.
-intros; destruct (E.eq_dec x y); auto.
-Qed.
-
-Lemma remove_eq_b : forall m x y,
- E.eq x y -> mem y (remove x m) = false.
-Proof.
-intros; rewrite mem_find_b; rewrite remove_eq_o; auto.
-Qed.
-
-Lemma remove_neq_b : forall m x y,
- ~ E.eq x y -> mem y (remove x m) = mem y m.
-Proof.
-intros; do 2 rewrite mem_find_b; rewrite remove_neq_o; auto.
-Qed.
-
-Lemma remove_b : forall m x y,
- mem y (remove x m) = negb (eqb x y) && mem y m.
-Proof.
-intros; do 2 rewrite mem_find_b; rewrite remove_o; unfold eqb.
-destruct (E.eq_dec x y); auto.
-Qed.
-
-Definition option_map (A:Set)(B:Set)(f:A->B)(o:option A) : option B :=
- match o with
- | Some a => Some (f a)
- | None => None
- end.
-
-Lemma map_o : forall m x (f:elt->elt'),
- find x (map f m) = option_map f (find x m).
-Proof.
-intros.
-generalize (find_mapsto_iff (map f m) x) (find_mapsto_iff m x)
- (fun b => map_mapsto_iff m x b f).
-destruct (find x (map f m)); destruct (find x m); simpl; auto; intros.
-rewrite <- H; rewrite H1; exists e0; rewrite H0; auto.
-destruct (H e) as [_ H2].
-rewrite H1 in H2.
-destruct H2 as (a,(_,H2)); auto.
-rewrite H0 in H2; discriminate.
-rewrite <- H; rewrite H1; exists e; rewrite H0; auto.
-Qed.
-
-Lemma map_b : forall m x (f:elt->elt'),
- mem x (map f m) = mem x m.
-Proof.
-intros; do 2 rewrite mem_find_b; rewrite map_o.
-destruct (find x m); simpl; auto.
-Qed.
-
-Lemma mapi_b : forall m x (f:key->elt->elt'),
- mem x (mapi f m) = mem x m.
-Proof.
-intros.
-generalize (mem_in_iff (mapi f m) x) (mem_in_iff m x) (mapi_in_iff m x f).
-destruct (mem x (mapi f m)); destruct (mem x m); simpl; auto; intros.
-symmetry; rewrite <- H0; rewrite <- H1; rewrite H; auto.
-rewrite <- H; rewrite H1; rewrite H0; auto.
-Qed.
-
-Lemma mapi_o : forall m x (f:key->elt->elt'),
- (forall x y e, E.eq x y -> f x e = f y e) ->
- find x (mapi f m) = option_map (f x) (find x m).
-Proof.
-intros.
-generalize (find_mapsto_iff (mapi f m) x) (find_mapsto_iff m x)
- (fun b => mapi_mapsto_iff m x b H).
-destruct (find x (mapi f m)); destruct (find x m); simpl; auto; intros.
-rewrite <- H0; rewrite H2; exists e0; rewrite H1; auto.
-destruct (H0 e) as [_ H3].
-rewrite H2 in H3.
-destruct H3 as (a,(_,H3)); auto.
-rewrite H1 in H3; discriminate.
-rewrite <- H0; rewrite H2; exists e; rewrite H1; auto.
-Qed.
-
-Lemma map2_1bis : forall (m: t elt)(m': t elt') x
- (f:option elt->option elt'->option elt''),
- f None None = None ->
- find x (map2 f m m') = f (find x m) (find x m').
-Proof.
-intros.
-case_eq (find x m); intros.
-rewrite <- H0.
-apply map2_1; auto.
-left; exists e; auto.
-case_eq (find x m'); intros.
-rewrite <- H0; rewrite <- H1.
-apply map2_1; auto.
-right; exists e; auto.
-rewrite H.
-case_eq (find x (map2 f m m')); intros; auto.
-assert (In x (map2 f m m')) by (exists e; auto).
-destruct (map2_2 H3) as [(e0,H4)|(e0,H4)].
-rewrite (find_1 H4) in H0; discriminate.
-rewrite (find_1 H4) in H1; discriminate.
-Qed.
-
-Fixpoint findA (A B:Set)(f : A -> bool) (l:list (A*B)) : option B :=
- match l with
- | nil => None
- | (a,b)::l => if f a then Some b else findA f l
- end.
-
-Lemma findA_NoDupA :
- forall (A B:Set)
- (eqA:A->A->Prop)
- (eqA_sym: forall a b, eqA a b -> eqA b a)
- (eqA_trans: forall a b c, eqA a b -> eqA b c -> eqA a c)
- (eqA_dec : forall a a', { eqA a a' }+{~eqA a a' })
- (l:list (A*B))(x:A)(e:B),
- NoDupA (fun p p' => eqA (fst p) (fst p')) l ->
- (InA (fun p p' => eqA (fst p) (fst p') /\ snd p = snd p') (x,e) l <->
- findA (fun y:A => if eqA_dec x y then true else false) l = Some e).
-Proof.
-induction l; simpl; intros.
-split; intros; try discriminate.
-inversion H0.
-destruct a as (y,e').
-inversion_clear H.
-split; intros.
-inversion_clear H.
-simpl in *; destruct H2; subst e'.
-destruct (eqA_dec x y); intuition.
-destruct (eqA_dec x y); simpl.
-destruct H0.
-generalize e0 H2 eqA_trans eqA_sym; clear.
-induction l.
-inversion 2.
-inversion_clear 2; intros; auto.
-destruct a.
-compute in H; destruct H.
-subst b.
-constructor 1; auto.
-simpl.
-apply eqA_trans with x; auto.
-rewrite <- IHl; auto.
-destruct (eqA_dec x y); simpl in *.
-inversion H; clear H; intros; subst e'; auto.
-constructor 2.
-rewrite IHl; auto.
-Qed.
-
-Lemma elements_o : forall m x,
- find x m = findA (eqb x) (elements m).
-Proof.
-intros.
-assert (forall e, find x m = Some e <-> InA (eq_key_elt (elt:=elt)) (x,e) (elements m)).
- intros; rewrite <- find_mapsto_iff; apply elements_mapsto_iff.
-assert (NoDupA (eq_key (elt:=elt)) (elements m)).
- exact (elements_3 m).
-generalize (fun e => @findA_NoDupA _ _ _ E.eq_sym E.eq_trans E.eq_dec (elements m) x e H0).
-unfold eqb.
-destruct (find x m); destruct (findA (fun y : E.t => if E.eq_dec x y then true else false) (elements m));
- simpl; auto; intros.
-symmetry; rewrite <- H1; rewrite <- H; auto.
-symmetry; rewrite <- H1; rewrite <- H; auto.
-rewrite H; rewrite H1; auto.
-Qed.
-
-Lemma elements_b : forall m x, mem x m = existsb (fun p => eqb x (fst p)) (elements m).
-Proof.
-intros.
-generalize (mem_in_iff m x)(elements_in_iff m x)
- (existsb_exists (fun p => eqb x (fst p)) (elements m)).
-destruct (mem x m); destruct (existsb (fun p => eqb x (fst p)) (elements m)); auto; intros.
-symmetry; rewrite H1.
-destruct H0 as (H0,_).
-destruct H0 as (e,He); [ intuition |].
-rewrite InA_alt in He.
-destruct He as ((y,e'),(Ha1,Ha2)).
-compute in Ha1; destruct Ha1; subst e'.
-exists (y,e); split; simpl; auto.
-unfold eqb; destruct (E.eq_dec x y); intuition.
-rewrite <- H; rewrite H0.
-destruct H1 as (H1,_).
-destruct H1 as ((y,e),(Ha1,Ha2)); [intuition|].
-simpl in Ha2.
-unfold eqb in *; destruct (E.eq_dec x y); auto; try discriminate.
-exists e; rewrite InA_alt.
-exists (y,e); intuition.
-compute; auto.
-Qed.
-
-End BoolSpec.
-
-End Facts.
diff --git a/theories/FSets/FMapWeakInterface.v b/theories/FSets/FMapWeakInterface.v
deleted file mode 100644
index b6df4da5..00000000
--- a/theories/FSets/FMapWeakInterface.v
+++ /dev/null
@@ -1,201 +0,0 @@
-(***********************************************************************)
-(* 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: FMapWeakInterface.v 8639 2006-03-16 19:21:55Z letouzey $ *)
-
-(** * Finite map library *)
-
-(** This file proposes an interface for finite maps over keys with decidable
- equality, but no decidable order. *)
-
-Set Implicit Arguments.
-Unset Strict Implicit.
-Require Import FSetInterface.
-Require Import FSetWeakInterface.
-
-Module Type S.
-
- Declare Module E : DecidableType.
-
- Definition key := E.t.
-
- Parameter t : Set -> Set. (** the abstract type of maps *)
-
- Section Types.
-
- Variable elt:Set.
-
- Parameter empty : t elt.
- (** The empty map. *)
-
- Parameter is_empty : t elt -> bool.
- (** Test whether a map is empty or not. *)
-
- Parameter add : key -> elt -> t elt -> t elt.
- (** [add x y m] returns a map containing the same bindings as [m],
- plus a binding of [x] to [y]. If [x] was already bound in [m],
- its previous binding disappears. *)
-
- Parameter find : key -> t elt -> option elt.
- (** [find x m] returns the current binding of [x] in [m],
- or raises [Not_found] if no such binding exists.
- NB: in Coq, the exception mechanism becomes a option type. *)
-
- Parameter remove : key -> t elt -> t elt.
- (** [remove x m] returns a map containing the same bindings as [m],
- except for [x] which is unbound in the returned map. *)
-
- Parameter mem : key -> t elt -> bool.
- (** [mem x m] returns [true] if [m] contains a binding for [x],
- and [false] otherwise. *)
-
- (** Coq comment: [iter] is useless in a purely functional world *)
- (** val iter : (key -> 'a -> unit) -> 'a t -> unit *)
- (** iter f m applies f to all bindings in map m. f receives the key as
- first argument, and the associated value as second argument.
- The bindings are passed to f in increasing order with respect to the
- ordering over the type of the keys. Only current bindings are
- presented to f: bindings hidden by more recent bindings are not
- passed to f. *)
-
- Variable elt' : Set.
- Variable elt'': Set.
-
- Parameter map : (elt -> elt') -> t elt -> t elt'.
- (** [map f m] returns a map with same domain as [m], where the associated
- value a of all bindings of [m] has been replaced by the result of the
- application of [f] to [a]. The bindings are passed to [f] in
- increasing order with respect to the ordering over the type of the
- keys. *)
-
- Parameter mapi : (key -> elt -> elt') -> t elt -> t elt'.
- (** Same as [S.map], but the function receives as arguments both the
- key and the associated value for each binding of the map. *)
-
- Parameter map2 : (option elt -> option elt' -> option elt'') -> t elt -> t elt' -> t elt''.
- (** Not present in Ocaml.
- [map f m m'] creates a new map whose bindings belong to the ones of either
- [m] or [m']. The presence and value for a key [k] is determined by [f e e']
- where [e] and [e'] are the (optional) bindings of [k] in [m] and [m']. *)
-
- Parameter elements : t elt -> list (key*elt).
- (** Not present in Ocaml.
- [elements m] returns an assoc list corresponding to the bindings of [m].
- Elements of this list are sorted with respect to their first components.
- Useful to specify [fold] ... *)
-
- Parameter fold : forall A: Set, (key -> elt -> A -> A) -> t elt -> A -> A.
- (** [fold f m a] computes [(f kN dN ... (f k1 d1 a)...)],
- where [k1] ... [kN] are the keys of all bindings in [m]
- (in increasing order), and [d1] ... [dN] are the associated data. *)
-
- Parameter equal : (elt -> elt -> bool) -> t elt -> t elt -> bool.
- (** [equal cmp m1 m2] tests whether the maps [m1] and [m2] are equal,
- that is, contain equal keys and associate them with equal data.
- [cmp] is the equality predicate used to compare the data associated
- with the keys. *)
-
- Section Spec.
-
- Variable m m' m'' : t elt.
- Variable x y z : key.
- Variable e e' : elt.
-
- Parameter MapsTo : key -> elt -> t elt -> Prop.
-
- Definition In (k:key)(m: t elt) : Prop := exists e:elt, MapsTo k e m.
-
- Definition Empty m := forall (a : key)(e:elt) , ~ MapsTo a e m.
-
- Definition eq_key (p p':key*elt) := E.eq (fst p) (fst p').
-
- Definition eq_key_elt (p p':key*elt) :=
- E.eq (fst p) (fst p') /\ (snd p) = (snd p').
-
- (** Specification of [MapsTo] *)
- Parameter MapsTo_1 : E.eq x y -> MapsTo x e m -> MapsTo y e m.
-
- (** Specification of [mem] *)
- Parameter mem_1 : In x m -> mem x m = true.
- Parameter mem_2 : mem x m = true -> In x m.
-
- (** Specification of [empty] *)
- Parameter empty_1 : Empty empty.
-
- (** Specification of [is_empty] *)
- Parameter is_empty_1 : Empty m -> is_empty m = true.
- Parameter is_empty_2 : is_empty m = true -> Empty m.
-
- (** Specification of [add] *)
- Parameter add_1 : E.eq x y -> MapsTo y e (add x e m).
- Parameter add_2 : ~ E.eq x y -> MapsTo y e m -> MapsTo y e (add x e' m).
- Parameter add_3 : ~ E.eq x y -> MapsTo y e (add x e' m) -> MapsTo y e m.
-
- (** Specification of [remove] *)
- Parameter remove_1 : E.eq x y -> ~ In y (remove x m).
- Parameter remove_2 : ~ E.eq x y -> MapsTo y e m -> MapsTo y e (remove x m).
- Parameter remove_3 : MapsTo y e (remove x m) -> MapsTo y e m.
-
- (** Specification of [find] *)
- Parameter find_1 : MapsTo x e m -> find x m = Some e.
- Parameter find_2 : find x m = Some e -> MapsTo x e m.
-
- (** Specification of [elements] *)
- Parameter elements_1 :
- MapsTo x e m -> InA eq_key_elt (x,e) (elements m).
- Parameter elements_2 :
- InA eq_key_elt (x,e) (elements m) -> MapsTo x e m.
- Parameter elements_3 : NoDupA eq_key (elements m).
-
- (** Specification of [fold] *)
- Parameter fold_1 :
- forall (A : Set) (i : A) (f : key -> elt -> A -> A),
- fold f m i = fold_left (fun a p => f (fst p) (snd p) a) (elements m) i.
-
- Definition Equal cmp m m' :=
- (forall k, In k m <-> In k m') /\
- (forall k e e', MapsTo k e m -> MapsTo k e' m' -> cmp e e' = true).
-
- Variable cmp : elt -> elt -> bool.
-
- (** Specification of [equal] *)
- Parameter equal_1 : Equal cmp m m' -> equal cmp m m' = true.
- Parameter equal_2 : equal cmp m m' = true -> Equal cmp m m'.
-
- End Spec.
- End Types.
-
- (** Specification of [map] *)
- Parameter map_1 : forall (elt elt':Set)(m: t elt)(x:key)(e:elt)(f:elt->elt'),
- MapsTo x e m -> MapsTo x (f e) (map f m).
- Parameter map_2 : forall (elt elt':Set)(m: t elt)(x:key)(f:elt->elt'),
- In x (map f m) -> In x m.
-
- (** Specification of [mapi] *)
- Parameter mapi_1 : forall (elt elt':Set)(m: t elt)(x:key)(e:elt)
- (f:key->elt->elt'), MapsTo x e m ->
- exists y, E.eq y x /\ MapsTo x (f y e) (mapi f m).
- Parameter mapi_2 : forall (elt elt':Set)(m: t elt)(x:key)
- (f:key->elt->elt'), In x (mapi f m) -> In x m.
-
- (** Specification of [map2] *)
- Parameter map2_1 : forall (elt elt' elt'':Set)(m: t elt)(m': t elt')
- (x:key)(f:option elt->option elt'->option elt''),
- In x m \/ In x m' ->
- find x (map2 f m m') = f (find x m) (find x m').
-
- Parameter map2_2 : forall (elt elt' elt'':Set)(m: t elt)(m': t elt')
- (x:key)(f:option elt->option elt'->option elt''),
- In x (map2 f m m') -> In x m \/ In x m'.
-
- Hint Immediate MapsTo_1 mem_2 is_empty_2.
-
- Hint Resolve mem_1 is_empty_1 is_empty_2 add_1 add_2 add_3 remove_1
- remove_2 remove_3 find_1 find_2 fold_1 map_1 map_2 mapi_1 mapi_2.
-
-End S.
diff --git a/theories/FSets/FMapWeakList.v b/theories/FSets/FMapWeakList.v
index 890485a8..be09e41a 100644
--- a/theories/FSets/FMapWeakList.v
+++ b/theories/FSets/FMapWeakList.v
@@ -6,39 +6,28 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: FMapWeakList.v 8985 2006-06-23 16:12:45Z jforest $ *)
+(* $Id: FMapWeakList.v 10616 2008-03-04 17:33:35Z letouzey $ *)
(** * Finite map library *)
(** This file proposes an implementation of the non-dependant interface
- [FMapInterface.S] using lists of pairs, unordered but without redundancy. *)
+ [FMapInterface.WS] using lists of pairs, unordered but without redundancy. *)
-Require Import FSetInterface.
-Require Import FSetWeakInterface.
-Require Import FMapWeakInterface.
+Require Import FMapInterface.
Set Implicit Arguments.
Unset Strict Implicit.
-Arguments Scope list [type_scope].
-
Module Raw (X:DecidableType).
-Module PX := KeyDecidableType X.
-Import PX.
+Module Import PX := KeyDecidableType X.
Definition key := X.t.
-Definition t (elt:Set) := list (X.t * elt).
+Definition t (elt:Type) := list (X.t * elt).
Section Elt.
-Variable elt : Set.
-
-(* now in KeyDecidableType:
-Definition eqk (p p':key*elt) := X.eq (fst p) (fst p').
-Definition eqke (p p':key*elt) :=
- X.eq (fst p) (fst p') /\ (snd p) = (snd p').
-*)
+Variable elt : Type.
Notation eqk := (eqk (elt:=elt)).
Notation eqke := (eqke (elt:=elt)).
@@ -221,10 +210,10 @@ Proof.
destruct a as (x',e').
simpl; case (X.eq_dec x x'); inversion_clear Hm; auto.
constructor; auto.
- swap H.
+ contradict H.
apply InA_eqk with (x,e); auto.
constructor; auto.
- swap H; apply add_3' with x e; auto.
+ contradict H; apply add_3' with x e; auto.
Qed.
(* Not part of the exported specifications, used later for [combine]. *)
@@ -272,8 +261,8 @@ Proof.
inversion_clear Hm.
subst.
- swap H0.
- destruct H2 as (e,H2); unfold PX.MapsTo in H2.
+ contradict H0.
+ destruct H0 as (e,H2); unfold PX.MapsTo in H2.
apply InA_eqk with (y,e); auto.
compute; apply X.eq_trans with x; auto.
@@ -323,7 +312,7 @@ Proof.
destruct a as (x',e').
simpl; case (X.eq_dec x x'); auto.
constructor; auto.
- swap H; apply remove_3' with x; auto.
+ contradict H; apply remove_3' with x; auto.
Qed.
(** * [elements] *)
@@ -340,20 +329,20 @@ Proof.
auto.
Qed.
-Lemma elements_3 : forall m (Hm:NoDupA m), NoDupA (elements m).
+Lemma elements_3w : forall m (Hm:NoDupA m), NoDupA (elements m).
Proof.
auto.
Qed.
(** * [fold] *)
-Function fold (A:Set)(f:key->elt->A->A)(m:t elt) (acc : A) {struct m} : A :=
+Function fold (A:Type)(f:key->elt->A->A)(m:t elt) (acc : A) {struct m} : A :=
match m with
| nil => acc
| (k,e)::m' => fold f m' (f k e acc)
end.
-Lemma fold_1 : forall m (A:Set)(i:A)(f:key->elt->A->A),
+Lemma fold_1 : forall m (A:Type)(i:A)(f:key->elt->A->A),
fold f m i = fold_left (fun a p => f (fst p) (snd p) a) (elements m) i.
Proof.
intros; functional induction (@fold A f m i); auto.
@@ -377,7 +366,7 @@ Definition Submap cmp m m' :=
(forall k, In k m -> In k m') /\
(forall k e e', MapsTo k e m -> MapsTo k e' m' -> cmp e e' = true).
-Definition Equal cmp m m' :=
+Definition Equivb cmp m m' :=
(forall k, In k m <-> In k m') /\
(forall k e e', MapsTo k e m -> MapsTo k e' m' -> cmp e e' = true).
@@ -444,17 +433,17 @@ Qed.
(** Specification of [equal] *)
Lemma equal_1 : forall m (Hm:NoDupA m) m' (Hm': NoDupA m') cmp,
- Equal cmp m m' -> equal cmp m m' = true.
+ Equivb cmp m m' -> equal cmp m m' = true.
Proof.
- unfold Equal, equal.
+ unfold Equivb, equal.
intuition.
apply andb_true_intro; split; apply submap_1; unfold Submap; firstorder.
Qed.
Lemma equal_2 : forall m (Hm:NoDupA m) m' (Hm':NoDupA m') cmp,
- equal cmp m m' = true -> Equal cmp m m'.
+ equal cmp m m' = true -> Equivb cmp m m'.
Proof.
- unfold Equal, equal.
+ unfold Equivb, equal.
intros.
destruct (andb_prop _ _ H); clear H.
generalize (submap_2 Hm Hm' H0).
@@ -462,7 +451,7 @@ Proof.
firstorder.
Qed.
-Variable elt':Set.
+Variable elt':Type.
(** * [map] and [mapi] *)
@@ -483,7 +472,7 @@ Section Elt2.
(* A new section is necessary for previous definitions to work
with different [elt], especially [MapsTo]... *)
-Variable elt elt' : Set.
+Variable elt elt' : Type.
(** Specification of [map] *)
@@ -533,12 +522,12 @@ Proof.
destruct a as (x',e').
inversion_clear Hm.
constructor; auto.
- swap H.
+ contradict H.
(* il faut un map_1 avec eqk au lieu de eqke *)
clear IHm H0.
induction m; simpl in *; auto.
- inversion H1.
- destruct a; inversion H1; auto.
+ inversion H.
+ destruct a; inversion H; auto.
Qed.
(** Specification of [mapi] *)
@@ -593,17 +582,17 @@ Proof.
destruct a as (x',e').
inversion_clear Hm; auto.
constructor; auto.
- swap H.
+ contradict H.
clear IHm H0.
induction m; simpl in *; auto.
- inversion_clear H1.
- destruct a; inversion_clear H1; auto.
+ inversion_clear H.
+ destruct a; inversion_clear H; auto.
Qed.
End Elt2.
Section Elt3.
-Variable elt elt' elt'' : Set.
+Variable elt elt' elt'' : Type.
Notation oee' := (option elt * option elt')%type.
@@ -613,7 +602,7 @@ Definition combine_l (m:t elt)(m':t elt') : t oee' :=
Definition combine_r (m:t elt)(m':t elt') : t oee' :=
mapi (fun k e' => (find k m, Some e')) m'.
-Definition fold_right_pair (A B C:Set)(f:A->B->C->C)(l:list (A*B))(i:C) :=
+Definition fold_right_pair (A B C:Type)(f:A->B->C->C)(l:list (A*B))(i:C) :=
List.fold_right (fun p => f (fst p) (snd p)) i l.
Definition combine (m:t elt)(m':t elt') : t oee' :=
@@ -737,7 +726,7 @@ Qed.
Variable f : option elt -> option elt' -> option elt''.
-Definition option_cons (A:Set)(k:key)(o:option A)(l:list (key*A)) :=
+Definition option_cons (A:Type)(k:key)(o:option A)(l:list (key*A)) :=
match o with
| Some e => (k,e)::l
| None => l
@@ -765,13 +754,13 @@ Proof.
inversion_clear H1.
destruct a; destruct o; simpl; auto.
constructor; auto.
- swap H.
+ contradict H.
clear IHl1.
induction l1.
- inversion H1.
+ inversion H.
inversion_clear H0.
destruct a; destruct o; simpl in *; auto.
- inversion_clear H1; auto.
+ inversion_clear H; auto.
Qed.
Definition at_least_one_then_f (o:option elt)(o':option elt') :=
@@ -807,7 +796,7 @@ Proof.
rewrite H2.
unfold f'; simpl.
destruct (f oo oo'); simpl.
- destruct (X.eq_dec x k); try absurd_hyp n; auto.
+ destruct (X.eq_dec x k); try contradict n; auto.
destruct (IHm0 H1) as (_,H4); apply H4; auto.
case_eq (find x m0); intros; auto.
elim H0.
@@ -817,7 +806,7 @@ Proof.
(* k < x *)
unfold f'; simpl.
destruct (f oo oo'); simpl.
- destruct (X.eq_dec x k); [ absurd_hyp n; auto | auto].
+ destruct (X.eq_dec x k); [ contradict n; auto | auto].
destruct (IHm0 H1) as (H3,_); apply H3; auto.
destruct (IHm0 H1) as (H3,_); apply H3; auto.
@@ -831,7 +820,7 @@ Proof.
(* k < x *)
unfold f'; simpl.
destruct (f oo oo'); simpl.
- destruct (X.eq_dec x k); [ absurd_hyp n; auto | auto].
+ destruct (X.eq_dec x k); [ contradict n; auto | auto].
destruct (IHm0 H1) as (_,H4); apply H4; auto.
destruct (IHm0 H1) as (_,H4); apply H4; auto.
Qed.
@@ -873,18 +862,18 @@ End Elt3.
End Raw.
-Module Make (X: DecidableType) <: S with Module E:=X.
+Module Make (X: DecidableType) <: WS with Module E:=X.
Module Raw := Raw X.
Module E := X.
Definition key := E.t.
- Record slist (elt:Set) : Set :=
+ Record slist (elt:Type) :=
{this :> Raw.t elt; NoDup : NoDupA (@Raw.PX.eqk elt) this}.
- Definition t (elt:Set) := slist elt.
+ Definition t (elt:Type) := slist elt.
Section Elt.
- Variable elt elt' elt'':Set.
+ Variable elt elt' elt'':Type.
Implicit Types m : t elt.
Implicit Types x y : key.
@@ -901,13 +890,18 @@ Section Elt.
Definition map2 f m (m':t elt') : t elt'' :=
Build_slist (Raw.map2_NoDup f m.(NoDup) m'.(NoDup)).
Definition elements m : list (key*elt) := @Raw.elements elt m.(this).
- Definition fold (A:Set)(f:key->elt->A->A) m (i:A) : A := @Raw.fold elt A f m.(this) i.
+ Definition cardinal m := length m.(this).
+ Definition fold (A:Type)(f:key->elt->A->A) m (i:A) : A := @Raw.fold elt A f m.(this) i.
Definition equal cmp m m' : bool := @Raw.equal elt cmp m.(this) m'.(this).
-
Definition MapsTo x e m : Prop := Raw.PX.MapsTo x e m.(this).
Definition In x m : Prop := Raw.PX.In x m.(this).
Definition Empty m : Prop := Raw.Empty m.(this).
- Definition Equal cmp m m' : Prop := @Raw.Equal elt cmp m.(this) m'.(this).
+
+ Definition Equal m m' := forall y, find y m = find y m'.
+ Definition Equiv (eq_elt:elt->elt->Prop) m m' :=
+ (forall k, In k m <-> In k m') /\
+ (forall k e e', MapsTo k e m -> MapsTo k e' m' -> eq_elt e e').
+ Definition Equivb cmp m m' : Prop := @Raw.Equivb elt cmp m.(this) m'.(this).
Definition eq_key : (key*elt) -> (key*elt) -> Prop := @Raw.PX.eqk elt.
Definition eq_key_elt : (key*elt) -> (key*elt) -> Prop:= @Raw.PX.eqke elt.
@@ -951,36 +945,39 @@ Section Elt.
Proof. intros m; exact (@Raw.elements_1 elt m.(this)). Qed.
Lemma elements_2 : forall m x e, InA eq_key_elt (x,e) (elements m) -> MapsTo x e m.
Proof. intros m; exact (@Raw.elements_2 elt m.(this)). Qed.
- Lemma elements_3 : forall m, NoDupA eq_key (elements m).
- Proof. intros m; exact (@Raw.elements_3 elt m.(this) m.(NoDup)). Qed.
+ Lemma elements_3w : forall m, NoDupA eq_key (elements m).
+ Proof. intros m; exact (@Raw.elements_3w elt m.(this) m.(NoDup)). Qed.
+
+ Lemma cardinal_1 : forall m, cardinal m = length (elements m).
+ Proof. intros; reflexivity. Qed.
- Lemma fold_1 : forall m (A : Set) (i : A) (f : key -> elt -> A -> A),
+ Lemma fold_1 : forall m (A : Type) (i : A) (f : key -> elt -> A -> A),
fold f m i = fold_left (fun a p => f (fst p) (snd p) a) (elements m) i.
Proof. intros m; exact (@Raw.fold_1 elt m.(this)). Qed.
- Lemma equal_1 : forall m m' cmp, Equal cmp m m' -> equal cmp m m' = true.
+ Lemma equal_1 : forall m m' cmp, Equivb cmp m m' -> equal cmp m m' = true.
Proof. intros m m'; exact (@Raw.equal_1 elt m.(this) m.(NoDup) m'.(this) m'.(NoDup)). Qed.
- Lemma equal_2 : forall m m' cmp, equal cmp m m' = true -> Equal cmp m m'.
+ Lemma equal_2 : forall m m' cmp, equal cmp m m' = true -> Equivb cmp m m'.
Proof. intros m m'; exact (@Raw.equal_2 elt m.(this) m.(NoDup) m'.(this) m'.(NoDup)). Qed.
End Elt.
- Lemma map_1 : forall (elt elt':Set)(m: t elt)(x:key)(e:elt)(f:elt->elt'),
+ Lemma map_1 : forall (elt elt':Type)(m: t elt)(x:key)(e:elt)(f:elt->elt'),
MapsTo x e m -> MapsTo x (f e) (map f m).
Proof. intros elt elt' m; exact (@Raw.map_1 elt elt' m.(this)). Qed.
- Lemma map_2 : forall (elt elt':Set)(m: t elt)(x:key)(f:elt->elt'),
+ Lemma map_2 : forall (elt elt':Type)(m: t elt)(x:key)(f:elt->elt'),
In x (map f m) -> In x m.
Proof. intros elt elt' m; exact (@Raw.map_2 elt elt' m.(this)). Qed.
- Lemma mapi_1 : forall (elt elt':Set)(m: t elt)(x:key)(e:elt)
+ Lemma mapi_1 : forall (elt elt':Type)(m: t elt)(x:key)(e:elt)
(f:key->elt->elt'), MapsTo x e m ->
exists y, E.eq y x /\ MapsTo x (f y e) (mapi f m).
Proof. intros elt elt' m; exact (@Raw.mapi_1 elt elt' m.(this)). Qed.
- Lemma mapi_2 : forall (elt elt':Set)(m: t elt)(x:key)
+ Lemma mapi_2 : forall (elt elt':Type)(m: t elt)(x:key)
(f:key->elt->elt'), In x (mapi f m) -> In x m.
Proof. intros elt elt' m; exact (@Raw.mapi_2 elt elt' m.(this)). Qed.
- Lemma map2_1 : forall (elt elt' elt'':Set)(m: t elt)(m': t elt')
+ Lemma map2_1 : forall (elt elt' elt'':Type)(m: t elt)(m': t elt')
(x:key)(f:option elt->option elt'->option elt''),
In x m \/ In x m' ->
find x (map2 f m m') = f (find x m) (find x m').
@@ -988,7 +985,7 @@ Section Elt.
intros elt elt' elt'' m m' x f;
exact (@Raw.map2_1 elt elt' elt'' f m.(this) m.(NoDup) m'.(this) m'.(NoDup) x).
Qed.
- Lemma map2_2 : forall (elt elt' elt'':Set)(m: t elt)(m': t elt')
+ Lemma map2_2 : forall (elt elt' elt'':Type)(m: t elt)(m': t elt')
(x:key)(f:option elt->option elt'->option elt''),
In x (map2 f m m') -> In x m \/ In x m'.
Proof.
diff --git a/theories/FSets/FMaps.v b/theories/FSets/FMaps.v
index 72ccad3f..75904202 100644
--- a/theories/FSets/FMaps.v
+++ b/theories/FSets/FMaps.v
@@ -6,13 +6,13 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: FMaps.v 8844 2006-05-22 17:22:36Z letouzey $ *)
+(* $Id: FMaps.v 10699 2008-03-19 20:56:43Z letouzey $ *)
-Require Export OrderedType.
-Require Export OrderedTypeEx.
-Require Export OrderedTypeAlt.
+
+Require Export OrderedType OrderedTypeEx OrderedTypeAlt.
+Require Export DecidableType DecidableTypeEx.
Require Export FMapInterface.
-Require Export FMapList.
Require Export FMapPositive.
-Require Export FMapIntMap.
-Require Export FMapFacts. \ No newline at end of file
+Require Export FMapFacts.
+Require Export FMapWeakList.
+Require Export FMapList.
diff --git a/theories/FSets/FSetAVL.v b/theories/FSets/FSetAVL.v
index d5ce54d9..faa705f6 100644
--- a/theories/FSets/FSetAVL.v
+++ b/theories/FSets/FSetAVL.v
@@ -1,4 +1,3 @@
-
(***********************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *)
@@ -12,41 +11,555 @@
* Institution: LRI, CNRS UMR 8623 - Université Paris Sud
* 91405 Orsay, France *)
-(* $Id: FSetAVL.v 9862 2007-05-25 16:57:06Z letouzey $ *)
+(* $Id: FSetAVL.v 10811 2008-04-17 16:29:49Z letouzey $ *)
+
+(** * FSetAVL *)
(** This module implements sets using AVL trees.
- It follows the implementation from Ocaml's standard library. *)
+ 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 FSet interface. As a consequence,
+ balancing results are not part of this file anymore, they can
+ now be found in [FSetFullAVL].
+
+ 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
+ [FSetFullAVL]. The structural variants compute faster in Coq,
+ whereas the other variants produce nicer and/or (slightly) faster
+ code after extraction.
+*)
-Require Import FSetInterface.
-Require Import FSetList.
-Require Import ZArith.
-Require Import Int.
+Require Import FSetInterface FSetList ZArith Int.
-Set Firstorder Depth 3.
+Set Implicit Arguments.
+Unset Strict Implicit.
-Module Raw (I:Int)(X:OrderedType).
-Import I.
-Module II:=MoreInt(I).
-Import II.
-Open Local Scope Int_scope.
+(** Notations and helper lemma about pairs *)
+
+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.
-Module E := X.
-Module MX := OrderedTypeFacts X.
+(** * Raw
+
+ Functor of pure functions + a posteriori proofs of invariant
+ preservation *)
+
+Module Raw (Import I:Int)(X:OrderedType).
+Open Local Scope pair_scope.
+Open Local Scope lazy_bool_scope.
+Open Local Scope Int_scope.
Definition elt := X.t.
-(** * Trees *)
+(** * Trees
-Inductive tree : Set :=
+ The fourth field of [Node] is the height of the tree *)
+
+Inductive tree :=
| Leaf : tree
| Node : tree -> X.t -> tree -> int -> tree.
Notation t := tree.
-(** The fourth field of [Node] is the height of the tree *)
+(** * Basic functions on trees: height and cardinal *)
+
+Definition height (s : tree) : int :=
+ match s with
+ | Leaf => 0
+ | Node _ _ _ h => h
+ end.
+
+Fixpoint cardinal (s : tree) : 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).
+Notation "t #l" := (t_left t) (at level 9, format "t '#l'").
+Notation "t #b" := (t_in t) (at level 9, format "t '#b'").
+Notation "t #r" := (t_right t) (at level 9, format "t '#r'").
+
+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 [FSetFullAVL].
+*)
+
+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) (t : tree) : list X.t :=
+ match t 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 : tree) : 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
+ [FSetFullAVL].
+ *)
+
+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 [FSetFullAVL]. 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 -> tree -> 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.
+
+
+
+
+(** * Invariants *)
+
+(** ** Occurrence in a tree *)
+
+Inductive In (x : elt) : tree -> Prop :=
+ | IsRoot : forall l r h y, X.eq x y -> In x (Node l y r h)
+ | InLeft : forall l r h y, In x l -> In x (Node l y r h)
+ | InRight : forall l r h y, In x r -> In x (Node l y r h).
+
+(** ** 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, In y s -> X.lt y x.
+Definition gt_tree x s := forall y, In 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).
+
+
+
+
+(** * Some shortcuts *)
+
+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.
+
+
+
+(** * Correctness proofs, isolated in a sub-module *)
+
+Module Proofs.
+ Module MX := OrderedTypeFacts X.
+ Module L := FSetList.Raw X.
+
+(** * Automation and dedicated tactics *)
+
+Hint Constructors In bst.
+Hint Unfold lt_tree gt_tree.
+
+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.
(** A tactic to repeat [inversion_clear] on all hyps of the
form [(f (Node _ _ _ _))] *)
+
Ltac inv f :=
match goal with
| H:f Leaf |- _ => inversion_clear H; inv f
@@ -56,30 +569,18 @@ Ltac inv f :=
| _ => idtac
end.
-(** Same, but with a backup of the original hypothesis. *)
+Ltac intuition_in := repeat progress (intuition; inv In).
-Ltac safe_inv f := match goal with
- | H:f (Node _ _ _ _) |- _ =>
- generalize H; inversion_clear H; safe_inv f
- | _ => intros
- end.
+(** Helper tactic concerning order of elements. *)
-(** * Occurrence in a tree *)
+Ltac order := match goal with
+ | U: lt_tree _ ?s, V: In _ ?s |- _ => generalize (U _ V); clear U; order
+ | U: gt_tree _ ?s, V: In _ ?s |- _ => generalize (U _ V); clear U; order
+ | _ => MX.order
+end.
-Inductive In (x : elt) : tree -> Prop :=
- | IsRoot :
- forall (l r : tree) (h : int) (y : elt),
- X.eq x y -> In x (Node l y r h)
- | InLeft :
- forall (l r : tree) (h : int) (y : elt),
- In x l -> In x (Node l y r h)
- | InRight :
- forall (l r : tree) (h : int) (y : elt),
- In x r -> In x (Node l y r h).
-
-Hint Constructors In.
-Ltac intuition_in := repeat progress (intuition; inv In).
+(** * Basic results about [In], [lt_tree], [gt_tree], [height] *)
(** [In] is compatible with [X.eq] *)
@@ -90,48 +591,37 @@ Proof.
Qed.
Hint Immediate In_1.
-(** * Binary search trees *)
-
-(** [lt_tree x s]: all elements in [s] are smaller than [x]
- (resp. greater for [gt_tree]) *)
-
-Definition lt_tree (x : elt) (s : tree) :=
- forall y:elt, In y s -> X.lt y x.
-Definition gt_tree (x : elt) (s : tree) :=
- forall y:elt, In y s -> X.lt x y.
-
-Hint Unfold lt_tree gt_tree.
-
-Ltac order := match goal with
- | H: lt_tree ?x ?s, H1: In ?y ?s |- _ => generalize (H _ H1); clear H; order
- | H: gt_tree ?x ?s, H1: In ?y ?s |- _ => generalize (H _ H1); clear H; order
- | _ => MX.order
-end.
+Lemma In_node_iff :
+ forall l x r h y,
+ In y (Node l x r h) <-> In y l \/ X.eq y x \/ In y r.
+Proof.
+ intuition_in.
+Qed.
(** Results about [lt_tree] and [gt_tree] *)
Lemma lt_leaf : forall x : elt, lt_tree x Leaf.
Proof.
- unfold lt_tree in |- *; intros; inversion H.
+ red; inversion 1.
Qed.
Lemma gt_leaf : forall x : elt, gt_tree x Leaf.
Proof.
- unfold gt_tree in |- *; intros; inversion H.
+ 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 in *; intuition_in; order.
+ 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 in *; intuition_in; order.
+ unfold gt_tree; intuition_in; order.
Qed.
Hint Resolve lt_leaf gt_leaf lt_tree_node gt_tree_node.
@@ -145,7 +635,7 @@ Qed.
Lemma lt_tree_trans :
forall x y, X.lt x y -> forall t, lt_tree x t -> lt_tree y t.
Proof.
- firstorder eauto.
+ eauto.
Qed.
Lemma gt_tree_not_in :
@@ -157,120 +647,43 @@ Qed.
Lemma gt_tree_trans :
forall x y, X.lt y x -> forall t, gt_tree x t -> gt_tree y t.
Proof.
- firstorder eauto.
+ eauto.
Qed.
Hint Resolve lt_tree_not_in lt_tree_trans gt_tree_not_in gt_tree_trans.
-(** [bst t] : [t] is a binary search tree *)
-
-Inductive bst : tree -> Prop :=
- | BSLeaf : bst Leaf
- | BSNode :
- forall (x : elt) (l r : tree) (h : int),
- bst l -> bst r -> lt_tree x l -> gt_tree x r -> bst (Node l x r h).
+(** * Inductions principles *)
-Hint Constructors bst.
+Functional Scheme mem_ind := Induction for mem Sort Prop.
+Functional Scheme bal_ind := Induction for bal Sort Prop.
+Functional Scheme add_ind := Induction for add Sort Prop.
+Functional Scheme remove_min_ind := Induction for remove_min Sort Prop.
+Functional Scheme merge_ind := Induction for merge Sort Prop.
+Functional Scheme remove_ind := Induction for remove 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 split_ind := Induction for split 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.
-(** * AVL trees *)
-(** [avl s] : [s] is a properly balanced AVL tree,
- i.e. for any node the heights of the two children
- differ by at most 2 *)
-
-Definition height (s : tree) : int :=
- match s with
- | Leaf => 0
- | Node _ _ _ h => h
- end.
-
-Inductive avl : tree -> Prop :=
- | RBLeaf : avl Leaf
- | RBNode :
- forall (x : elt) (l r : tree) (h : int),
- avl l ->
- avl r ->
- -(2) <= height l - height r <= 2 ->
- h = max (height l) (height r) + 1 ->
- avl (Node l x r h).
-
-Hint Constructors avl.
-
-(** Results about [avl] *)
-
-Lemma avl_node :
- forall (x : elt) (l r : tree),
- avl l ->
- avl r ->
- -(2) <= height l - height r <= 2 ->
- avl (Node l x r (max (height l) (height r) + 1)).
-Proof.
- intros; auto.
-Qed.
-Hint Resolve avl_node.
-
-(** The tactics *)
+(** * Empty set *)
-Lemma height_non_negative : forall s : tree, avl s -> height s >= 0.
+Lemma empty_1 : Empty empty.
Proof.
- induction s; simpl; intros; auto with zarith.
- inv avl; intuition; omega_max.
+ intro; intro.
+ inversion H.
Qed.
-Implicit Arguments height_non_negative.
-
-(** When [H:avl r], typing [avl_nn H] or [avl_nn r] adds [height r>=0] *)
-
-Ltac avl_nn_hyp H :=
- let nz := fresh "nz" in assert (nz := height_non_negative H).
-
-Ltac avl_nn h :=
- let t := type of h in
- match type of t with
- | Prop => avl_nn_hyp h
- | _ => match goal with H : avl h |- _ => avl_nn_hyp H end
- end.
-
-(* Repeat the previous tactic.
- Drawback: need to clear the [avl _] hyps ... Thank you Ltac *)
-
-Ltac avl_nns :=
- match goal with
- | H:avl _ |- _ => avl_nn_hyp H; clear H; avl_nns
- | _ => idtac
- end.
-
-(** * Some shortcuts. *)
-
-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.
-
-(** * Empty set *)
-
-Definition empty := Leaf.
Lemma empty_bst : bst empty.
Proof.
auto.
Qed.
-Lemma empty_avl : avl empty.
-Proof.
- auto.
-Qed.
-
-Lemma empty_1 : Empty empty.
-Proof.
- intro; intro.
- inversion H.
-Qed.
-
(** * Emptyness test *)
-Definition is_empty (s:t) := match s with Leaf => true | _ => false end.
-
Lemma is_empty_1 : forall s, Empty s -> is_empty s = true.
Proof.
destruct s as [|r x l h]; simpl; auto.
@@ -282,54 +695,28 @@ Proof.
destruct s; simpl; intros; try discriminate; red; auto.
Qed.
-(** * Appartness *)
-(** The [mem] function is deciding appartness. It exploits the [bst] property
- to achieve logarithmic complexity. *)
-Function mem (x:elt)(s:t) { struct s } : bool :=
- 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.
+(** * Appartness *)
Lemma mem_1 : forall s x, bst s -> In x s -> mem x s = true.
-Proof.
- intros s x.
- functional induction (mem x s); inversion_clear 1; auto.
- inversion_clear 1.
- inversion_clear 1; auto; absurd (X.lt x y); eauto.
- inversion_clear 1; auto; absurd (X.lt y x); eauto.
+Proof.
+ intros s x; functional induction mem x s; auto; intros; try clear e0;
+ inv bst; intuition_in; order.
Qed.
Lemma mem_2 : forall s x, mem x s = true -> In x s.
Proof.
- intros s x.
- functional induction (mem x s); auto; intros; try discriminate.
+ intros s x; functional induction mem x s; auto; intros; discriminate.
Qed.
-(** * Singleton set *)
-Definition singleton (x : elt) := Node Leaf x Leaf 1.
-
-Lemma singleton_bst : forall x : elt, bst (singleton x).
-Proof.
- unfold singleton; auto.
-Qed.
-Lemma singleton_avl : forall x : elt, avl (singleton x).
-Proof.
- unfold singleton; intro.
- constructor; auto; try red; simpl; omega_max.
-Qed.
+(** * Singleton set *)
Lemma singleton_1 : forall x y, In y (singleton x) -> X.eq x y.
Proof.
- unfold singleton; inversion_clear 1; auto; inversion_clear H0.
+ unfold singleton; intros; inv In; order.
Qed.
Lemma singleton_2 : forall x y, X.eq x y -> In y (singleton x).
@@ -337,35 +724,14 @@ Proof.
unfold singleton; auto.
Qed.
-(** * 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).
-
-Lemma create_bst :
- forall l x r, bst l -> bst r -> lt_tree x l -> gt_tree x r ->
- bst (create l x r).
+Lemma singleton_bst : forall x : elt, bst (singleton x).
Proof.
- unfold create; auto.
+ unfold singleton; auto.
Qed.
-Hint Resolve create_bst.
-Lemma create_avl :
- forall l x r, avl l -> avl r -> -(2) <= height l - height r <= 2 ->
- avl (create l x r).
-Proof.
- unfold create; auto.
-Qed.
-Lemma create_height :
- forall l x r, avl l -> avl r -> -(2) <= height l - height r <= 2 ->
- height (create l x r) = max (height l) (height r) + 1.
-Proof.
- unfold create; intros; auto.
-Qed.
+
+(** * Helper functions *)
Lemma create_in :
forall l x r y, In y (create l x r) <-> X.eq y x \/ In y l \/ In y r.
@@ -373,196 +739,69 @@ Proof.
unfold create; split; [ inversion_clear 1 | ]; intuition.
Qed.
-(** trick for emulating [assert false] in Coq *)
-
-Definition assert_false := Leaf.
-
-(** [bal l x r] acts as [create], but performs one step of
- rebalancing if necessary, i.e. assumes [|height l - height r| <= 3]. *)
-
-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
- | 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
- | 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
- | 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
- | Node rll rlx rlr _ =>
- create (create l x rll) rlx (create rlr rx rr)
- end
- end
- else
- create l x r.
-
-Ltac bal_tac :=
- intros l x r;
- unfold bal;
- destruct (gt_le_dec (height l) (height r + 2));
- [ destruct l as [ |ll lx lr lh];
- [ | destruct (ge_lt_dec (height ll) (height lr));
- [ | destruct lr ] ]
- | destruct (gt_le_dec (height r) (height l + 2));
- [ destruct r as [ |rl rx rr rh];
- [ | destruct (ge_lt_dec (height rr) (height rl));
- [ | destruct rl ] ]
- | ] ]; intros.
-
-Lemma bal_bst : forall l x r, bst l -> bst r ->
- lt_tree x l -> gt_tree x r -> bst (bal l x r).
-Proof.
- (* intros l x r; functional induction bal l x r. MARCHE PAS !*)
- bal_tac;
- inv bst; repeat apply create_bst; auto; unfold create;
- apply lt_tree_node || apply gt_tree_node; auto;
- eapply lt_tree_trans || eapply gt_tree_trans || eauto; eauto.
-Qed.
-
-Lemma bal_avl : forall l x r, avl l -> avl r ->
- -(3) <= height l - height r <= 3 -> avl (bal l x r).
+Lemma create_bst :
+ forall l x r, bst l -> bst r -> lt_tree x l -> gt_tree x r ->
+ bst (create l x r).
Proof.
- bal_tac; inv avl; repeat apply create_avl; simpl in *; auto; omega_max.
+ unfold create; auto.
Qed.
+Hint Resolve create_bst.
-Lemma bal_height_1 : forall l x r, avl l -> avl r ->
- -(3) <= height l - height r <= 3 ->
- 0 <= height (bal l x r) - max (height l) (height r) <= 1.
+Lemma bal_in : forall l x r y,
+ In y (bal l x r) <-> X.eq y x \/ In y l \/ In y r.
Proof.
- bal_tac; inv avl; avl_nns; simpl in *; omega_max.
+ intros l x r; functional induction bal l x r; intros; try clear e0;
+ rewrite !create_in; intuition_in.
Qed.
-Lemma bal_height_2 :
- forall l x r, avl l -> avl r -> -(2) <= height l - height r <= 2 ->
- height (bal l x r) == max (height l) (height r) +1.
+Lemma bal_bst : forall l x r, bst l -> bst r ->
+ lt_tree x l -> gt_tree x r -> bst (bal l x r).
Proof.
- bal_tac; inv avl; simpl in *; omega_max.
+ intros l x r; functional induction bal l x r; intros;
+ inv bst; repeat apply create_bst; auto; unfold create;
+ (apply lt_tree_node || apply gt_tree_node); auto;
+ (eapply lt_tree_trans || eapply gt_tree_trans); eauto.
Qed.
+Hint Resolve bal_bst.
-Lemma bal_in : forall l x r y, avl l -> avl r ->
- (In y (bal l x r) <-> X.eq y x \/ In y l \/ In y r).
-Proof.
- bal_tac;
- solve [repeat rewrite create_in; intuition_in
- |inv avl; avl_nns; simpl in *; false_omega].
-Qed.
-Ltac omega_bal := match goal with
- | H:avl ?l, H':avl ?r |- context [ bal ?l ?x ?r ] =>
- generalize (bal_height_1 l x r H H') (bal_height_2 l x r H H');
- omega_max
- end.
(** * Insertion *)
-Function add (x:elt)(s:t) { struct s } : t := 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.
-
-Lemma add_avl_1 : forall s x, avl s ->
- avl (add x s) /\ 0 <= height (add x s) - height s <= 1.
-Proof.
- intros s x; functional induction (add x s); subst;intros; inv avl; simpl in *.
- intuition; try constructor; simpl; auto; try omega_max.
- (* LT *)
- destruct IHt; auto.
- split.
- apply bal_avl; auto; omega_max.
- omega_bal.
- (* EQ *)
- intuition; omega_max.
- (* GT *)
- destruct IHt; auto.
- split.
- apply bal_avl; auto; omega_max.
- omega_bal.
-Qed.
-
-Lemma add_avl : forall s x, avl s -> avl (add x s).
-Proof.
- intros; generalize (add_avl_1 s x H); intuition.
-Qed.
-Hint Resolve add_avl.
-
-Lemma add_in : forall s x y, avl s ->
- (In y (add x s) <-> X.eq y x \/ In y s).
+Lemma add_in : forall s x y,
+ In y (add x s) <-> X.eq y x \/ In y s.
Proof.
- intros s x; functional induction (add x s); auto; intros.
- intuition_in.
- (* LT *)
- inv avl.
- rewrite bal_in; auto.
- rewrite (IHt y0 H0); intuition_in.
- (* EQ *)
- inv avl.
- intuition.
+ intros s x; functional induction (add x s); auto; intros;
+ try rewrite bal_in, IHt; intuition_in.
eapply In_1; eauto.
- (* GT *)
- inv avl.
- rewrite bal_in; auto.
- rewrite (IHt y0 H1); intuition_in.
Qed.
-Lemma add_bst : forall s x, bst s -> avl s -> bst (add x s).
+Lemma add_bst : forall s x, bst s -> bst (add x s).
Proof.
- intros s x; functional induction (add x s); auto; intros.
- inv bst; inv avl; apply bal_bst; auto.
+ intros s x; functional induction (add x s); auto; intros;
+ inv bst; apply bal_bst; auto.
(* lt_tree -> lt_tree (add ...) *)
- red; red in H4.
+ red; red in H3.
intros.
- rewrite (add_in l x y0 H) in H0.
+ rewrite add_in in H.
intuition.
eauto.
- inv bst; inv avl; apply bal_bst; auto.
+ inv bst; auto using bal_bst.
(* gt_tree -> gt_tree (add ...) *)
- red; red in H4.
+ red; red in H3.
intros.
- rewrite (add_in r x y0 H5) in H0.
+ rewrite add_in in H.
intuition.
apply MX.lt_eq with x; auto.
Qed.
+Hint Resolve add_bst.
-(** * Join
- Same as [bal] but does not assume anything regarding heights
- of [l] and [r].
-*)
-Fixpoint join (l:t) : 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.
+(** * 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];
@@ -579,437 +818,200 @@ Ltac join_tac :=
end
| ] ] ] ]; intros.
-Lemma join_avl_1 : forall l x r, avl l -> avl r -> avl (join l x r) /\
- 0<= height (join l x r) - max (height l) (height r) <= 1.
-Proof.
- (* intros l x r; functional induction join l x r. AUTRE PROBLEME! *)
- join_tac.
-
- split; simpl; auto.
- destruct (add_avl_1 r x H0).
- avl_nns; omega_max.
- split; auto.
- set (l:=Node ll lx lr lh) in *.
- destruct (add_avl_1 l x H).
- simpl (height Leaf).
- avl_nns; omega_max.
-
- inversion_clear H.
- assert (height (Node rl rx rr rh) = rh); auto.
- set (r := Node rl rx rr rh) in *; clearbody r.
- destruct (Hlr x r H2 H0); clear Hrl Hlr.
- set (j := join lr x r) in *; clearbody j.
- simpl.
- assert (-(3) <= height ll - height j <= 3) by omega_max.
- split.
- apply bal_avl; auto.
- omega_bal.
-
- inversion_clear H0.
- assert (height (Node ll lx lr lh) = lh); auto.
- set (l := Node ll lx lr lh) in *; clearbody l.
- destruct (Hrl H H1); clear Hrl Hlr.
- set (j := join l x rl) in *; clearbody j.
- simpl.
- assert (-(3) <= height j - height rr <= 3) by omega_max.
- split.
- apply bal_avl; auto.
- omega_bal.
-
- clear Hrl Hlr.
- assert (height (Node ll lx lr lh) = lh); auto.
- assert (height (Node rl rx rr rh) = rh); auto.
- set (l := Node ll lx lr lh) in *; clearbody l.
- set (r := Node rl rx rr rh) in *; clearbody r.
- assert (-(2) <= height l - height r <= 2) by omega_max.
- split.
- apply create_avl; auto.
- rewrite create_height; auto; omega_max.
-Qed.
-
-Lemma join_avl : forall l x r, avl l -> avl r -> avl (join l x r).
-Proof.
- intros; generalize (join_avl_1 l x r H H0); intuition.
-Qed.
-Hint Resolve join_avl.
-
-Lemma join_in : forall l x r y, avl l -> avl r ->
- (In y (join l x r) <-> X.eq y x \/ In y l \/ In y r).
+Lemma join_in : forall l x r y,
+ In y (join l x r) <-> X.eq y x \/ In y l \/ In y r.
Proof.
join_tac.
simpl.
rewrite add_in; intuition_in.
-
rewrite add_in; intuition_in.
-
- inv avl.
- rewrite bal_in; auto.
- rewrite Hlr; clear Hlr Hrl; intuition_in.
-
- inv avl.
- rewrite bal_in; auto.
- rewrite Hrl; clear Hlr Hrl; intuition_in.
-
+ rewrite bal_in, Hlr; clear Hlr Hrl; intuition_in.
+ rewrite bal_in, Hrl; clear Hlr Hrl; intuition_in.
apply create_in.
Qed.
-Lemma join_bst : forall l x r, bst l -> avl l -> bst r -> avl r ->
+Lemma join_bst : forall l x r, bst l -> bst r ->
lt_tree x l -> gt_tree x r -> bst (join l x r).
Proof.
- join_tac.
- apply add_bst; auto.
- apply add_bst; auto.
-
- inv bst; safe_inv avl.
- apply bal_bst; auto.
- clear Hrl Hlr H13 H14 H16 H17 z; intro; intros.
- set (r:=Node rl rx rr rh) in *; clearbody r.
- rewrite (join_in lr x r y) in H13; auto.
- intuition.
- apply MX.lt_eq with x; eauto.
- eauto.
-
- inv bst; safe_inv avl.
- apply bal_bst; auto.
- clear Hrl Hlr H13 H14 H16 H17 z; intro; intros.
- set (l:=Node ll lx lr lh) in *; clearbody l.
- rewrite (join_in l x rl y) in H13; auto.
- intuition.
- apply MX.eq_lt with x; eauto.
- eauto.
-
- apply create_bst; auto.
+ join_tac; auto; inv bst; apply bal_bst; auto;
+ clear Hrl Hlr z; intro; intros; rewrite join_in in *.
+ intuition; [ apply MX.lt_eq with x | ]; eauto.
+ intuition; [ apply MX.eq_lt with x | ]; eauto.
Qed.
+Hint Resolve join_bst.
-(** * 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]).
-*)
-
-Function remove_min (l:t)(x:elt)(r:t) { struct l } : t*elt :=
- match l with
- | Leaf => (r,x)
- | Node ll lx lr lh => let (l',m) := (remove_min ll lx lr : t*elt) in (bal l' x r, m)
- end.
-
-Lemma remove_min_avl_1 : forall l x r h, avl (Node l x r h) ->
- avl (fst (remove_min l x r)) /\
- 0 <= height (Node l x r h) - height (fst (remove_min l x r)) <= 1.
-Proof.
- intros l x r; functional induction (remove_min l x r); subst;simpl in *; intros.
- inv avl; simpl in *; split; auto.
- avl_nns; omega_max.
- (* l = Node *)
- inversion_clear H.
- rewrite e0 in IHp;simpl in IHp;destruct (IHp lh); auto.
- split; simpl in *.
- apply bal_avl; auto; omega_max.
- omega_bal.
-Qed.
-Lemma remove_min_avl : forall l x r h, avl (Node l x r h) ->
- avl (fst (remove_min l x r)).
-Proof.
- intros; generalize (remove_min_avl_1 l x r h H); intuition.
-Qed.
+(** * Extraction of minimum element *)
-Lemma remove_min_in : forall l x r h y, avl (Node l x r h) ->
- (In y (Node l x r h) <->
- X.eq y (snd (remove_min l x r)) \/ In y (fst (remove_min l x r))).
+Lemma remove_min_in : forall l x r h y,
+ In y (Node l x r h) <->
+ X.eq y (remove_min l x r)#2 \/ In y (remove_min l x r)#1.
Proof.
intros l x r; functional induction (remove_min l x r); simpl in *; intros.
intuition_in.
- (* l = Node *)
- inversion_clear H.
- generalize (remove_min_avl ll lx lr lh H0).
- rewrite e0; simpl; intros.
- rewrite bal_in; auto.
- rewrite e0 in IHp;generalize (IHp lh y H0).
- intuition.
- inversion_clear H7; intuition.
+ rewrite bal_in, In_node_iff, IHp, e0; simpl; intuition.
Qed.
Lemma remove_min_bst : forall l x r h,
- bst (Node l x r h) -> avl (Node l x r h) -> bst (fst (remove_min l x r)).
+ bst (Node l x r h) -> bst (remove_min l x r)#1.
Proof.
- intros l x r; functional induction (remove_min l x r); subst;simpl in *; intros.
+ intros l x r; functional induction (remove_min l x r); simpl; intros.
inv bst; auto.
- inversion_clear H; inversion_clear H0.
- rewrite_all e0;simpl in *.
+ inversion_clear H.
+ specialize IHp with (1:=H0); rewrite e0 in IHp; auto.
apply bal_bst; auto.
- firstorder.
- intro; intros.
- generalize (remove_min_in ll lx lr lh y H).
- rewrite e0; simpl.
- destruct 1.
- apply H3; intuition.
+ intro y; specialize (H2 y).
+ rewrite remove_min_in, e0 in H2; simpl in H2; intuition.
Qed.
Lemma remove_min_gt_tree : forall l x r h,
- bst (Node l x r h) -> avl (Node l x r h) ->
- gt_tree (snd (remove_min l x r)) (fst (remove_min l x r)).
+ bst (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); subst;simpl in *; intros.
+ intros l x r; functional induction (remove_min l x r); simpl; intros.
inv bst; auto.
- inversion_clear H; inversion_clear H0.
- intro; intro.
- generalize (IHp lh H1 H); clear H6 H7 IHp.
- generalize (remove_min_avl ll lx lr lh H).
- generalize (remove_min_in ll lx lr lh m H).
- rewrite e0; simpl; intros.
- rewrite (bal_in l' x r y H7 H5) in H0.
- destruct H6.
- firstorder.
- apply MX.lt_eq with x; auto.
- apply X.lt_trans with x; auto.
+ inversion_clear H.
+ specialize IHp with (1:=H0); rewrite e0 in IHp; simpl in IHp.
+ intro y; rewrite bal_in; intuition;
+ specialize (H2 m); rewrite remove_min_in, e0 in H2; simpl in H2;
+ [ apply MX.lt_eq with x | ]; eauto.
Qed.
+Hint Resolve remove_min_bst remove_min_gt_tree.
-(** * 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].
-*)
-Function merge (s1 s2 :t) : t:= 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.
-Lemma merge_avl_1 : forall s1 s2, avl s1 -> avl s2 ->
- -(2) <= height s1 - height s2 <= 2 ->
- avl (merge s1 s2) /\
- 0<= height (merge s1 s2) - max (height s1) (height s2) <=1.
-Proof.
- intros s1 s2; functional induction (merge s1 s2); subst;simpl in *; intros.
- split; auto; avl_nns; omega_max.
- split; auto; avl_nns; simpl in *; omega_max.
- destruct s1;try contradiction;clear y.
- generalize (remove_min_avl_1 l2 x2 r2 h2 H0).
- rewrite e1; simpl; destruct 1.
- split.
- apply bal_avl; auto.
- simpl; omega_max.
- omega_bal.
-Qed.
-
-Lemma merge_avl : forall s1 s2, avl s1 -> avl s2 ->
- -(2) <= height s1 - height s2 <= 2 -> avl (merge s1 s2).
-Proof.
- intros; generalize (merge_avl_1 s1 s2 H H0 H1); intuition.
-Qed.
+(** * Merging two trees *)
-Lemma merge_in : forall s1 s2 y, bst s1 -> avl s1 -> bst s2 -> avl s2 ->
- (In y (merge s1 s2) <-> In y s1 \/ In y s2).
-Proof.
- intros s1 s2; functional induction (merge s1 s2); subst; simpl in *; intros.
+Lemma merge_in : forall s1 s2 y,
+ In y (merge s1 s2) <-> In y s1 \/ In y s2.
+Proof.
+ intros s1 s2; functional induction (merge s1 s2); intros;
+ try factornode _x _x0 _x1 _x2 as s1.
intuition_in.
intuition_in.
- destruct s1;try contradiction;clear y.
- replace s2' with (fst (remove_min l2 x2 r2)); [|rewrite e1; auto].
- rewrite bal_in; auto.
- generalize (remove_min_avl l2 x2 r2 h2); rewrite e1; simpl; auto.
- generalize (remove_min_in l2 x2 r2 h2 y0); rewrite e1; simpl; intro.
- rewrite H3 ; intuition.
+ rewrite bal_in, remove_min_in, e1; simpl; intuition.
Qed.
-Lemma merge_bst : forall s1 s2, bst s1 -> avl s1 -> bst s2 -> avl s2 ->
+Lemma merge_bst : forall s1 s2, bst s1 -> bst s2 ->
(forall y1 y2 : elt, In y1 s1 -> In y2 s2 -> X.lt y1 y2) ->
- bst (merge s1 s2).
+ bst (merge s1 s2).
Proof.
- intros s1 s2; functional induction (merge s1 s2); subst;simpl in *; intros; auto.
- destruct s1;try contradiction;clear y.
+ intros s1 s2; functional induction (merge s1 s2); intros; auto;
+ try factornode _x _x0 _x1 _x2 as s1.
apply bal_bst; auto.
- generalize (remove_min_bst l2 x2 r2 h2); rewrite e1; simpl in *; auto.
- intro; intro.
- apply H3; auto.
- generalize (remove_min_in l2 x2 r2 h2 m); rewrite e1; simpl; intuition.
- generalize (remove_min_gt_tree l2 x2 r2 h2); rewrite e1; simpl; auto.
-Qed.
+ change s2' with ((s2',m)#1); rewrite <-e1; eauto.
+ intros y Hy.
+ apply H1; auto.
+ rewrite remove_min_in, e1; simpl; auto.
+ change (gt_tree (s2',m)#2 (s2',m)#1); rewrite <-e1; eauto.
+Qed.
+Hint Resolve merge_bst.
-(** * Deletion *)
-Function remove (x:elt)(s:tree) { struct s } : t := 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.
-Lemma remove_avl_1 : forall s x, avl s ->
- avl (remove x s) /\ 0 <= height s - height (remove x s) <= 1.
-Proof.
- intros s x; functional induction (remove x s); subst;simpl; intros.
- intuition; omega_max.
- (* LT *)
- inv avl.
- destruct (IHt H0).
- split.
- apply bal_avl; auto.
- omega_max.
- omega_bal.
- (* EQ *)
- inv avl.
- generalize (merge_avl_1 l r H0 H1 H2).
- intuition omega_max.
- (* GT *)
- inv avl.
- destruct (IHt H1).
- split.
- apply bal_avl; auto.
- omega_max.
- omega_bal.
-Qed.
-
-Lemma remove_avl : forall s x, avl s -> avl (remove x s).
-Proof.
- intros; generalize (remove_avl_1 s x H); intuition.
-Qed.
-Hint Resolve remove_avl.
+(** * Deletion *)
-Lemma remove_in : forall s x y, bst s -> avl s ->
+Lemma remove_in : forall s x y, bst s ->
(In y (remove x s) <-> ~ X.eq y x /\ In y s).
Proof.
- intros s x; functional induction (remove x s); subst;simpl; intros.
+ intros s x; functional induction (remove x s); intros; inv bst.
intuition_in.
- (* LT *)
- inv avl; inv bst; clear e0.
- rewrite bal_in; auto.
- generalize (IHt y0 H0); intuition; [ order | order | intuition_in ].
- (* EQ *)
- inv avl; inv bst; clear e0.
- rewrite merge_in; intuition; [ order | order | intuition_in ].
- elim H9; eauto.
- (* GT *)
- inv avl; inv bst; clear e0.
- rewrite bal_in; auto.
- generalize (IHt y0 H5); intuition; [ order | order | intuition_in ].
+ rewrite bal_in, IHt; clear e0 IHt; intuition; [order|order|intuition_in].
+ rewrite merge_in; clear e0; intuition; [order|order|intuition_in].
+ elim H4; eauto.
+ rewrite bal_in, IHt; clear e0 IHt; intuition; [order|order|intuition_in].
Qed.
-Lemma remove_bst : forall s x, bst s -> avl s -> bst (remove x s).
+Lemma remove_bst : forall s x, bst s -> bst (remove x s).
Proof.
- intros s x; functional induction (remove x s); simpl; intros.
+ intros s x; functional induction (remove x s); intros; inv bst.
auto.
(* LT *)
- inv avl; inv bst.
apply bal_bst; auto.
- intro; intro.
- rewrite (remove_in l x y0) in H; auto.
- destruct H; eauto.
+ intro z; rewrite remove_in; auto; destruct 1; eauto.
(* EQ *)
- inv avl; inv bst.
- apply merge_bst; eauto.
+ eauto.
(* GT *)
- inv avl; inv bst.
apply bal_bst; auto.
- intro; intro.
- rewrite (remove_in r x y0) in H; auto.
- destruct H; eauto.
+ intro z; rewrite remove_in; auto; destruct 1; eauto.
Qed.
+Hint Resolve remove_bst.
- (** * Minimum element *)
-Function min_elt (s:t) : option elt := match s with
- | Leaf => None
- | Node Leaf y _ _ => Some y
- | Node l _ _ _ => min_elt l
-end.
+(** * Minimum element *)
Lemma min_elt_1 : forall s x, min_elt s = Some x -> In x s.
Proof.
- intro s; functional induction (min_elt s); subst; simpl.
- inversion 1.
- inversion 1; auto.
- intros.
- destruct l; auto.
+ intro s; functional induction (min_elt s); auto; inversion 1; auto.
Qed.
-Lemma min_elt_2 : forall s x y, bst s ->
+Lemma min_elt_2 : forall s x y, bst s ->
min_elt s = Some x -> In y s -> ~ X.lt y x.
Proof.
- intro s; functional induction (min_elt s); subst;simpl.
+ intro s; functional induction (min_elt s);
+ try rename _x1 into l1, _x2 into x1, _x3 into r1, _x4 into h1.
inversion_clear 2.
inversion_clear 1.
inversion 1; subst.
inversion_clear 1; auto.
inversion_clear H5.
- destruct l;try contradiction.
inversion_clear 1.
simpl.
destruct l1.
inversion 1; subst.
- assert (X.lt x _x) by (apply H2; auto).
+ assert (X.lt x y) by (apply H2; auto).
inversion_clear 1; auto; order.
- assert (X.lt t _x) by auto.
+ assert (X.lt x1 y) by auto.
inversion_clear 2; auto;
- (assert (~ X.lt t x) by auto); order.
+ (assert (~ X.lt x1 x) by auto); order.
Qed.
Lemma min_elt_3 : forall s, min_elt s = None -> Empty s.
Proof.
- intro s; functional induction (min_elt s); subst;simpl.
- red; auto.
+ intro s; functional induction (min_elt s).
+ red; red; inversion 2.
inversion 1.
- destruct l;try contradiction.
- clear y;intro H0.
- destruct (IHo H0 t); auto.
+ intro H0.
+ destruct (IHo H0 _x2); auto.
Qed.
-(** * Maximum element *)
-Function max_elt (s:t) : option elt := match s with
- | Leaf => None
- | Node _ y Leaf _ => Some y
- | Node _ _ r _ => max_elt r
-end.
+(** * Maximum element *)
Lemma max_elt_1 : forall s x, max_elt s = Some x -> In x s.
Proof.
- intro s; functional induction (max_elt s); subst;simpl.
- inversion 1.
- inversion 1; auto.
- destruct r;try contradiction; auto.
+ intro s; functional induction (max_elt s); auto; inversion 1; auto.
Qed.
Lemma max_elt_2 : forall s x y, bst s ->
max_elt s = Some x -> In y s -> ~ X.lt x y.
Proof.
- intro s; functional induction (max_elt s); subst;simpl.
+ intro s; functional induction (max_elt s);
+ try rename _x1 into l1, _x2 into x1, _x3 into r1, _x4 into h1.
inversion_clear 2.
inversion_clear 1.
inversion 1; subst.
inversion_clear 1; auto.
inversion_clear H5.
- destruct r;try contradiction.
inversion_clear 1.
-(* inversion 1; subst. *)
-(* assert (X.lt y x) by (apply H4; auto). *)
-(* inversion_clear 1; auto; order. *)
- assert (X.lt _x0 t) by auto.
+ assert (X.lt y x1) by auto.
inversion_clear 2; auto;
- (assert (~ X.lt x t) by auto); order.
+ (assert (~ X.lt x x1) by auto); order.
Qed.
Lemma max_elt_3 : forall s, max_elt s = None -> Empty s.
Proof.
- intro s; functional induction (max_elt s); subst;simpl.
+ intro s; functional induction (max_elt s).
red; auto.
inversion 1.
- destruct r;try contradiction.
- intros H0; destruct (IHo H0 t); auto.
+ intros H0; destruct (IHo H0 _x2); auto.
Qed.
-(** * Any element *)
-Definition choose := min_elt.
+
+(** * Any element *)
Lemma choose_1 : forall s x, choose s = Some x -> In x s.
Proof.
@@ -1021,353 +1023,215 @@ Proof.
exact min_elt_3.
Qed.
-(** * Concatenation
+Lemma choose_3 : forall s s', bst s -> bst s' ->
+ forall x x', choose s = Some x -> choose s' = Some x' ->
+ Equal s s' -> X.eq x x'.
+Proof.
+ unfold choose, Equal; intros s s' Hb Hb' x x' Hx Hx' H.
+ assert (~X.lt x x').
+ apply min_elt_2 with s'; auto.
+ rewrite <-H; auto using min_elt_1.
+ assert (~X.lt x' x).
+ apply min_elt_2 with s; auto.
+ rewrite H; auto using min_elt_1.
+ destruct (X.compare x x'); intuition.
+Qed.
- Same as [merge] but does not assume anything about heights.
-*)
-Function concat (s1 s2 : t) : t :=
- match s1, s2 with
- | Leaf, _ => s2
- | _, Leaf => s1
- | _, Node l2 x2 r2 h2 =>
- let (s2',m) := remove_min l2 x2 r2 in
- join s1 m s2'
- end.
+(** * Concatenation *)
-Lemma concat_avl : forall s1 s2, avl s1 -> avl s2 -> avl (concat s1 s2).
+Lemma concat_in : forall s1 s2 y,
+ In y (concat s1 s2) <-> In y s1 \/ In y s2.
Proof.
- intros s1 s2; functional induction (concat s1 s2); subst;auto.
- destruct s1;try contradiction;clear y.
- intros; apply join_avl; auto.
- generalize (remove_min_avl l2 x2 r2 h2 H0); rewrite e1; simpl; auto.
+ intros s1 s2; functional induction (concat s1 s2); intros;
+ try factornode _x _x0 _x1 _x2 as s1.
+ intuition_in.
+ intuition_in.
+ rewrite join_in, remove_min_in, e1; simpl; intuition.
Qed.
-Lemma concat_bst : forall s1 s2, bst s1 -> avl s1 -> bst s2 -> avl s2 ->
+Lemma concat_bst : forall s1 s2, bst s1 -> bst s2 ->
(forall y1 y2 : elt, In y1 s1 -> In y2 s2 -> X.lt y1 y2) ->
bst (concat s1 s2).
Proof.
- intros s1 s2; functional induction (concat s1 s2); subst ;auto.
- destruct s1;try contradiction;clear y.
- intros; apply join_bst; auto.
- generalize (remove_min_bst l2 x2 r2 h2 H1 H2); rewrite e1; simpl; auto.
- generalize (remove_min_avl l2 x2 r2 h2 H2); rewrite e1; simpl; auto.
- generalize (remove_min_in l2 x2 r2 h2 m H2); rewrite e1; simpl; auto.
- destruct 1; intuition.
- generalize (remove_min_gt_tree l2 x2 r2 h2 H1 H2); rewrite e1; simpl; auto.
-Qed.
-
-Lemma concat_in : forall s1 s2 y, bst s1 -> avl s1 -> bst s2 -> avl s2 ->
- (forall y1 y2 : elt, In y1 s1 -> In y2 s2 -> X.lt y1 y2) ->
- (In y (concat s1 s2) <-> In y s1 \/ In y s2).
-Proof.
- intros s1 s2; functional induction (concat s1 s2);subst;simpl.
- intuition.
- inversion_clear H5.
- destruct s1;try contradiction;clear y;intuition.
- inversion_clear H5.
- destruct s1;try contradiction;clear y; intros.
- rewrite (join_in (Node s1_1 t s1_2 i) m s2' y H0).
- generalize (remove_min_avl l2 x2 r2 h2 H2); rewrite e1; simpl; auto.
- generalize (remove_min_in l2 x2 r2 h2 y H2); rewrite e1; simpl.
- intro EQ; rewrite EQ; intuition.
+ intros s1 s2; functional induction (concat s1 s2); intros; auto;
+ try factornode _x _x0 _x1 _x2 as s1.
+ apply join_bst; auto.
+ change (bst (s2',m)#1); rewrite <-e1; eauto.
+ intros y Hy.
+ apply H1; auto.
+ rewrite remove_min_in, e1; simpl; auto.
+ change (gt_tree (s2',m)#2 (s2',m)#1); rewrite <-e1; eauto.
Qed.
+Hint Resolve concat_bst.
-(** * 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].
-*)
+(** * Splitting *)
-Function split (x:elt)(s:t) {struct s} : t * (bool * t) := match s with
- | Leaf => (Leaf, (false, Leaf))
- | Node l y r h =>
- match X.compare x y with
- | LT _ => match split x l with
- | (ll,(pres,rl)) => (ll, (pres, join rl y r))
- end
- | EQ _ => (l, (true, r))
- | GT _ => match split x r with
- | (rl,(pres,rr)) => (join l y rl, (pres, rr))
- end
- end
- end.
-
-Lemma split_avl : forall s x, avl s ->
- avl (fst (split x s)) /\ avl (snd (snd (split x s))).
-Proof.
- intros s x; functional induction (split x s);subst;simpl in *.
- auto.
- rewrite e1 in IHp;simpl in IHp;inversion_clear 1; intuition.
- simpl; inversion_clear 1; auto.
- rewrite e1 in IHp;simpl in IHp;inversion_clear 1; intuition.
-Qed.
-
-Lemma split_in_1 : forall s x y, bst s -> avl s ->
- (In y (fst (split x s)) <-> In y s /\ X.lt y x).
-Proof.
- intros s x; functional induction (split x s);subst;simpl in *.
- intuition; try inversion_clear H1.
- (* LT *)
- rewrite e1 in IHp;simpl in *; inversion_clear 1; inversion_clear 1; clear H7 H6.
- rewrite (IHp y0 H0 H4); clear IHp e0.
- intuition.
- inversion_clear H6; auto; order.
- (* EQ *)
- simpl in *; inversion_clear 1; inversion_clear 1; clear H6 H5 e0.
- intuition.
- order.
+Lemma split_in_1 : forall s x y, bst s ->
+ (In y (split x s)#l <-> In y s /\ X.lt y x).
+Proof.
+ intros s x; functional induction (split x s); simpl; intros;
+ inv bst; try clear e0.
+ intuition_in.
+ rewrite e1 in IHt; simpl in IHt; rewrite IHt; intuition_in; order.
intuition_in; order.
- (* GT *)
- rewrite e1 in IHp;simpl in *; inversion_clear 1; inversion_clear 1; clear H7 H6.
- rewrite join_in; auto.
- generalize (split_avl r x H5); rewrite e1; simpl; intuition.
- rewrite (IHp y0 H1 H5); clear e1.
- intuition; [ eauto | eauto | intuition_in ].
+ rewrite join_in.
+ rewrite e1 in IHt; simpl in IHt; rewrite IHt; intuition_in; order.
Qed.
-Lemma split_in_2 : forall s x y, bst s -> avl s ->
- (In y (snd (snd (split x s))) <-> In y s /\ X.lt x y).
+Lemma split_in_2 : forall s x y, bst s ->
+ (In y (split x s)#r <-> In y s /\ X.lt x y).
Proof.
- intros s x; functional induction (split x s);subst;simpl in *.
- intuition; try inversion_clear H1.
- (* LT *)
- rewrite e1 in IHp; simpl in *; inversion_clear 1; inversion_clear 1; clear H7 H6.
- rewrite join_in; auto.
- generalize (split_avl l x H4); rewrite e1; simpl; intuition.
- rewrite (IHp y0 H0 H4); clear IHp e0.
- intuition; [ order | order | intuition_in ].
- (* EQ *)
- simpl in *; inversion_clear 1; inversion_clear 1; clear H6 H5 e0.
- intuition; [ order | intuition_in; order ].
- (* GT *)
- rewrite e1 in IHp; simpl in *; inversion_clear 1; inversion_clear 1; clear H7 H6.
- rewrite (IHp y0 H1 H5); clear IHp e0.
- intuition; intuition_in; order.
+ intros s x; functional induction (split x s); subst; simpl; intros;
+ inv bst; try clear e0.
+ intuition_in.
+ rewrite join_in.
+ rewrite e1 in IHt; simpl in IHt; rewrite IHt; intuition_in; order.
+ intuition_in; order.
+ rewrite e1 in IHt; simpl in IHt; rewrite IHt; intuition_in; order.
Qed.
-Lemma split_in_3 : forall s x, bst s -> avl s ->
- (fst (snd (split x s)) = true <-> In x s).
+Lemma split_in_3 : forall s x, bst s ->
+ ((split x s)#b = true <-> In x s).
Proof.
- intros s x; functional induction (split x s);subst;simpl in *.
- intuition; try inversion_clear H1.
- (* LT *)
- rewrite e1 in IHp; simpl in *; inversion_clear 1; inversion_clear 1; clear H7 H6.
- rewrite IHp; auto.
- intuition_in; absurd (X.lt x y); eauto.
- (* EQ *)
- simpl in *; inversion_clear 1; inversion_clear 1; intuition.
- (* GT *)
- rewrite e1 in IHp; simpl in *; inversion_clear 1; inversion_clear 1; clear H7 H6.
- rewrite IHp; auto.
- intuition_in; absurd (X.lt y x); eauto.
+ intros s x; functional induction (split x s); subst; simpl; intros;
+ inv bst; try clear e0.
+ intuition_in; try discriminate.
+ rewrite e1 in IHt; simpl in IHt; rewrite IHt; intuition_in; order.
+ intuition.
+ rewrite e1 in IHt; simpl in IHt; rewrite IHt; intuition_in; order.
Qed.
-Lemma split_bst : forall s x, bst s -> avl s ->
- bst (fst (split x s)) /\ bst (snd (snd (split x s))).
+Lemma split_bst : forall s x, bst s ->
+ bst (split x s)#l /\ bst (split x s)#r.
Proof.
- intros s x; functional induction (split x s);subst;simpl in *.
- intuition.
- (* LT *)
- rewrite e1 in IHp; simpl in *; inversion_clear 1; inversion_clear 1.
- intuition.
- apply join_bst; auto.
- generalize (split_avl l x H4); rewrite e1; simpl; intuition.
- intro; intro.
- generalize (split_in_2 l x y0 H0 H4); rewrite e1; simpl; intuition.
- (* EQ *)
- simpl in *; inversion_clear 1; inversion_clear 1; intuition.
- (* GT *)
- rewrite e1 in IHp; simpl in *; inversion_clear 1; inversion_clear 1.
- intuition.
- apply join_bst; auto.
- generalize (split_avl r x H5); rewrite e1; simpl; intuition.
- intro; intro.
- generalize (split_in_1 r x y0 H1 H5); rewrite e1; simpl; intuition.
+ intros s x; functional induction (split x s); subst; simpl; intros;
+ inv bst; try clear e0; try rewrite e1 in *; simpl in *; intuition;
+ apply join_bst; auto.
+ intros y0.
+ generalize (split_in_2 x y0 H0); rewrite e1; simpl; intuition.
+ intros y0.
+ generalize (split_in_1 x y0 H1); rewrite e1; simpl; intuition.
Qed.
-(** * Intersection *)
-Fixpoint inter (s1 s2 : t) {struct s1} : t := match s1, s2 with
- | Leaf,_ => Leaf
- | _,Leaf => Leaf
- | Node l1 x1 r1 h1, _ =>
- match split x1 s2 with
- | (l2',(true,r2')) => join (inter l1 l2') x1 (inter r1 r2')
- | (l2',(false,r2')) => concat (inter l1 l2') (inter r1 r2')
- end
- end.
-Lemma inter_avl : forall s1 s2, avl s1 -> avl s2 -> avl (inter s1 s2).
-Proof.
- (* intros s1 s2; functional induction inter s1 s2; auto. BOF BOF *)
- induction s1 as [ | l1 Hl1 x1 r1 Hr1 h1]; simpl; auto.
- destruct s2 as [ | l2 x2 r2 h2]; intros; auto.
- generalize H0; inv avl.
- set (r:=Node l2 x2 r2 h2) in *; clearbody r; intros.
- destruct (split_avl r x1 H8).
- destruct (split x1 r) as [l2' (b,r2')]; simpl in *.
- destruct b; [ apply join_avl | apply concat_avl ]; auto.
-Qed.
+(** * Intersection *)
-Lemma inter_bst_in : forall s1 s2, bst s1 -> avl s1 -> bst s2 -> avl s2 ->
+Lemma inter_bst_in : forall s1 s2, bst s1 -> bst s2 ->
bst (inter s1 s2) /\ (forall y, In y (inter s1 s2) <-> In y s1 /\ In y s2).
-Proof.
- induction s1 as [ | l1 Hl1 x1 r1 Hr1 h1]; simpl; auto.
- intuition; inversion_clear H3.
- destruct s2 as [ | l2 x2 r2 h2]; intros.
- simpl; intuition; inversion_clear H3.
- generalize H1 H2; inv avl; inv bst.
- set (r:=Node l2 x2 r2 h2) in *; clearbody r; intros.
- destruct (split_avl r x1 H17).
- destruct (split_bst r x1 H16 H17).
- split.
- (* bst *)
- destruct (split x1 r) as [l2' (b,r2')]; simpl in *.
- destruct (Hl1 l2'); auto.
- destruct (Hr1 r2'); auto.
- destruct b.
+Proof.
+ intros s1 s2; functional induction inter s1 s2; intros B1 B2;
+ [intuition_in|intuition_in | | ];
+ factornode _x0 _x1 _x2 _x3 as s2;
+ generalize (split_bst x1 B2);
+ rewrite e1; simpl; destruct 1; inv bst;
+ destruct IHt as (IHb1,IHi1); auto;
+ destruct IHt0 as (IHb2,IHi2); auto;
+ generalize (@split_in_1 s2 x1)(@split_in_2 s2 x1)
+ (split_in_3 x1 B2)(split_bst x1 B2);
+ rewrite e1; simpl; split; intros.
(* bst join *)
- apply join_bst; try apply inter_avl; firstorder.
- (* bst concat *)
- apply concat_bst; try apply inter_avl; auto.
- intros; generalize (H22 y1) (H24 y2); intuition eauto.
- (* in *)
- intros.
- destruct (split_in_1 r x1 y H16 H17).
- destruct (split_in_2 r x1 y H16 H17).
- destruct (split_in_3 r x1 H16 H17).
- destruct (split x1 r) as [l2' (b,r2')]; simpl in *.
- destruct (Hl1 l2'); auto.
- destruct (Hr1 r2'); auto.
- destruct b.
- (* in join *)
- rewrite join_in; try apply inter_avl; auto.
- rewrite H30.
- rewrite H28.
- intuition_in.
+ apply join_bst; auto; intro y; [rewrite IHi1|rewrite IHi2]; intuition. (* In join *)
+ rewrite join_in, IHi1, IHi2, H5, H6; intuition_in.
apply In_1 with x1; auto.
- (* in concat *)
- rewrite concat_in; try apply inter_avl; auto.
- intros.
- intros; generalize (H28 y1) (H30 y2); intuition eauto.
- rewrite H30.
- rewrite H28.
+ (* bst concat *)
+ apply concat_bst; auto; intros y1 y2; rewrite IHi1, IHi2; intuition; order.
+ (* In concat *)
+ rewrite concat_in, IHi1, IHi2, H5, H6; auto.
+ assert (~In x1 s2) by (rewrite <- H7; auto).
intuition_in.
- generalize (H26 (In_1 _ _ _ H22 H35)); intro; discriminate.
+ elim H9.
+ apply In_1 with y; auto.
Qed.
-Lemma inter_bst : forall s1 s2, bst s1 -> avl s1 -> bst s2 -> avl s2 ->
- bst (inter s1 s2).
+Lemma inter_in : forall s1 s2 y, bst s1 -> bst s2 ->
+ (In y (inter s1 s2) <-> In y s1 /\ In y s2).
Proof.
- intros; generalize (inter_bst_in s1 s2); intuition.
+ intros s1 s2 y B1 B2; destruct (inter_bst_in B1 B2); auto.
Qed.
-Lemma inter_in : forall s1 s2 y, bst s1 -> avl s1 -> bst s2 -> avl s2 ->
- (In y (inter s1 s2) <-> In y s1 /\ In y s2).
+Lemma inter_bst : forall s1 s2, bst s1 -> bst s2 -> bst (inter s1 s2).
Proof.
- intros; generalize (inter_bst_in s1 s2); firstorder.
+ intros s1 s2 B1 B2; destruct (inter_bst_in B1 B2); auto.
Qed.
-(** * Difference *)
-
-Fixpoint diff (s1 s2 : t) { struct s1 } : t := match s1, s2 with
- | Leaf, _ => Leaf
- | _, Leaf => s1
- | Node l1 x1 r1 h1, _ =>
- match split x1 s2 with
- | (l2',(true,r2')) => concat (diff l1 l2') (diff r1 r2')
- | (l2',(false,r2')) => join (diff l1 l2') x1 (diff r1 r2')
- end
-end.
-Lemma diff_avl : forall s1 s2, avl s1 -> avl s2 -> avl (diff s1 s2).
-Proof.
- (* intros s1 s2; functional induction diff s1 s2; auto. BOF BOF *)
- induction s1 as [ | l1 Hl1 x1 r1 Hr1 h1]; simpl; auto.
- destruct s2 as [ | l2 x2 r2 h2]; intros; auto.
- generalize H0; inv avl.
- set (r:=Node l2 x2 r2 h2) in *; clearbody r; intros.
- destruct (split_avl r x1 H8).
- destruct (split x1 r) as [l2' (b,r2')]; simpl in *.
- destruct b; [ apply concat_avl | apply join_avl ]; auto.
-Qed.
+(** * Difference *)
-Lemma diff_bst_in : forall s1 s2, bst s1 -> avl s1 -> bst s2 -> avl s2 ->
+Lemma diff_bst_in : forall s1 s2, bst s1 -> bst s2 ->
bst (diff s1 s2) /\ (forall y, In y (diff s1 s2) <-> In y s1 /\ ~In y s2).
-Proof.
- induction s1 as [ | l1 Hl1 x1 r1 Hr1 h1]; simpl; auto.
- intuition; inversion_clear H3.
- destruct s2 as [ | l2 x2 r2 h2]; intros; auto.
- intuition; inversion_clear H4.
- generalize H1 H2; inv avl; inv bst.
- set (r:=Node l2 x2 r2 h2) in *; clearbody r; intros.
- destruct (split_avl r x1 H17).
- destruct (split_bst r x1 H16 H17).
- split.
- (* bst *)
- destruct (split x1 r) as [l2' (b,r2')]; simpl in *.
- destruct (Hl1 l2'); auto.
- destruct (Hr1 r2'); auto.
- destruct b.
+Proof.
+ intros s1 s2; functional induction diff s1 s2; intros B1 B2;
+ [intuition_in|intuition_in | | ];
+ factornode _x0 _x1 _x2 _x3 as s2;
+ generalize (split_bst x1 B2);
+ rewrite e1; simpl; destruct 1;
+ inv avl; inv bst;
+ destruct IHt as (IHb1,IHi1); auto;
+ destruct IHt0 as (IHb2,IHi2); auto;
+ generalize (@split_in_1 s2 x1)(@split_in_2 s2 x1)
+ (split_in_3 x1 B2)(split_bst x1 B2);
+ rewrite e1; simpl; split; intros.
(* bst concat *)
- apply concat_bst; try apply diff_avl; auto.
- intros; generalize (H22 y1) (H24 y2); intuition eauto.
+ apply concat_bst; auto; intros y1 y2; rewrite IHi1, IHi2; intuition; order.
+ (* In concat *)
+ rewrite concat_in, IHi1, IHi2, H5, H6; intuition_in.
+ elim H13.
+ apply In_1 with x1; auto.
(* bst join *)
- apply join_bst; try apply diff_avl; firstorder.
- (* in *)
- intros.
- destruct (split_in_1 r x1 y H16 H17).
- destruct (split_in_2 r x1 y H16 H17).
- destruct (split_in_3 r x1 H16 H17).
- destruct (split x1 r) as [l2' (b,r2')]; simpl in *.
- destruct (Hl1 l2'); auto.
- destruct (Hr1 r2'); auto.
- destruct b.
- (* in concat *)
- rewrite concat_in; try apply diff_avl; auto.
- intros.
- intros; generalize (H28 y1) (H30 y2); intuition eauto.
- rewrite H30.
- rewrite H28.
+ apply join_bst; auto; intro y; [rewrite IHi1|rewrite IHi2]; intuition. (* In join *)
+ rewrite join_in, IHi1, IHi2, H5, H6; auto.
+ assert (~In x1 s2) by (rewrite <- H7; auto).
intuition_in.
- elim H35; apply In_1 with x1; auto.
- (* in join *)
- rewrite join_in; try apply diff_avl; auto.
- rewrite H30.
- rewrite H28.
- intuition_in.
- generalize (H26 (In_1 _ _ _ H34 H24)); intro; discriminate.
+ elim H9.
+ apply In_1 with y; auto.
Qed.
-Lemma diff_bst : forall s1 s2, bst s1 -> avl s1 -> bst s2 -> avl s2 ->
- bst (diff s1 s2).
+Lemma diff_in : forall s1 s2 y, bst s1 -> bst s2 ->
+ (In y (diff s1 s2) <-> In y s1 /\ ~In y s2).
Proof.
- intros; generalize (diff_bst_in s1 s2); intuition.
+ intros s1 s2 y B1 B2; destruct (diff_bst_in B1 B2); auto.
Qed.
-Lemma diff_in : forall s1 s2 y, bst s1 -> avl s1 -> bst s2 -> avl s2 ->
- (In y (diff s1 s2) <-> In y s1 /\ ~In y s2).
+Lemma diff_bst : forall s1 s2, bst s1 -> bst s2 -> bst (diff s1 s2).
Proof.
- intros; generalize (diff_bst_in s1 s2); firstorder.
+ intros s1 s2 B1 B2; destruct (diff_bst_in B1 B2); auto.
Qed.
-(** * Elements *)
-(** [elements_tree_aux acc t] catenates the elements of [t] in infix
- order to the list [acc] *)
+(** * Union *)
-Fixpoint elements_aux (acc : list X.t) (t : tree) {struct t} : list X.t :=
- match t with
- | Leaf => acc
- | Node l x r _ => elements_aux (x :: elements_aux acc r) l
- end.
+Lemma union_in : forall s1 s2 y, bst s1 -> bst s2 ->
+ (In y (union s1 s2) <-> In y s1 \/ In 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.
+ generalize (split_in_1 x1 y B2)(split_in_2 x1 y B2)(split_bst x1 B2).
+ rewrite e1; simpl.
+ destruct 3; inv bst.
+ rewrite join_in, IHt, IHt0, H, H0; auto.
+ case (X.compare y x1); intuition_in.
+Qed.
-(** then [elements] is an instanciation with an empty [acc] *)
+Lemma union_bst : forall s1 s2, bst s1 -> bst s2 ->
+ bst (union s1 s2).
+Proof.
+ intros s1 s2; functional induction union s1 s2; intros B1 B2; auto.
+ factornode _x0 _x1 _x2 _x3 as s2.
+ generalize (@split_in_1 s2 x1)(@split_in_2 s2 x1)(split_bst x1 B2).
+ rewrite e1; simpl; destruct 3.
+ inv bst.
+ apply join_bst; auto.
+ intro y; rewrite union_in, H; intuition_in.
+ intro y; rewrite union_in, H0; intuition_in.
+Qed.
-Definition elements := elements_aux nil.
+
+(** * Elements *)
Lemma elements_aux_in : forall s acc x,
InA X.eq x (elements_aux acc s) <-> In x s \/ InA X.eq x acc.
@@ -1411,246 +1275,190 @@ Proof.
Qed.
Hint Resolve elements_sort.
-(** * Filter *)
-
-Section F.
-Variable f : elt -> bool.
+Lemma elements_nodup : forall s : tree, bst s -> NoDupA X.eq (elements s).
+Proof.
+ auto.
+Qed.
-Fixpoint filter_acc (acc:t)(s:t) { struct s } : t := match s with
- | Leaf => acc
- | Node l x r h =>
- filter_acc (filter_acc (if f x then add x acc else acc) l) r
- end.
+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.
-Definition filter := filter_acc Leaf.
+Lemma elements_cardinal : forall s : tree, cardinal s = length (elements s).
+Proof.
+ exact (fun s => elements_aux_cardinal s nil).
+Qed.
-Lemma filter_acc_avl : forall s acc, avl s -> avl acc ->
- avl (filter_acc acc s).
+Lemma elements_app :
+ forall s acc, elements_aux acc s = elements s ++ acc.
Proof.
- induction s; simpl; auto.
- intros.
- inv avl.
- apply IHs2; auto.
- apply IHs1; auto.
- destruct (f t); auto.
-Qed.
-Hint Resolve filter_acc_avl.
+ induction s; simpl; intros; auto.
+ rewrite IHs1, IHs2.
+ unfold elements; simpl.
+ rewrite 2 IHs1, IHs2, <- !app_nil_end, !app_ass; auto.
+Qed.
-Lemma filter_acc_bst : forall s acc, bst s -> avl s -> bst acc -> avl acc ->
- bst (filter_acc acc s).
+Lemma elements_node :
+ forall l x r h acc,
+ elements l ++ x :: elements r ++ acc =
+ elements (Node l x r h) ++ acc.
Proof.
- induction s; simpl; auto.
- intros.
- inv avl; inv bst.
- destruct (f t); auto.
- apply IHs2; auto.
- apply IHs1; auto.
- apply add_bst; auto.
-Qed.
+ unfold elements; simpl; intros; auto.
+ rewrite !elements_app, <- !app_nil_end, !app_ass; auto.
+Qed.
+
-Lemma filter_acc_in : forall s acc, avl s -> avl acc ->
+(** * Filter *)
+
+Section F.
+Variable f : elt -> bool.
+
+Lemma filter_acc_in : forall s acc,
compat_bool X.eq f -> forall x : elt,
- In x (filter_acc acc s) <-> In x acc \/ In x s /\ f x = true.
+ In x (filter_acc f acc s) <-> In x acc \/ In x s /\ f x = true.
Proof.
induction s; simpl; intros.
intuition_in.
- inv bst; inv avl.
- rewrite IHs2; auto.
- destruct (f t); auto.
- rewrite IHs1; auto.
- destruct (f t); auto.
+ rewrite IHs2, IHs1 by (destruct (f t); auto).
case_eq (f t); intros.
rewrite (add_in); auto.
intuition_in.
- rewrite (H1 _ _ H8).
+ rewrite (H _ _ H2).
intuition.
intuition_in.
- rewrite (H1 _ _ H8) in H9.
- rewrite H in H9; discriminate.
-Qed.
-
-Lemma filter_avl : forall s, avl s -> avl (filter s).
-Proof.
- unfold filter; intros; apply filter_acc_avl; auto.
+ rewrite (H _ _ H2) in H3.
+ rewrite H0 in H3; discriminate.
Qed.
-Lemma filter_bst : forall s, bst s -> avl s -> bst (filter s).
+Lemma filter_acc_bst : forall s acc, bst s -> bst acc ->
+ bst (filter_acc f acc s).
Proof.
- unfold filter; intros; apply filter_acc_bst; auto.
+ induction s; simpl; auto.
+ intros.
+ inv bst.
+ destruct (f t); auto.
Qed.
-Lemma filter_in : forall s, avl s ->
+Lemma filter_in : forall s,
compat_bool X.eq f -> forall x : elt,
- In x (filter s) <-> In x s /\ f x = true.
+ In x (filter f s) <-> In x s /\ f x = true.
Proof.
unfold filter; intros; rewrite filter_acc_in; intuition_in.
-Qed.
-
-(** * Partition *)
-
-Fixpoint partition_acc (acc : t*t)(s : t) { struct s } : t*t :=
- match s with
- | Leaf => acc
- | Node l x r _ =>
- let (acct,accf) := acc in
- partition_acc
- (partition_acc
- (if f x then (add x acct, accf) else (acct, add x accf)) l) r
- end.
-
-Definition partition := partition_acc (Leaf,Leaf).
+Qed.
-Lemma partition_acc_avl_1 : forall s acc, avl s ->
- avl (fst acc) -> avl (fst (partition_acc acc s)).
+Lemma filter_bst : forall s, bst s -> bst (filter f s).
Proof.
- induction s; simpl; auto.
- destruct acc as [acct accf]; simpl in *.
- intros.
- inv avl.
- apply IHs2; auto.
- apply IHs1; auto.
- destruct (f t); simpl; auto.
-Qed.
+ unfold filter; intros; apply filter_acc_bst; auto.
+Qed.
-Lemma partition_acc_avl_2 : forall s acc, avl s ->
- avl (snd acc) -> avl (snd (partition_acc acc s)).
-Proof.
- induction s; simpl; auto.
- destruct acc as [acct accf]; simpl in *.
- intros.
- inv avl.
- apply IHs2; auto.
- apply IHs1; auto.
- destruct (f t); simpl; auto.
-Qed.
-Hint Resolve partition_acc_avl_1 partition_acc_avl_2.
-Lemma partition_acc_bst_1 : forall s acc, bst s -> avl s ->
- bst (fst acc) -> avl (fst acc) ->
- bst (fst (partition_acc acc s)).
-Proof.
- induction s; simpl; auto.
- destruct acc as [acct accf]; simpl in *.
- intros.
- inv avl; inv bst.
- destruct (f t); auto.
- apply IHs2; simpl; auto.
- apply IHs1; simpl; auto.
- apply add_bst; auto.
- apply partition_acc_avl_1; simpl; auto.
-Qed.
-Lemma partition_acc_bst_2 : forall s acc, bst s -> avl s ->
- bst (snd acc) -> avl (snd acc) ->
- bst (snd (partition_acc acc s)).
-Proof.
- induction s; simpl; auto.
- destruct acc as [acct accf]; simpl in *.
- intros.
- inv avl; inv bst.
- destruct (f t); auto.
- apply IHs2; simpl; auto.
- apply IHs1; simpl; auto.
- apply add_bst; auto.
- apply partition_acc_avl_2; simpl; auto.
-Qed.
+(** * Partition *)
-Lemma partition_acc_in_1 : forall s acc, avl s -> avl (fst acc) ->
+Lemma partition_acc_in_1 : forall s acc,
compat_bool X.eq f -> forall x : elt,
- In x (fst (partition_acc acc s)) <->
- In x (fst acc) \/ In x s /\ f x = true.
+ In x (partition_acc f acc s)#1 <->
+ In x acc#1 \/ In x s /\ f x = true.
Proof.
induction s; simpl; intros.
intuition_in.
destruct acc as [acct accf]; simpl in *.
- inv bst; inv avl.
- rewrite IHs2; auto.
- destruct (f t); auto.
- apply partition_acc_avl_1; simpl; auto.
- rewrite IHs1; auto.
- destruct (f t); simpl; auto.
+ rewrite IHs2 by
+ (destruct (f t); auto; apply partition_acc_avl_1; simpl; auto).
+ rewrite IHs1 by (destruct (f t); simpl; auto).
case_eq (f t); simpl; intros.
rewrite (add_in); auto.
intuition_in.
- rewrite (H1 _ _ H8).
+ rewrite (H _ _ H2).
intuition.
intuition_in.
- rewrite (H1 _ _ H8) in H9.
- rewrite H in H9; discriminate.
-Qed.
+ rewrite (H _ _ H2) in H3.
+ rewrite H0 in H3; discriminate.
+Qed.
-Lemma partition_acc_in_2 : forall s acc, avl s -> avl (snd acc) ->
+Lemma partition_acc_in_2 : forall s acc,
compat_bool X.eq f -> forall x : elt,
- In x (snd (partition_acc acc s)) <->
- In x (snd acc) \/ In x s /\ f x = false.
+ In x (partition_acc f acc s)#2 <->
+ In x acc#2 \/ In x s /\ f x = false.
Proof.
induction s; simpl; intros.
intuition_in.
destruct acc as [acct accf]; simpl in *.
- inv bst; inv avl.
- rewrite IHs2; auto.
- destruct (f t); auto.
- apply partition_acc_avl_2; simpl; auto.
- rewrite IHs1; auto.
- destruct (f t); simpl; auto.
+ rewrite IHs2 by
+ (destruct (f t); auto; apply partition_acc_avl_2; simpl; auto).
+ rewrite IHs1 by (destruct (f t); simpl; auto).
case_eq (f t); simpl; intros.
intuition.
intuition_in.
- rewrite (H1 _ _ H8) in H9.
- rewrite H in H9; discriminate.
+ rewrite (H _ _ H2) in H3.
+ rewrite H0 in H3; discriminate.
rewrite (add_in); auto.
intuition_in.
- rewrite (H1 _ _ H8).
+ rewrite (H _ _ H2).
intuition.
+Qed.
+
+Lemma partition_in_1 : forall s,
+ compat_bool X.eq f -> forall x : elt,
+ In x (partition f s)#1 <-> In x s /\ f x = true.
+Proof.
+ unfold partition; intros; rewrite partition_acc_in_1;
+ simpl in *; intuition_in.
Qed.
-Lemma partition_avl_1 : forall s, avl s -> avl (fst (partition s)).
+Lemma partition_in_2 : forall s,
+ compat_bool X.eq f -> forall x : elt,
+ In x (partition f s)#2 <-> In x s /\ f x = false.
Proof.
- unfold partition; intros; apply partition_acc_avl_1; auto.
+ unfold partition; intros; rewrite partition_acc_in_2;
+ simpl in *; intuition_in.
+Qed.
+
+Lemma partition_acc_bst_1 : forall s acc, bst s -> bst acc#1 ->
+ bst (partition_acc f acc s)#1.
+Proof.
+ induction s; simpl; auto.
+ destruct acc as [acct accf]; simpl in *.
+ intros.
+ inv bst.
+ destruct (f t); auto.
+ apply IHs2; simpl; auto.
+ apply IHs1; simpl; auto.
Qed.
-Lemma partition_avl_2 : forall s, avl s -> avl (snd (partition s)).
+Lemma partition_acc_bst_2 : forall s acc, bst s -> bst acc#2 ->
+ bst (partition_acc f acc s)#2.
Proof.
- unfold partition; intros; apply partition_acc_avl_2; auto.
+ induction s; simpl; auto.
+ destruct acc as [acct accf]; simpl in *.
+ intros.
+ inv bst.
+ destruct (f t); auto.
+ apply IHs2; simpl; auto.
+ apply IHs1; simpl; auto.
Qed.
-Lemma partition_bst_1 : forall s, bst s -> avl s ->
- bst (fst (partition s)).
+Lemma partition_bst_1 : forall s, bst s -> bst (partition f s)#1.
Proof.
unfold partition; intros; apply partition_acc_bst_1; auto.
Qed.
-Lemma partition_bst_2 : forall s, bst s -> avl s ->
- bst (snd (partition s)).
+Lemma partition_bst_2 : forall s, bst s -> bst (partition f s)#2.
Proof.
unfold partition; intros; apply partition_acc_bst_2; auto.
Qed.
-Lemma partition_in_1 : forall s, avl s ->
- compat_bool X.eq f -> forall x : elt,
- In x (fst (partition s)) <-> In x s /\ f x = true.
-Proof.
- unfold partition; intros; rewrite partition_acc_in_1;
- simpl in *; intuition_in.
-Qed.
-
-Lemma partition_in_2 : forall s, avl s ->
- compat_bool X.eq f -> forall x : elt,
- In x (snd (partition s)) <-> In x s /\ f x = false.
-Proof.
- unfold partition; intros; rewrite partition_acc_in_2;
- simpl in *; intuition_in.
-Qed.
-(** [for_all] and [exists] *)
-Fixpoint for_all (s:t) : bool := match s with
- | Leaf => true
- | Node l x r _ => f x && for_all l && for_all r
-end.
+(** * [for_all] and [exists] *)
-Lemma for_all_1 : forall s, compat_bool E.eq f ->
- For_all (fun x => f x = true) s -> for_all s = true.
+Lemma for_all_1 : forall s, compat_bool X.eq f ->
+ For_all (fun x => f x = true) s -> for_all f s = true.
Proof.
induction s; simpl; auto.
intros.
@@ -1660,8 +1468,8 @@ Proof.
destruct (f t); simpl; auto.
Qed.
-Lemma for_all_2 : forall s, compat_bool E.eq f ->
- for_all s = true -> For_all (fun x => f x = true) s.
+Lemma for_all_2 : forall s, compat_bool X.eq f ->
+ for_all f s = true -> For_all (fun x => f x = true) s.
Proof.
induction s; simpl; auto; intros; red; intros; inv In.
destruct (andb_prop _ _ H0); auto.
@@ -1673,52 +1481,40 @@ Proof.
destruct (andb_prop _ _ H0); auto.
Qed.
-Fixpoint exists_ (s:t) : bool := match s with
- | Leaf => false
- | Node l x r _ => f x || exists_ l || exists_ r
-end.
-
-Lemma exists_1 : forall s, compat_bool E.eq f ->
- Exists (fun x => f x = true) s -> exists_ s = true.
+Lemma exists_1 : forall s, compat_bool X.eq f ->
+ Exists (fun x => f x = true) s -> exists_ f s = true.
Proof.
- induction s; simpl; destruct 2 as (x,(U,V)); inv In.
+ induction s; simpl; destruct 2 as (x,(U,V)); inv In; rewrite <- ?orb_lazy_alt.
rewrite (H _ _ (X.eq_sym H0)); rewrite V; auto.
apply orb_true_intro; left.
- apply orb_true_intro; right; apply IHs1; firstorder.
- apply orb_true_intro; right; apply IHs2; firstorder.
+ apply orb_true_intro; right; apply IHs1; auto; exists x; auto.
+ apply orb_true_intro; right; apply IHs2; auto; exists x; auto.
Qed.
-Lemma exists_2 : forall s, compat_bool E.eq f ->
- exists_ s = true -> Exists (fun x => f x = true) s.
+Lemma exists_2 : forall s, compat_bool X.eq f ->
+ exists_ f s = true -> Exists (fun x => f x = true) s.
Proof.
- induction s; simpl; intros.
+ 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 t; auto.
- destruct (IHs1 H H2); firstorder.
- destruct (IHs2 H H1); firstorder.
-Qed.
+ destruct (IHs1 H H2); auto; exists x; intuition.
+ destruct (IHs2 H H1); auto; exists x; intuition.
+Qed.
End F.
-(** * Fold *)
-Module L := FSetList.Raw X.
-Fixpoint fold (A : Set) (f : elt -> A -> A)(s : tree) {struct s} : A -> A :=
- fun a => match s with
- | Leaf => a
- | Node l x r _ => fold A f r (f x (fold A f l a))
- end.
-Implicit Arguments fold [A].
+(** * Fold *)
-Definition fold' (A : Set) (f : elt -> A -> A)(s : tree) :=
+Definition fold' (A : Type) (f : elt -> A -> A)(s : tree) :=
L.fold f (elements s).
Implicit Arguments fold' [A].
Lemma fold_equiv_aux :
- forall (A : Set) (s : tree) (f : elt -> A -> A) (a : A) (acc : list elt),
+ forall (A : Type) (s : tree) (f : elt -> A -> A) (a : A) (acc : list elt),
L.fold f (elements_aux acc s) a = L.fold f acc (fold f s a).
Proof.
simple induction s.
@@ -1730,7 +1526,7 @@ Proof.
Qed.
Lemma fold_equiv :
- forall (A : Set) (s : tree) (f : elt -> A -> A) (a : A),
+ forall (A : Type) (s : tree) (f : elt -> A -> A) (a : A),
fold f s a = fold' f s a.
Proof.
unfold fold', elements in |- *.
@@ -1741,7 +1537,7 @@ Proof.
Qed.
Lemma fold_1 :
- forall (s:t)(Hs:bst s)(A : Set)(f : elt -> A -> A)(i : A),
+ forall (s:t)(Hs:bst s)(A : Type)(f : elt -> A -> A)(i : A),
fold f s i = fold_left (fun a e => f e a) (elements s) i.
Proof.
intros.
@@ -1752,416 +1548,168 @@ Proof.
apply elements_sort; auto.
Qed.
-(** * Cardinal *)
-
-Fixpoint cardinal (s : tree) : nat :=
- match s with
- | Leaf => 0%nat
- | Node l _ r _ => S (cardinal l + cardinal r)
- end.
+(** * Subset *)
-Lemma cardinal_elements_aux_1 :
- forall s acc, (length acc + cardinal s)%nat = length (elements_aux acc s).
+Lemma subsetl_12 : forall subset_l1 l1 x1 h1 s2,
+ bst (Node l1 x1 Leaf h1) -> bst s2 ->
+ (forall s, bst s -> (subset_l1 s = true <-> Subset l1 s)) ->
+ (subsetl subset_l1 x1 s2 = true <-> Subset (Node l1 x1 Leaf h1) s2 ).
Proof.
- simple induction s; simpl in |- *; intuition.
- rewrite <- H.
- simpl in |- *.
- rewrite <- H0; omega.
-Qed.
+ induction s2 as [|l2 IHl2 x2 r2 IHr2 h2]; simpl; intros.
+ unfold Subset; intuition; try discriminate.
+ assert (H': In x1 Leaf) by auto; inversion H'.
+ inversion_clear H0.
+ specialize (IHl2 H H2 H1).
+ specialize (IHr2 H H3 H1).
+ inv bst. clear H8.
+ destruct X.compare.
+
+ rewrite IHl2; clear H1 IHl2 IHr2.
+ unfold Subset. intuition_in.
+ assert (In a (Node l2 x2 r2 h2)) by auto; intuition_in; order.
+ assert (In a (Node l2 x2 r2 h2)) by auto; intuition_in; order.
-Lemma cardinal_elements_1 : forall s : tree, cardinal s = length (elements s).
-Proof.
- exact (fun s => cardinal_elements_aux_1 s nil).
-Qed.
+ rewrite H1 by auto; clear H1 IHl2 IHr2.
+ unfold Subset. intuition_in.
+ assert (X.eq a x2) by order; intuition_in.
+ assert (In a (Node l2 x2 r2 h2)) by auto; intuition_in; order.
-(** NB: the remaining functions (union, subset, compare) are still defined
- in a dependent style, due to the use of well-founded induction. *)
+ rewrite <-andb_lazy_alt, andb_true_iff, H1 by auto; clear H1 IHl2 IHr2.
+ unfold Subset. intuition_in.
+ assert (H':=mem_2 H6); apply In_1 with x1; auto.
+ apply mem_1; auto.
+ assert (In x1 (Node l2 x2 r2 h2)) by auto; intuition_in; order.
+Qed.
-(** Induction over cardinals *)
-Lemma sorted_subset_cardinal : forall l' l : list X.t,
- sort X.lt l -> sort X.lt l' ->
- (forall x : elt, InA X.eq x l -> InA X.eq x l') -> (length l <= length l')%nat.
+Lemma subsetr_12 : 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.
- simple induction l'; simpl in |- *; intuition.
- destruct l; trivial; intros.
- absurd (InA X.eq t nil); intuition.
- inversion_clear H2.
- inversion_clear H1.
- destruct l0; simpl in |- *; intuition.
+ induction s2 as [|l2 IHl2 x2 r2 IHr2 h2]; simpl; intros.
+ unfold Subset; intuition; try discriminate.
+ assert (H': In x1 Leaf) by auto; inversion H'.
inversion_clear H0.
- apply le_n_S.
- case (X.compare t a); intro.
- absurd (InA X.eq t (a :: l)).
- intro.
- inversion_clear H0.
- order.
- assert (X.lt a t).
- apply MX.Sort_Inf_In with l; auto.
- order.
- firstorder.
- apply H; auto.
- intros.
- assert (InA X.eq x (a :: l)).
- apply H2; auto.
- inversion_clear H6; auto.
- assert (X.lt t x).
- apply MX.Sort_Inf_In with l0; auto.
- order.
- apply le_trans with (length (t :: l0)).
- simpl in |- *; omega.
- apply (H (t :: l0)); auto.
- intros.
- assert (InA X.eq x (a :: l)); firstorder.
- inversion_clear H6; auto.
- assert (X.lt a x).
- apply MX.Sort_Inf_In with (t :: l0); auto.
- elim (X.lt_not_eq (x:=a) (y:=x)); auto.
-Qed.
-
-Lemma cardinal_subset : forall a b : tree, bst a -> bst b ->
- (forall y : elt, In y a -> In y b) ->
- (cardinal a <= cardinal b)%nat.
-Proof.
- intros.
- do 2 rewrite cardinal_elements_1.
- apply sorted_subset_cardinal; auto.
- intros.
- generalize (elements_in a x) (elements_in b x).
- intuition.
+ specialize (IHl2 H H2 H1).
+ specialize (IHr2 H H3 H1).
+ inv bst. clear H7.
+ destruct X.compare.
+
+ rewrite <-andb_lazy_alt, andb_true_iff, H1 by auto; clear H1 IHl2 IHr2.
+ unfold Subset. intuition_in.
+ assert (H':=mem_2 H1); apply In_1 with x1; auto.
+ apply mem_1; auto.
+ assert (In x1 (Node l2 x2 r2 h2)) by auto; intuition_in; order.
+
+ rewrite H1 by auto; clear H1 IHl2 IHr2.
+ unfold Subset. intuition_in.
+ assert (X.eq a x2) by order; intuition_in.
+ assert (In a (Node l2 x2 r2 h2)) by auto; intuition_in; order.
+
+ rewrite IHr2; clear H1 IHl2 IHr2.
+ unfold Subset. intuition_in.
+ assert (In a (Node l2 x2 r2 h2)) by auto; intuition_in; order.
+ assert (In a (Node l2 x2 r2 h2)) by auto; intuition_in; order.
Qed.
-Lemma cardinal_left : forall (l r : tree) (x : elt) (h : int),
- (cardinal l < cardinal (Node l x r h))%nat.
-Proof.
- simpl in |- *; intuition.
-Qed.
-Lemma cardinal_right :
- forall (l r : tree) (x : elt) (h : int),
- (cardinal r < cardinal (Node l x r h))%nat.
+Lemma subset_12 : forall s1 s2, bst s1 -> bst s2 ->
+ (subset s1 s2 = true <-> Subset s1 s2).
Proof.
- simpl in |- *; intuition.
-Qed.
+ 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': In x1 Leaf) by auto; inversion H'.
+ inv bst.
+ destruct X.compare.
+
+ rewrite <-andb_lazy_alt, andb_true_iff, IHr1 by auto.
+ rewrite (@subsetl_12 (subset l1) l1 x1 h1) by auto.
+ clear IHl1 IHr1.
+ unfold Subset; intuition_in.
+ assert (In a (Node l2 x2 r2 h2)) by auto; intuition_in; order.
+ assert (In a (Node l2 x2 r2 h2)) by auto; intuition_in; order.
+
+ 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 (In a (Node l2 x2 r2 h2)) by auto; intuition_in; order.
+ assert (In a (Node l2 x2 r2 h2)) by auto; intuition_in; order.
+
+ rewrite <-andb_lazy_alt, andb_true_iff, IHl1 by auto.
+ rewrite (@subsetr_12 (subset r1) r1 x1 h1) by auto.
+ clear IHl1 IHr1.
+ unfold Subset; intuition_in.
+ assert (In a (Node l2 x2 r2 h2)) by auto; intuition_in; order.
+ assert (In a (Node l2 x2 r2 h2)) by auto; intuition_in; order.
+Qed.
-Lemma cardinal_rec2 : forall P : tree -> tree -> Set,
- (forall s1 s2 : tree,
- (forall t1 t2 : tree,
- (cardinal t1 + cardinal t2 < cardinal s1 + cardinal s2)%nat -> P t1 t2)
- -> P s1 s2) ->
- forall s1 s2 : tree, P s1 s2.
-Proof.
- intros P H s1 s2.
- apply well_founded_induction_type_2
- with (R := fun yy' xx' : tree * tree =>
- (cardinal (fst yy') + cardinal (snd yy') <
- cardinal (fst xx') + cardinal (snd xx'))%nat); auto.
- apply (Wf_nat.well_founded_ltof _
- (fun xx' : tree * tree => (cardinal (fst xx') + cardinal (snd xx'))%nat)).
-Qed.
-
-Lemma height_0 : forall s, avl s -> height s = 0 -> s = Leaf.
-Proof.
- destruct 1; intuition; simpl in *.
- avl_nns; simpl in *; false_omega_max.
-Qed.
-
-(** * Union
-
- [union s1 s2] does an induction over the sum of the cardinals of
- [s1] and [s2]. Code is
-<<
- let rec union s1 s2 =
- match (s1, s2) with
- (Empty, t2) -> t2
- | (t1, Empty) -> t1
- | (Node(l1, v1, r1, h1), Node(l2, v2, r2, h2)) ->
- if h1 >= h2 then
- if h2 = 1 then add v2 s1 else begin
- let (l2', _, r2') = split v1 s2 in
- join (union l1 l2') v1 (union r1 r2')
- end
- else
- if h1 = 1 then add v1 s2 else begin
- let (l1', _, r1') = split v2 s1 in
- join (union l1' l2) v2 (union r1' r2)
- end
->>
-*)
-Definition union :
- forall s1 s2, bst s1 -> avl s1 -> bst s2 -> avl s2 ->
- {s' : t | bst s' /\ avl s' /\ forall x : elt, In x s' <-> In x s1 \/ In x s2}.
-Proof.
- intros s1 s2; pattern s1, s2; apply cardinal_rec2; clear s1 s2.
- destruct s1 as [| l1 x1 r1 h1]; intros.
- (* s = Leaf *)
- clear H.
- exists s2; intuition_in.
- (* s1 = Node l1 x1 r1 *)
- destruct s2 as [| l2 x2 r2 h2]; simpl in |- *.
- (* s2 = Leaf *)
- clear H.
- exists (Node l1 x1 r1 h1); simpl; intuition_in.
- (* x' = Node l2 x2 r2 *)
- case (ge_lt_dec h1 h2); intro.
- (* h1 >= h2 *)
- case (eq_dec h2 1); intro.
- (* h2 = 1 *)
- clear H.
- exists (add x2 (Node l1 x1 r1 h1)); auto.
- inv avl; inv bst.
- avl_nn l2; avl_nn r2.
- rewrite (height_0 _ H); [ | omega_max].
- rewrite (height_0 _ H4); [ | omega_max].
- split; [apply add_bst; auto|].
- split; [apply add_avl; auto|].
- intros.
- rewrite (add_in (Node l1 x1 r1 h1) x2 x); intuition_in.
- (* h2 <> 1 *)
- (* split x1 s2 = l2',_,r2' *)
- case_eq (split x1 (Node l2 x2 r2 h2)); intros l2' (b,r2') EqSplit.
- set (s2 := Node l2 x2 r2 h2) in *; clearbody s2.
- generalize (split_avl s2 x1 H3); rewrite EqSplit; simpl in *; intros (A,B).
- generalize (split_bst s2 x1 H2 H3); rewrite EqSplit; simpl in *; intros (C,D).
- generalize (split_in_1 s2 x1); rewrite EqSplit; simpl in *; intros.
- generalize (split_in_2 s2 x1); rewrite EqSplit; simpl in *; intros.
- (* union l1 l2' = l0 *)
- destruct (H l1 l2') as [l0 (H7,(H8,H9))]; inv avl; inv bst; auto.
- assert (cardinal l2' <= cardinal s2)%nat.
- apply cardinal_subset; trivial.
- intros y; rewrite (H4 y); intuition.
- omega.
- (* union r1 r2' = r0 *)
- destruct (H r1 r2') as [r0 (H10,(H11,H12))]; inv avl; inv bst; auto.
- assert (cardinal r2' <= cardinal s2)%nat.
- apply cardinal_subset; trivial.
- intros y; rewrite (H5 y); intuition.
- omega.
- exists (join l0 x1 r0).
- inv avl; inv bst; clear H.
- split.
- apply join_bst; auto.
- red; intros.
- rewrite (H9 y) in H.
- destruct H; auto.
- rewrite (H4 y) in H; intuition.
- red; intros.
- rewrite (H12 y) in H.
- destruct H; auto.
- rewrite (H5 y) in H; intuition.
- split.
- apply join_avl; auto.
- intros.
- rewrite join_in; auto.
- rewrite H9.
- rewrite H12.
- rewrite H4; auto.
- rewrite H5; auto.
- intuition_in.
- case (X.compare x x1); intuition.
- (* h1 < h2 *)
- case (eq_dec h1 1); intro.
- (* h1 = 1 *)
- exists (add x1 (Node l2 x2 r2 h2)); auto.
- inv avl; inv bst.
- avl_nn l1; avl_nn r1.
- rewrite (height_0 _ H3); [ | omega_max].
- rewrite (height_0 _ H8); [ | omega_max].
- split; [apply add_bst; auto|].
- split; [apply add_avl; auto|].
- intros.
- rewrite (add_in (Node l2 x2 r2 h2) x1 x); intuition_in.
- (* h1 <> 1 *)
- (* split x2 s1 = l1',_,r1' *)
- case_eq (split x2 (Node l1 x1 r1 h1)); intros l1' (b,r1') EqSplit.
- set (s1 := Node l1 x1 r1 h1) in *; clearbody s1.
- generalize (split_avl s1 x2 H1); rewrite EqSplit; simpl in *; intros (A,B).
- generalize (split_bst s1 x2 H0 H1); rewrite EqSplit; simpl in *; intros (C,D).
- generalize (split_in_1 s1 x2); rewrite EqSplit; simpl in *; intros.
- generalize (split_in_2 s1 x2); rewrite EqSplit; simpl in *; intros.
- (* union l1' l2 = l0 *)
- destruct (H l1' l2) as [l0 (H7,(H8,H9))]; inv avl; inv bst; auto.
- assert (cardinal l1' <= cardinal s1)%nat.
- apply cardinal_subset; trivial.
- intros y; rewrite (H4 y); intuition.
- omega.
- (* union r1' r2 = r0 *)
- destruct (H r1' r2) as [r0 (H10,(H11,H12))]; inv avl; inv bst; auto.
- assert (cardinal r1' <= cardinal s1)%nat.
- apply cardinal_subset; trivial.
- intros y; rewrite (H5 y); intuition.
- omega.
- exists (join l0 x2 r0).
- inv avl; inv bst; clear H.
- split.
- apply join_bst; auto.
- red; intros.
- rewrite (H9 y) in H.
- destruct H; auto.
- rewrite (H4 y) in H; intuition.
- red; intros.
- rewrite (H12 y) in H.
- destruct H; auto.
- rewrite (H5 y) in H; intuition.
- split.
- apply join_avl; auto.
- intros.
- rewrite join_in; auto.
- rewrite H9.
- rewrite H12.
- rewrite H4; auto.
- rewrite H5; auto.
- intuition_in.
- case (X.compare x x2); intuition.
-Qed.
-
-
-(** * Subset
-<<
- let rec subset s1 s2 =
- match (s1, s2) with
- Empty, _ -> true
- | _, Empty -> false
- | Node (l1, v1, r1, _), (Node (l2, v2, r2, _) as t2) ->
- let c = Ord.compare v1 v2 in
- if c = 0 then
- subset l1 l2 && subset r1 r2
- else if c < 0 then
- subset (Node (l1, v1, Empty, 0)) l2 && subset r1 t2
- else
- subset (Node (Empty, v1, r1, 0)) r2 && subset l1 t2
->>
-*)
-
-Definition subset : forall s1 s2 : t, bst s1 -> bst s2 ->
- {Subset s1 s2} + {~ Subset s1 s2}.
-Proof.
- intros s1 s2; pattern s1, s2; apply cardinal_rec2; clear s1 s2.
- destruct s1 as [| l1 x1 r1 h1]; intros.
- (* s1 = Leaf *)
- left; red; intros; inv In.
- (* s1 = Node l1 x1 r1 h1 *)
- destruct s2 as [| l2 x2 r2 h2].
- (* s2 = Leaf *)
- right; intros; intro.
- assert (In x1 Leaf); auto.
- inversion_clear H3.
- (* s2 = Node l2 x2 r2 h2 *)
- case (X.compare x1 x2); intro.
- (* x1 < x2 *)
- case (H (Node l1 x1 Leaf 0) l2); inv bst; auto; intros.
- simpl in |- *; omega.
- case (H r1 (Node l2 x2 r2 h2)); inv bst; auto; intros.
- simpl in |- *; omega.
- clear H; left; red; intuition.
- generalize (s a) (s0 a); clear s s0; intuition_in.
- clear H; right; red; firstorder.
- clear H; right; red; inv bst; intuition.
- apply n; red; intros.
- assert (In a (Node l2 x2 r2 h2)) by (inv In; auto).
- intuition_in; order.
- (* x1 = x2 *)
- case (H l1 l2); inv bst; auto; intros.
- simpl in |- *; omega.
- case (H r1 r2); inv bst; auto; intros.
- simpl in |- *; omega.
- clear H; left; red; intuition_in; eauto.
- clear H; right; red; inv bst; intuition.
- apply n; red; intros.
- assert (In a (Node l2 x2 r2 h2)) by auto.
- intuition_in; order.
- clear H; right; red; inv bst; intuition.
- apply n; red; intros.
- assert (In a (Node l2 x2 r2 h2)) by auto.
- intuition_in; order.
- (* x1 > x2 *)
- case (H (Node Leaf x1 r1 0) r2); inv bst; auto; intros.
- simpl in |- *; omega.
- intros; case (H l1 (Node l2 x2 r2 h2)); inv bst; auto; intros.
- simpl in |- *; omega.
- clear H; left; red; intuition.
- generalize (s a) (s0 a); clear s s0; intuition_in.
- clear H; right; red; firstorder.
- clear H; right; red; inv bst; intuition.
- apply n; red; intros.
- assert (In a (Node l2 x2 r2 h2)) by (inv In; auto).
- intuition_in; order.
-Qed.
(** * Comparison *)
(** ** Relations [eq] and [lt] over trees *)
-Definition eq : t -> t -> Prop := Equal.
+Definition eq := Equal.
+Definition lt (s1 s2 : t) : Prop := L.lt (elements s1) (elements s2).
-Lemma eq_refl : forall s : t, eq s s.
+Lemma eq_refl : forall s : t, Equal s s.
Proof.
- unfold eq, Equal in |- *; intuition.
+ unfold Equal; intuition.
Qed.
-Lemma eq_sym : forall s s' : t, eq s s' -> eq s' s.
+Lemma eq_sym : forall s s' : t, Equal s s' -> Equal s' s.
Proof.
- unfold eq, Equal in |- *; firstorder.
+ unfold Equal; intros s s' H x; destruct (H x); split; auto.
Qed.
-Lemma eq_trans : forall s s' s'' : t, eq s s' -> eq s' s'' -> eq s s''.
+Lemma eq_trans : forall s s' s'' : t,
+ Equal s s' -> Equal s' s'' -> Equal s s''.
Proof.
- unfold eq, Equal in |- *; firstorder.
+ unfold Equal; intros s s' s'' H1 H2 x;
+ destruct (H1 x); destruct (H2 x); split; auto.
Qed.
Lemma eq_L_eq :
- forall s s' : t, eq s s' -> L.eq (elements s) (elements s').
+ forall s s' : t, Equal s s' -> L.eq (elements s) (elements s').
Proof.
- unfold eq, Equal, L.eq, L.Equal in |- *; intros.
- generalize (elements_in s a) (elements_in s' a).
- firstorder.
+ unfold Equal, L.eq, L.Equal; intros; do 2 rewrite elements_in; auto.
Qed.
Lemma L_eq_eq :
- forall s s' : t, L.eq (elements s) (elements s') -> eq s s'.
+ forall s s' : t, L.eq (elements s) (elements s') -> Equal s s'.
Proof.
- unfold eq, Equal, L.eq, L.Equal in |- *; intros.
- generalize (elements_in s a) (elements_in s' a).
- firstorder.
+ unfold Equal, L.eq, L.Equal; intros; do 2 rewrite <-elements_in; auto.
Qed.
Hint Resolve eq_L_eq L_eq_eq.
-Definition lt (s1 s2 : t) : Prop := L.lt (elements s1) (elements s2).
-
Definition lt_trans (s s' s'' : t) (h : lt s s')
(h' : lt s' s'') : lt s s'' := L.lt_trans h h'.
-Lemma lt_not_eq : forall s s' : t, bst s -> bst s' -> lt s s' -> ~ eq s s'.
+Lemma lt_not_eq : forall s s' : t,
+ bst s -> bst s' -> lt s s' -> ~ Equal s s'.
Proof.
unfold lt in |- *; intros; intro.
apply L.lt_not_eq with (s := elements s) (s' := elements s'); auto.
Qed.
-(** A new comparison algorithm suggested by Xavier Leroy:
-
-type enumeration = End | More of elt * t * enumeration
-
-let rec cons s e = match s with
- | Empty -> e
- | Node(l, v, r, _) -> cons l (More(v, r, e))
-
-let rec compare_aux e1 e2 = match (e1, e2) with
- | (End, End) -> 0
- | (End, More _) -> -1
- | (More _, End) -> 1
- | (More(v1, r1, k1), More(v2, r2, k2)) ->
- let c = Ord.compare v1 v2 in
- if c <> 0 then c else compare_aux (cons r1 k1) (cons r2 k2)
-
-let compare s1 s2 = compare_aux (cons s1 End) (cons s2 End)
-*)
+Lemma L_eq_cons :
+ forall (l1 l2 : list elt) (x y : elt),
+ X.eq x y -> L.eq l1 l2 -> L.eq (x :: l1) (y :: l2).
+Proof.
+ unfold L.eq, L.Equal in |- *; intuition.
+ inversion_clear H1; generalize (H0 a); clear H0; intuition.
+ apply InA_eqA with x; eauto.
+ inversion_clear H1; generalize (H0 a); clear H0; intuition.
+ apply InA_eqA with y; eauto.
+Qed.
+Hint Resolve L_eq_cons.
-(** ** Enumeration of the elements of a tree *)
-Inductive enumeration : Set :=
- | End : enumeration
- | More : elt -> tree -> enumeration -> enumeration.
+(** * A new comparison algorithm suggested by Xavier Leroy *)
(** [flatten_e e] returns the list of elements of [e] i.e. the list
of elements actually compared *)
@@ -2171,462 +1719,166 @@ Fixpoint flatten_e (e : enumeration) : list elt := match e with
| More x t r => x :: elements t ++ flatten_e r
end.
-(** [sorted_e e] expresses that elements in the enumeration [e] are
- sorted, and that all trees in [e] are binary search trees. *)
-
-Inductive In_e (x:elt) : enumeration -> Prop :=
- | InEHd1 :
- forall (y : elt) (s : tree) (e : enumeration),
- X.eq x y -> In_e x (More y s e)
- | InEHd2 :
- forall (y : elt) (s : tree) (e : enumeration),
- In x s -> In_e x (More y s e)
- | InETl :
- forall (y : elt) (s : tree) (e : enumeration),
- In_e x e -> In_e x (More y s e).
-
-Hint Constructors In_e.
-
-Inductive sorted_e : enumeration -> Prop :=
- | SortedEEnd : sorted_e End
- | SortedEMore :
- forall (x : elt) (s : tree) (e : enumeration),
- bst s ->
- (gt_tree x s) ->
- sorted_e e ->
- (forall y : elt, In_e y e -> X.lt x y) ->
- (forall y : elt,
- In y s -> forall z : elt, In_e z e -> X.lt y z) ->
- sorted_e (More x s e).
-
-Hint Constructors sorted_e.
-
-Lemma in_app :
- forall (x : elt) (l1 l2 : list elt),
- InA X.eq x (l1 ++ l2) -> InA X.eq x l1 \/ InA X.eq x l2.
-Proof.
- simple induction l1; simpl in |- *; intuition.
- inversion_clear H0; auto.
- elim (H l2 H1); auto.
-Qed.
-
-Lemma in_flatten_e :
- forall (x : elt) (e : enumeration), InA X.eq x (flatten_e e) -> In_e x e.
-Proof.
- simple induction e; simpl in |- *; intuition.
- inversion_clear H.
- inversion_clear H0; auto.
- elim (in_app x _ _ H1); auto.
- destruct (elements_in t x); auto.
-Qed.
-
-Lemma sort_app :
- forall l1 l2 : list elt, sort X.lt l1 -> sort X.lt l2 ->
- (forall x y : elt, InA X.eq x l1 -> InA X.eq y l2 -> X.lt x y) ->
- sort X.lt (l1 ++ l2).
-Proof.
- simple induction l1; simpl in |- *; intuition.
- apply cons_sort; auto.
- apply H; auto.
- inversion_clear H0; trivial.
- induction l as [| a0 l Hrecl]; simpl in |- *; intuition.
- induction l2 as [| a0 l2 Hrecl2]; simpl in |- *; intuition.
- inversion_clear H0; inversion_clear H4; auto.
-Qed.
-
-Lemma sorted_flatten_e :
- forall e : enumeration, sorted_e e -> sort X.lt (flatten_e e).
-Proof.
- simple induction e; simpl in |- *; intuition.
- apply cons_sort.
- apply sort_app; inversion H0; auto.
- intros; apply H8; auto.
- destruct (elements_in t x0); auto.
- apply in_flatten_e; auto.
- apply L.MX.ListIn_Inf.
- inversion_clear H0.
- intros; elim (in_app_or _ _ _ H0); intuition.
- destruct (elements_in t y); auto.
- apply H4; apply in_flatten_e; auto.
-Qed.
-
-Lemma elements_app :
- forall (s : tree) (acc : list elt), elements_aux acc s = elements s ++ acc.
-Proof.
- simple induction s; simpl in |- *; intuition.
- rewrite H0.
- rewrite H.
- unfold elements; simpl.
- do 2 rewrite H.
- rewrite H0.
- repeat rewrite <- app_nil_end.
- repeat rewrite app_ass; auto.
-Qed.
-
-Lemma compare_flatten_1 :
- forall (t0 t2 : tree) (t1 : elt) (z : int) (l : list elt),
- elements t0 ++ t1 :: elements t2 ++ l =
- elements (Node t0 t1 t2 z) ++ l.
+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.
- simpl in |- *; unfold elements in |- *; simpl in |- *; intuition.
- repeat rewrite elements_app.
- repeat rewrite <- app_nil_end.
- repeat rewrite app_ass; auto.
+ intros; simpl; apply elements_node.
Qed.
-(** key lemma for correctness *)
-
-Lemma flatten_e_elements :
- forall (x : elt) (l r : tree) (z : int) (e : enumeration),
- elements l ++ flatten_e (More x r e) = elements (Node l x r z) ++ flatten_e e.
+Lemma cons_1 : forall s e,
+ flatten_e (cons s e) = elements s ++ flatten_e e.
Proof.
- intros; simpl.
- apply compare_flatten_1.
+ induction s; simpl; auto; intros.
+ rewrite IHs1; apply flatten_e_elements.
Qed.
-(** termination of [compare_aux] *)
+(** Correctness of this comparison *)
-Open Local Scope Z_scope.
-
-Fixpoint measure_e_t (s : tree) : Z := match s with
- | Leaf => 0
- | Node l _ r _ => 1 + measure_e_t l + measure_e_t r
- end.
-
-Fixpoint measure_e (e : enumeration) : Z := match e with
- | End => 0
- | More _ s r => 1 + measure_e_t s + measure_e r
+Definition Cmp c :=
+ match c with
+ | Eq => L.eq
+ | Lt => L.lt
+ | Gt => (fun l1 l2 => L.lt l2 l1)
end.
-Ltac Measure_e_t := unfold measure_e_t in |- *; fold measure_e_t in |- *.
-Ltac Measure_e := unfold measure_e in |- *; fold measure_e in |- *.
-
-Lemma measure_e_t_0 : forall s : tree, measure_e_t s >= 0.
+Lemma cons_Cmp : forall c x1 x2 l1 l2, X.eq x1 x2 ->
+ Cmp c l1 l2 -> Cmp c (x1::l1) (x2::l2).
Proof.
- simple induction s.
- simpl in |- *; omega.
- intros.
- Measure_e_t; omega. (* BUG Simpl! *)
+ destruct c; simpl; auto.
Qed.
+Hint Resolve cons_Cmp.
-Ltac Measure_e_t_0 s := generalize (measure_e_t_0 s); intro.
-
-Lemma measure_e_0 : forall e : enumeration, measure_e e >= 0.
+Lemma compare_end_Cmp :
+ forall e2, Cmp (compare_end e2) nil (flatten_e e2).
Proof.
- simple induction e.
- simpl in |- *; omega.
- intros.
- Measure_e; Measure_e_t_0 t; omega.
+ destruct e2; simpl; auto.
+ apply L.eq_refl.
Qed.
-Ltac Measure_e_0 e := generalize (measure_e_0 e); intro.
-
-(** Induction principle over the sum of the measures for two lists *)
-
-Definition compare_rec2 :
- forall P : enumeration -> enumeration -> Set,
- (forall x x' : enumeration,
- (forall y y' : enumeration,
- measure_e y + measure_e y' < measure_e x + measure_e x' -> P y y') ->
- P x x') ->
- forall x x' : enumeration, P x x'.
+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.
- intros P H x x'.
- apply well_founded_induction_type_2
- with (R := fun yy' xx' : enumeration * enumeration =>
- measure_e (fst yy') + measure_e (snd yy') <
- measure_e (fst xx') + measure_e (snd xx')); auto.
- apply Wf_nat.well_founded_lt_compat
- with (f := fun xx' : enumeration * enumeration =>
- Zabs_nat (measure_e (fst xx') + measure_e (snd xx'))).
- intros; apply Zabs.Zabs_nat_lt.
- Measure_e_0 (fst x0); Measure_e_0 (snd x0); Measure_e_0 (fst y);
- Measure_e_0 (snd y); intros; omega.
-Qed.
-
-(** [cons t e] adds the elements of tree [t] on the head of
- enumeration [e]. Code:
-
-let rec cons s e = match s with
- | Empty -> e
- | Node(l, v, r, _) -> cons l (More(v, r, e))
-*)
-
-Definition cons : forall (s : tree) (e : enumeration), bst s -> sorted_e e ->
- (forall (x y : elt), In x s -> In_e y e -> X.lt x y) ->
- { r : enumeration
- | sorted_e r /\
- measure_e r = measure_e_t s + measure_e e /\
- flatten_e r = elements s ++ flatten_e e
- }.
-Proof.
- simple induction s; intuition.
- (* s = Leaf *)
- exists e; intuition.
- (* s = Node t t0 t1 z *)
- clear H0.
- case (H (More t0 t1 e)); clear H; intuition.
- inv bst; auto.
- constructor; inversion_clear H1; auto.
- inversion_clear H0; inv bst; intuition; order.
- exists x; intuition.
- generalize H4; Measure_e; intros; Measure_e_t; omega.
- rewrite H5.
- apply flatten_e_elements.
+ simpl; intros; destruct X.compare; simpl; auto.
Qed.
-Lemma l_eq_cons :
- forall (l1 l2 : list elt) (x y : elt),
- X.eq x y -> L.eq l1 l2 -> L.eq (x :: l1) (y :: l2).
+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.
- unfold L.eq, L.Equal in |- *; intuition.
- inversion_clear H1; generalize (H0 a); clear H0; intuition.
- apply InA_eqA with x; eauto.
- inversion_clear H1; generalize (H0 a); clear H0; intuition.
- apply InA_eqA with y; eauto.
+ 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.
-Definition compare_aux :
- forall e1 e2 : enumeration, sorted_e e1 -> sorted_e e2 ->
- Compare L.lt L.eq (flatten_e e1) (flatten_e e2).
-Proof.
- intros e1 e2; pattern e1, e2 in |- *; apply compare_rec2.
- simple destruct x; simple destruct x'; intuition.
- (* x = x' = End *)
- constructor 2; unfold L.eq, L.Equal in |- *; intuition.
- (* x = End x' = More *)
- constructor 1; simpl in |- *; auto.
- (* x = More x' = End *)
- constructor 3; simpl in |- *; auto.
- (* x = More e t e0, x' = More e3 t0 e4 *)
- case (X.compare e e3); intro.
- (* e < e3 *)
- constructor 1; simpl; auto.
- (* e = e3 *)
- destruct (cons t e0) as [c1 (H2,(H3,H4))]; try inversion_clear H0; auto.
- destruct (cons t0 e4) as [c2 (H5,(H6,H7))]; try inversion_clear H1; auto.
- destruct (H c1 c2); clear H; intuition.
- Measure_e; omega.
- constructor 1; simpl.
- apply L.lt_cons_eq; auto.
- rewrite H4 in l; rewrite H7 in l; auto.
- constructor 2; simpl.
- apply l_eq_cons; auto.
- rewrite H4 in e6; rewrite H7 in e6; auto.
- constructor 3; simpl.
- apply L.lt_cons_eq; auto.
- rewrite H4 in l; rewrite H7 in l; auto.
- (* e > e3 *)
- constructor 3; simpl; auto.
-Qed.
-
-Definition compare : forall s1 s2, bst s1 -> bst s2 -> Compare lt eq s1 s2.
-Proof.
- intros s1 s2 s1_bst s2_bst; unfold lt, eq; simpl.
- destruct (cons s1 End); intuition.
- inversion_clear H0.
- destruct (cons s2 End); intuition.
- inversion_clear H3.
- simpl in H2; rewrite <- app_nil_end in H2.
- simpl in H5; rewrite <- app_nil_end in H5.
- destruct (compare_aux x x0); intuition.
- constructor 1; simpl in |- *.
- rewrite H2 in l; rewrite H5 in l; auto.
- constructor 2; apply L_eq_eq; simpl in |- *.
- rewrite H2 in e; rewrite H5 in e; auto.
- constructor 3; simpl in |- *.
- rewrite H2 in l; rewrite H5 in l; auto.
+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.
(** * Equality test *)
-Definition equal : forall s s' : t, bst s -> bst s' -> {Equal s s'} + {~ Equal s s'}.
+Lemma equal_1 : forall s1 s2, bst s1 -> bst s2 ->
+ Equal s1 s2 -> equal s1 s2 = true.
Proof.
- intros s s' Hs Hs'; case (compare s s'); auto; intros.
- right; apply lt_not_eq; auto.
- right; intro; apply (lt_not_eq s' s); auto; apply eq_sym; auto.
-Qed.
-
-(** We provide additionally a different implementation for union, subset and
- equal, which is less efficient, but uses structural induction, hence computes
- within Coq. *)
-
-(** Alternative union based on fold.
- Complexity : [min(|s|,|s'|)*log(max(|s|,|s'|))] *)
-
-Definition union' s s' :=
- if ge_lt_dec (height s) (height s') then fold add s' s else fold add s s'.
-
-Lemma fold_add_avl : forall s s', avl s -> avl s' -> avl (fold add s s').
-Proof.
- induction s; simpl; intros; inv avl; auto.
-Qed.
-Hint Resolve fold_add_avl.
-
-Lemma union'_avl : forall s s', avl s -> avl s' -> avl (union' s s').
-Proof.
- unfold union'; intros; destruct (ge_lt_dec (height s) (height s')); auto.
-Qed.
-
-Lemma fold_add_bst : forall s s', bst s -> avl s -> bst s' -> avl s' ->
- bst (fold add s s').
-Proof.
- induction s; simpl; intros; inv avl; inv bst; auto.
- apply IHs2; auto.
- apply add_bst; auto.
-Qed.
-
-Lemma union'_bst : forall s s', bst s -> avl s -> bst s' -> avl s' ->
- bst (union' s s').
-Proof.
- unfold union'; intros; destruct (ge_lt_dec (height s) (height s'));
- apply fold_add_bst; auto.
+unfold equal; intros s1 s2 B1 B2 E.
+generalize (compare_Cmp s1 s2).
+destruct (compare s1 s2); simpl in *; auto; intros.
+elim (lt_not_eq B1 B2 H E); auto.
+elim (lt_not_eq B2 B1 H (eq_sym E)); auto.
Qed.
-Lemma fold_add_in : forall s s' y, bst s -> avl s -> bst s' -> avl s' ->
- (In y (fold add s s') <-> In y s \/ In y s').
-Proof.
- induction s; simpl; intros; inv avl; inv bst; auto.
- intuition_in.
- rewrite IHs2; auto.
- apply add_bst; auto.
- apply fold_add_bst; auto.
- rewrite add_in; auto.
- rewrite IHs1; auto.
- intuition_in.
-Qed.
-
-Lemma union'_in : forall s s' y, bst s -> avl s -> bst s' -> avl s' ->
- (In y (union' s s') <-> In y s \/ In y s').
+Lemma equal_2 : forall s1 s2,
+ equal s1 s2 = true -> Equal s1 s2.
Proof.
- unfold union'; intros; destruct (ge_lt_dec (height s) (height s')).
- rewrite fold_add_in; intuition.
- apply fold_add_in; auto.
-Qed.
-
-(** Alternative subset based on diff. *)
-
-Definition subset' s s' := is_empty (diff s s').
-
-Lemma subset'_1 : forall s s', bst s -> avl s -> bst s' -> avl s' ->
- Subset s s' -> subset' s s' = true.
-Proof.
- unfold subset', Subset; intros; apply is_empty_1; red; intros.
- rewrite (diff_in); intuition.
+unfold equal; intros s1 s2 E.
+generalize (compare_Cmp s1 s2);
+ destruct compare; auto; discriminate.
Qed.
-Lemma subset'_2 : forall s s', bst s -> avl s -> bst s' -> avl s' ->
- subset' s s' = true -> Subset s s'.
-Proof.
- unfold subset', Subset; intros; generalize (is_empty_2 _ H3 a); unfold Empty.
- rewrite (diff_in); intuition.
- generalize (mem_2 s' a) (mem_1 s' a); destruct (mem a s'); intuition.
-Qed.
+End Proofs.
-(** Alternative equal based on subset *)
+End Raw.
-Definition equal' s s' := subset' s s' && subset' s' s.
-Lemma equal'_1 : forall s s', bst s -> avl s -> bst s' -> avl s' ->
- Equal s s' -> equal' s s' = true.
-Proof.
- unfold equal', Equal; intros.
- rewrite subset'_1; firstorder; simpl.
- apply subset'_1; firstorder.
-Qed.
-
-Lemma equal'_2 : forall s s', bst s -> avl s -> bst s' -> avl s' ->
- equal' s s' = true -> Equal s s'.
-Proof.
- unfold equal', Equal; intros; destruct (andb_prop _ _ H3); split;
- apply subset'_2; auto.
-Qed.
-
-End Raw.
(** * Encapsulation
Now, in order to really provide a functor implementing [S], we
- need to encapsulate everything into a type of balanced binary search trees. *)
+ 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 [FSetFullAVL] if you need more than just the FSet interface.
+*)
Module IntMake (I:Int)(X: OrderedType) <: S with Module E := X.
Module E := X.
- Module Raw := Raw I X.
+ Module Raw := Raw I X.
+ Import Raw.Proofs.
- Record bbst : Set := Bbst {this :> Raw.t; is_bst : Raw.bst this; is_avl: Raw.avl this}.
- Definition t := bbst.
+ Record bst := Bst {this :> Raw.t; is_bst : Raw.bst this}.
+ Definition t := bst.
Definition elt := E.t.
- Definition In (x : elt) (s : t) : Prop := Raw.In x s.
- Definition Equal (s s':t) : Prop := forall a : elt, In a s <-> In a s'.
- Definition Subset (s s':t) : Prop := forall a : elt, In a s -> In a s'.
- Definition Empty (s:t) : Prop := forall a : elt, ~ In a s.
- Definition For_all (P : elt -> Prop) (s:t) : Prop := forall x, In x s -> P x.
- Definition Exists (P : elt -> Prop) (s:t) : Prop := exists x, In x s /\ P x.
-
+ Definition In (x : elt) (s : t) := Raw.In x s.
+ 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.
+
Lemma In_1 : forall (s:t)(x y:elt), E.eq x y -> In x s -> In y s.
- Proof. intro s; exact (Raw.In_1 s). Qed.
+ Proof. intro s; exact (@In_1 s). Qed.
Definition mem (x:elt)(s:t) : bool := Raw.mem x s.
- Definition empty : t := Bbst _ Raw.empty_bst Raw.empty_avl.
+ Definition empty : t := Bst empty_bst.
Definition is_empty (s:t) : bool := Raw.is_empty s.
- Definition singleton (x:elt) : t := Bbst _ (Raw.singleton_bst x) (Raw.singleton_avl x).
- Definition add (x:elt)(s:t) : t :=
- Bbst _ (Raw.add_bst s x (is_bst s) (is_avl s))
- (Raw.add_avl s x (is_avl s)).
- Definition remove (x:elt)(s:t) : t :=
- Bbst _ (Raw.remove_bst s x (is_bst s) (is_avl s))
- (Raw.remove_avl s x (is_avl s)).
- Definition inter (s s':t) : t :=
- Bbst _ (Raw.inter_bst _ _ (is_bst s) (is_avl s) (is_bst s') (is_avl s'))
- (Raw.inter_avl _ _ (is_avl s) (is_avl s')).
- Definition diff (s s':t) : t :=
- Bbst _ (Raw.diff_bst _ _ (is_bst s) (is_avl s) (is_bst s') (is_avl s'))
- (Raw.diff_avl _ _ (is_avl s) (is_avl s')).
+ Definition singleton (x:elt) : t := Bst (singleton_bst x).
+ Definition add (x:elt)(s:t) : t := Bst (add_bst x (is_bst s)).
+ Definition remove (x:elt)(s:t) : t := Bst (remove_bst x (is_bst s)).
+ Definition inter (s s':t) : t := Bst (inter_bst (is_bst s) (is_bst s')).
+ Definition union (s s':t) : t := Bst (union_bst (is_bst s) (is_bst s')).
+ Definition diff (s s':t) : t := Bst (diff_bst (is_bst s) (is_bst s')).
Definition elements (s:t) : list elt := Raw.elements s.
Definition min_elt (s:t) : option elt := Raw.min_elt s.
Definition max_elt (s:t) : option elt := Raw.max_elt s.
Definition choose (s:t) : option elt := Raw.choose s.
- Definition fold (B : Set) (f : elt -> B -> B) (s:t) : B -> B := Raw.fold f s.
+ Definition fold (B : Type) (f : elt -> B -> B) (s:t) : B -> B := Raw.fold f s.
Definition cardinal (s:t) : nat := Raw.cardinal s.
Definition filter (f : elt -> bool) (s:t) : t :=
- Bbst _ (Raw.filter_bst f _ (is_bst s) (is_avl s))
- (Raw.filter_avl f _ (is_avl s)).
+ Bst (filter_bst f (is_bst s)).
Definition for_all (f : elt -> bool) (s:t) : bool := Raw.for_all f s.
Definition exists_ (f : elt -> bool) (s:t) : bool := Raw.exists_ f s.
Definition partition (f : elt -> bool) (s:t) : t * t :=
let p := Raw.partition f s in
- (Bbst (fst p) (Raw.partition_bst_1 f _ (is_bst s) (is_avl s))
- (Raw.partition_avl_1 f _ (is_avl s)),
- Bbst (snd p) (Raw.partition_bst_2 f _ (is_bst s) (is_avl s))
- (Raw.partition_avl_2 f _ (is_avl s))).
-
- Definition union (s s':t) : t :=
- let (u,p) := Raw.union _ _ (is_bst s) (is_avl s) (is_bst s') (is_avl s') in
- let (b,p) := p in
- let (a,_) := p in
- Bbst u b a.
-
- Definition union' (s s' : t) : t :=
- Bbst _ (Raw.union'_bst _ _ (is_bst s) (is_avl s) (is_bst s') (is_avl s'))
- (Raw.union'_avl _ _ (is_avl s) (is_avl s')).
-
- Definition equal (s s': t) : bool := if Raw.equal _ _ (is_bst s) (is_bst s') then true else false.
- Definition equal' (s s':t) : bool := Raw.equal' s s'.
+ (@Bst (fst p) (partition_bst_1 f (is_bst s)),
+ @Bst (snd p) (partition_bst_2 f (is_bst s))).
- Definition subset (s s':t) : bool := if Raw.subset _ _ (is_bst s) (is_bst s') then true else false.
- Definition subset' (s s':t) : bool := Raw.subset' s s'.
+ Definition equal (s s':t) : bool := Raw.equal s s'.
+ Definition subset (s s':t) : bool := Raw.subset s s'.
- Definition eq (s s':t) : Prop := Raw.eq s s'.
- Definition lt (s s':t) : Prop := Raw.lt s s'.
+ Definition eq (s s':t) : Prop := Raw.Equal s s'.
+ Definition lt (s s':t) : Prop := Raw.Proofs.lt s s'.
- Definition compare (s s':t) : Compare lt eq s s'.
+ Definition compare (s s':t) : Compare lt eq s s'.
Proof.
- intros; elim (Raw.compare _ _ (is_bst s) (is_bst s'));
- [ constructor 1 | constructor 2 | constructor 3 ];
- auto.
+ intros (s,b) (s',b').
+ generalize (compare_Cmp s s').
+ destruct Raw.compare; intros; [apply EQ|apply LT|apply GT]; red; auto.
Defined.
(* specs *)
@@ -2634,260 +1886,164 @@ Module IntMake (I:Int)(X: OrderedType) <: S with Module E := X.
Variable s s' s'': t.
Variable x y : elt.
- Hint Resolve is_bst is_avl.
+ Hint Resolve is_bst.
Lemma mem_1 : In x s -> mem x s = true.
- Proof. exact (Raw.mem_1 s x (is_bst s)). Qed.
+ Proof. exact (mem_1 (is_bst s)). Qed.
Lemma mem_2 : mem x s = true -> In x s.
- Proof. exact (Raw.mem_2 s x). Qed.
+ Proof. exact (@mem_2 s x). Qed.
Lemma equal_1 : Equal s s' -> equal s s' = true.
- Proof.
- unfold equal; destruct (Raw.equal s s'); simpl; auto.
- Qed.
-
+ Proof. exact (equal_1 (is_bst s) (is_bst s')). Qed.
Lemma equal_2 : equal s s' = true -> Equal s s'.
- Proof.
- unfold equal; destruct (Raw.equal s s'); simpl; intuition; discriminate.
- Qed.
+ Proof. exact (@equal_2 s s'). Qed.
- Lemma equal'_1 : Equal s s' -> equal' s s' = true.
- Proof. exact (Raw.equal'_1 _ _ (is_bst s) (is_avl s) (is_bst s') (is_avl s')). Qed.
- Lemma equal'_2 : equal' s s' = true -> Equal s s'.
- Proof. exact (Raw.equal'_2 _ _ (is_bst s) (is_avl s) (is_bst s') (is_avl s')). Qed.
+ Ltac wrap t H := unfold t, In; simpl; rewrite H; auto; intuition.
Lemma subset_1 : Subset s s' -> subset s s' = true.
- Proof.
- unfold subset; destruct (Raw.subset s s'); simpl; intuition.
- Qed.
-
+ Proof. wrap subset subset_12. Qed.
Lemma subset_2 : subset s s' = true -> Subset s s'.
- Proof.
- unfold subset; destruct (Raw.subset s s'); simpl; intuition discriminate.
- Qed.
-
- Lemma subset'_1 : Subset s s' -> subset' s s' = true.
- Proof. exact (Raw.subset'_1 _ _ (is_bst s) (is_avl s) (is_bst s') (is_avl s')). Qed.
- Lemma subset'_2 : subset' s s' = true -> Subset s s'.
- Proof. exact (Raw.subset'_2 _ _ (is_bst s) (is_avl s) (is_bst s') (is_avl s')). Qed.
+ Proof. wrap subset subset_12. Qed.
Lemma empty_1 : Empty empty.
- Proof. exact Raw.empty_1. Qed.
+ Proof. exact empty_1. Qed.
Lemma is_empty_1 : Empty s -> is_empty s = true.
- Proof. exact (Raw.is_empty_1 s). Qed.
+ Proof. exact (@is_empty_1 s). Qed.
Lemma is_empty_2 : is_empty s = true -> Empty s.
- Proof. exact (Raw.is_empty_2 s). Qed.
+ Proof. exact (@is_empty_2 s). Qed.
Lemma add_1 : E.eq x y -> In y (add x s).
- Proof.
- unfold add, In; simpl; rewrite Raw.add_in; auto.
- Qed.
-
+ Proof. wrap add add_in. Qed.
Lemma add_2 : In y s -> In y (add x s).
- Proof.
- unfold add, In; simpl; rewrite Raw.add_in; auto.
- Qed.
-
+ Proof. wrap add add_in. Qed.
Lemma add_3 : ~ E.eq x y -> In y (add x s) -> In y s.
- Proof.
- unfold add, In; simpl; rewrite Raw.add_in; intuition.
- elim H; auto.
- Qed.
+ Proof. wrap add add_in. elim H; auto. Qed.
Lemma remove_1 : E.eq x y -> ~ In y (remove x s).
- Proof.
- unfold remove, In; simpl; rewrite Raw.remove_in; intuition.
- Qed.
-
+ Proof. wrap remove remove_in. Qed.
Lemma remove_2 : ~ E.eq x y -> In y s -> In y (remove x s).
- Proof.
- unfold remove, In; simpl; rewrite Raw.remove_in; intuition.
- Qed.
-
+ Proof. wrap remove remove_in. Qed.
Lemma remove_3 : In y (remove x s) -> In y s.
- Proof.
- unfold remove, In; simpl; rewrite Raw.remove_in; intuition.
- Qed.
+ Proof. wrap remove remove_in. Qed.
Lemma singleton_1 : In y (singleton x) -> E.eq x y.
- Proof. exact (Raw.singleton_1 x y). Qed.
+ Proof. exact (@singleton_1 x y). Qed.
Lemma singleton_2 : E.eq x y -> In y (singleton x).
- Proof. exact (Raw.singleton_2 x y). Qed.
+ Proof. exact (@singleton_2 x y). Qed.
Lemma union_1 : In x (union s s') -> In x s \/ In x s'.
- Proof.
- unfold union, In; simpl.
- destruct (Raw.union s s' (is_bst s) (is_avl s) (is_bst s') (is_avl s'))
- as (u,(b,(a,i))).
- simpl in *; rewrite i; auto.
- Qed.
-
+ Proof. wrap union union_in. Qed.
Lemma union_2 : In x s -> In x (union s s').
- Proof.
- unfold union, In; simpl.
- destruct (Raw.union s s' (is_bst s) (is_avl s) (is_bst s') (is_avl s'))
- as (u,(b,(a,i))).
- simpl in *; rewrite i; auto.
- Qed.
-
+ Proof. wrap union union_in. Qed.
Lemma union_3 : In x s' -> In x (union s s').
- Proof.
- unfold union, In; simpl.
- destruct (Raw.union s s' (is_bst s) (is_avl s) (is_bst s') (is_avl s'))
- as (u,(b,(a,i))).
- simpl in *; rewrite i; auto.
- Qed.
-
- Lemma union'_1 : In x (union' s s') -> In x s \/ In x s'.
- Proof.
- unfold union', In; simpl; rewrite Raw.union'_in; intuition.
- Qed.
-
- Lemma union'_2 : In x s -> In x (union' s s').
- Proof.
- unfold union', In; simpl; rewrite Raw.union'_in; intuition.
- Qed.
-
- Lemma union'_3 : In x s' -> In x (union' s s').
- Proof.
- unfold union', In; simpl; rewrite Raw.union'_in; intuition.
- Qed.
+ Proof. wrap union union_in. Qed.
Lemma inter_1 : In x (inter s s') -> In x s.
- Proof.
- unfold inter, In; simpl; rewrite Raw.inter_in; intuition.
- Qed.
-
+ Proof. wrap inter inter_in. Qed.
Lemma inter_2 : In x (inter s s') -> In x s'.
- Proof.
- unfold inter, In; simpl; rewrite Raw.inter_in; intuition.
- Qed.
-
+ Proof. wrap inter inter_in. Qed.
Lemma inter_3 : In x s -> In x s' -> In x (inter s s').
- Proof.
- unfold inter, In; simpl; rewrite Raw.inter_in; intuition.
- Qed.
+ Proof. wrap inter inter_in. Qed.
Lemma diff_1 : In x (diff s s') -> In x s.
- Proof.
- unfold diff, In; simpl; rewrite Raw.diff_in; intuition.
- Qed.
-
+ Proof. wrap diff diff_in. Qed.
Lemma diff_2 : In x (diff s s') -> ~ In x s'.
- Proof.
- unfold diff, In; simpl; rewrite Raw.diff_in; intuition.
- Qed.
-
+ Proof. wrap diff diff_in. Qed.
Lemma diff_3 : In x s -> ~ In x s' -> In x (diff s s').
- Proof.
- unfold diff, In; simpl; rewrite Raw.diff_in; intuition.
- Qed.
+ Proof. wrap diff diff_in. Qed.
- Lemma fold_1 : forall (A : Set) (i : A) (f : elt -> A -> A),
- fold A f s i = fold_left (fun a e => f e a) (elements s) i.
- Proof.
- unfold fold, elements; intros; apply Raw.fold_1; auto.
- Qed.
+ Lemma fold_1 : 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. unfold fold, elements; intros; apply fold_1; auto. Qed.
Lemma cardinal_1 : cardinal s = length (elements s).
Proof.
- unfold cardinal, elements; intros; apply Raw.cardinal_elements_1; auto.
+ unfold cardinal, elements; intros; apply elements_cardinal; auto.
Qed.
Section Filter.
Variable f : elt -> bool.
Lemma filter_1 : compat_bool E.eq f -> In x (filter f s) -> In x s.
- Proof.
- intro; unfold filter, In; simpl; rewrite Raw.filter_in; intuition.
- Qed.
-
+ Proof. intro. wrap filter filter_in. Qed.
Lemma filter_2 : compat_bool E.eq f -> In x (filter f s) -> f x = true.
- Proof.
- intro; unfold filter, In; simpl; rewrite Raw.filter_in; intuition.
- Qed.
-
+ Proof. intro. wrap filter filter_in. Qed.
Lemma filter_3 : compat_bool E.eq f -> In x s -> f x = true -> In x (filter f s).
- Proof.
- intro; unfold filter, In; simpl; rewrite Raw.filter_in; intuition.
- Qed.
+ Proof. intro. wrap filter filter_in. Qed.
Lemma for_all_1 : compat_bool E.eq f -> For_all (fun x => f x = true) s -> for_all f s = true.
- Proof. exact (Raw.for_all_1 f s). Qed.
+ Proof. exact (@for_all_1 f s). Qed.
Lemma for_all_2 : compat_bool E.eq f -> for_all f s = true -> For_all (fun x => f x = true) s.
- Proof. exact (Raw.for_all_2 f s). Qed.
+ Proof. exact (@for_all_2 f s). Qed.
Lemma exists_1 : compat_bool E.eq f -> Exists (fun x => f x = true) s -> exists_ f s = true.
- Proof. exact (Raw.exists_1 f s). Qed.
+ Proof. exact (@exists_1 f s). Qed.
Lemma exists_2 : compat_bool E.eq f -> exists_ f s = true -> Exists (fun x => f x = true) s.
- Proof. exact (Raw.exists_2 f s). Qed.
+ Proof. exact (@exists_2 f s). Qed.
Lemma partition_1 : compat_bool E.eq f ->
Equal (fst (partition f s)) (filter f s).
Proof.
unfold partition, filter, Equal, In; simpl ;intros H a.
- rewrite Raw.partition_in_1; auto.
- rewrite Raw.filter_in; intuition.
+ rewrite partition_in_1, filter_in; intuition.
Qed.
Lemma partition_2 : compat_bool E.eq f ->
Equal (snd (partition f s)) (filter (fun x => negb (f x)) s).
Proof.
unfold partition, filter, Equal, In; simpl ;intros H a.
- rewrite Raw.partition_in_2; auto.
- rewrite Raw.filter_in; intuition.
- red; intros.
- f_equal; auto.
- destruct (f a); auto.
+ rewrite partition_in_2, filter_in; intuition.
+ rewrite H2; auto.
destruct (f a); auto.
+ red; intros; f_equal.
+ rewrite (H _ _ H0); auto.
Qed.
End Filter.
Lemma elements_1 : In x s -> InA E.eq x (elements s).
- Proof.
- unfold elements, In; rewrite Raw.elements_in; auto.
- Qed.
-
+ Proof. wrap elements elements_in. Qed.
Lemma elements_2 : InA E.eq x (elements s) -> In x s.
- Proof.
- unfold elements, In; rewrite Raw.elements_in; auto.
- Qed.
-
+ Proof. wrap elements elements_in. Qed.
Lemma elements_3 : sort E.lt (elements s).
- Proof. exact (Raw.elements_sort _ (is_bst s)). Qed.
+ Proof. exact (elements_sort (is_bst s)). Qed.
+ Lemma elements_3w : NoDupA E.eq (elements s).
+ Proof. exact (elements_nodup (is_bst s)). Qed.
Lemma min_elt_1 : min_elt s = Some x -> In x s.
- Proof. exact (Raw.min_elt_1 s x). Qed.
+ Proof. exact (@min_elt_1 s x). Qed.
Lemma min_elt_2 : min_elt s = Some x -> In y s -> ~ E.lt y x.
- Proof. exact (Raw.min_elt_2 s x y (is_bst s)). Qed.
+ Proof. exact (@min_elt_2 s x y (is_bst s)). Qed.
Lemma min_elt_3 : min_elt s = None -> Empty s.
- Proof. exact (Raw.min_elt_3 s). Qed.
+ Proof. exact (@min_elt_3 s). Qed.
Lemma max_elt_1 : max_elt s = Some x -> In x s.
- Proof. exact (Raw.max_elt_1 s x). Qed.
+ Proof. exact (@max_elt_1 s x). Qed.
Lemma max_elt_2 : max_elt s = Some x -> In y s -> ~ E.lt x y.
- Proof. exact (Raw.max_elt_2 s x y (is_bst s)). Qed.
+ Proof. exact (@max_elt_2 s x y (is_bst s)). Qed.
Lemma max_elt_3 : max_elt s = None -> Empty s.
- Proof. exact (Raw.max_elt_3 s). Qed.
+ Proof. exact (@max_elt_3 s). Qed.
Lemma choose_1 : choose s = Some x -> In x s.
- Proof. exact (Raw.choose_1 s x). Qed.
+ Proof. exact (@choose_1 s x). Qed.
Lemma choose_2 : choose s = None -> Empty s.
- Proof. exact (Raw.choose_2 s). Qed.
+ Proof. exact (@choose_2 s). Qed.
+ Lemma choose_3 : choose s = Some x -> choose s' = Some y ->
+ Equal s s' -> E.eq x y.
+ Proof. exact (@choose_3 _ _ (is_bst s) (is_bst s') x y). Qed.
Lemma eq_refl : eq s s.
- Proof. exact (Raw.eq_refl s). Qed.
+ Proof. exact (eq_refl s). Qed.
Lemma eq_sym : eq s s' -> eq s' s.
- Proof. exact (Raw.eq_sym s s'). Qed.
+ Proof. exact (@eq_sym s s'). Qed.
Lemma eq_trans : eq s s' -> eq s' s'' -> eq s s''.
- Proof. exact (Raw.eq_trans s s' s''). Qed.
+ Proof. exact (@eq_trans s s' s''). Qed.
Lemma lt_trans : lt s s' -> lt s' s'' -> lt s s''.
- Proof. exact (Raw.lt_trans s s' s''). Qed.
+ Proof. exact (@lt_trans s s' s''). Qed.
Lemma lt_not_eq : lt s s' -> ~eq s s'.
- Proof. exact (Raw.lt_not_eq _ _ (is_bst s) (is_bst s')). Qed.
+ Proof. exact (@lt_not_eq _ _ (is_bst s) (is_bst s')). Qed.
End Specs.
End IntMake.
@@ -2897,4 +2053,3 @@ End IntMake.
Module Make (X: OrderedType) <: S with Module E := X
:=IntMake(Z_as_Int)(X).
-
diff --git a/theories/FSets/FSetBridge.v b/theories/FSets/FSetBridge.v
index 08985cfc..0622451f 100644
--- a/theories/FSets/FSetBridge.v
+++ b/theories/FSets/FSetBridge.v
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: FSetBridge.v 8834 2006-05-20 00:41:35Z letouzey $ *)
+(* $Id: FSetBridge.v 10601 2008-02-28 00:20:33Z letouzey $ *)
(** * Finite sets library *)
@@ -27,7 +27,7 @@ Module DepOfNodep (M: S) <: Sdep with Module E := M.E.
Definition empty : {s : t | Empty s}.
Proof.
- exists empty; auto.
+ exists empty; auto with set.
Qed.
Definition is_empty : forall s : t, {Empty s} + {~ Empty s}.
@@ -66,12 +66,12 @@ Module DepOfNodep (M: S) <: Sdep with Module E := M.E.
{s' : t | forall y : elt, In y s' <-> ~ E.eq x y /\ In y s}.
Proof.
intros; exists (remove x s); intuition.
- absurd (In x (remove x s)); auto.
+ absurd (In x (remove x s)); auto with set.
apply In_1 with y; auto.
elim (ME.eq_dec x y); intros; auto.
- absurd (In x (remove x s)); auto.
+ absurd (In x (remove x s)); auto with set.
apply In_1 with y; auto.
- eauto.
+ eauto with set.
Qed.
Definition union :
@@ -83,14 +83,14 @@ Module DepOfNodep (M: S) <: Sdep with Module E := M.E.
Definition inter :
forall s s' : t, {s'' : t | forall x : elt, In x s'' <-> In x s /\ In x s'}.
Proof.
- intros; exists (inter s s'); intuition; eauto.
+ intros; exists (inter s s'); intuition; eauto with set.
Qed.
Definition diff :
forall s s' : t, {s'' : t | forall x : elt, In x s'' <-> In x s /\ ~ In x s'}.
Proof.
- intros; exists (diff s s'); intuition; eauto.
- absurd (In x s'); eauto.
+ intros; exists (diff s s'); intuition; eauto with set.
+ absurd (In x s'); eauto with set.
Qed.
Definition equal : forall s s' : t, {Equal s s'} + {~ Equal s s'}.
@@ -115,7 +115,7 @@ Module DepOfNodep (M: S) <: Sdep with Module E := M.E.
Defined.
Definition fold :
- forall (A : Set) (f : elt -> A -> A) (s : t) (i : A),
+ forall (A : Type) (f : elt -> A -> A) (s : t) (i : A),
{r : A | let (l,_) := elements s in
r = fold_left (fun a e => f e a) l i}.
Proof.
@@ -150,7 +150,7 @@ Module DepOfNodep (M: S) <: Sdep with Module E := M.E.
exists (filter (fdec Pdec) s).
intro H; assert (compat_bool E.eq (fdec Pdec)); auto.
intuition.
- eauto.
+ eauto with set.
generalize (filter_2 H0 H1).
unfold fdec in |- *.
case (Pdec x); intuition.
@@ -226,7 +226,7 @@ Module DepOfNodep (M: S) <: Sdep with Module E := M.E.
generalize H4; unfold For_all, Equal in |- *; intuition.
elim (H0 x); intros.
assert (fdec Pdec x = true).
- eauto.
+ eapply filter_2; eauto with set.
generalize H8; unfold fdec in |- *; case (Pdec x); intuition.
inversion H9.
generalize H; unfold For_all, Equal in |- *; intuition.
@@ -238,8 +238,8 @@ Module DepOfNodep (M: S) <: Sdep with Module E := M.E.
set (b := fdec Pdec x) in *; generalize (refl_equal b);
pattern b at -1 in |- *; case b; unfold b in |- *;
[ left | right ].
- elim (H4 x); intros _ B; apply B; auto.
- elim (H x); intros _ B; apply B; auto.
+ elim (H4 x); intros _ B; apply B; auto with set.
+ elim (H x); intros _ B; apply B; auto with set.
apply filter_3; auto.
rewrite H5; auto.
eapply (filter_1 (s:=s) (x:=x) H2); elim (H4 x); intros B _; apply B;
@@ -247,12 +247,63 @@ Module DepOfNodep (M: S) <: Sdep with Module E := M.E.
eapply (filter_1 (s:=s) (x:=x) H3); elim (H x); intros B _; apply B; auto.
Qed.
+ Definition choose_aux: forall s : t,
+ { x : elt | M.choose s = Some x } + { M.choose s = None }.
+ Proof.
+ intros.
+ destruct (M.choose s); [left | right]; auto.
+ exists e; auto.
+ Qed.
+
Definition choose : forall s : t, {x : elt | In x s} + {Empty s}.
- Proof.
- intros.
- generalize (choose_1 (s:=s)) (choose_2 (s:=s)).
- case (choose s); [ left | right ]; auto.
- exists e; auto.
+ Proof.
+ intros; destruct (choose_aux s) as [(x,Hx)|H].
+ left; exists x; apply choose_1; auto.
+ right; apply choose_2; auto.
+ Defined.
+
+ Lemma choose_ok1 :
+ forall s x, M.choose s = Some x <-> exists H:In x s,
+ choose s = inleft _ (exist (fun x => In x s) x H).
+ Proof.
+ intros s x.
+ unfold choose; split; intros.
+ destruct (choose_aux s) as [(y,Hy)|H']; try congruence.
+ replace x with y in * by congruence.
+ exists (choose_1 Hy); auto.
+ destruct H.
+ destruct (choose_aux s) as [(y,Hy)|H']; congruence.
+ Qed.
+
+ Lemma choose_ok2 :
+ forall s, M.choose s = None <-> exists H:Empty s,
+ choose s = inright _ H.
+ Proof.
+ intros s.
+ unfold choose; split; intros.
+ destruct (choose_aux s) as [(y,Hy)|H']; try congruence.
+ exists (choose_2 H'); auto.
+ destruct H.
+ destruct (choose_aux s) as [(y,Hy)|H']; congruence.
+ Qed.
+
+ Lemma choose_equal : forall s s', Equal s s' ->
+ match choose s, choose s' with
+ | inleft (exist x _), inleft (exist x' _) => E.eq x x'
+ | inright _, inright _ => True
+ | _, _ => False
+ end.
+ Proof.
+ intros.
+ generalize (@M.choose_1 s)(@M.choose_2 s)
+ (@M.choose_1 s')(@M.choose_2 s')(@M.choose_3 s s')
+ (choose_ok1 s)(choose_ok2 s)(choose_ok1 s')(choose_ok2 s').
+ destruct (choose s) as [(x,Hx)|Hx]; destruct (choose s') as [(x',Hx')|Hx']; auto; intros.
+ apply H4; auto.
+ rewrite H5; exists Hx; auto.
+ rewrite H7; exists Hx'; auto.
+ apply Hx' with x; unfold Equal in H; rewrite <-H; auto.
+ apply Hx with x'; unfold Equal in H; rewrite H; auto.
Qed.
Definition min_elt :
@@ -391,6 +442,15 @@ Module NodepOfDep (M: Sdep) <: S with Module E := M.E.
intro s; unfold choose in |- *; case (M.choose s); auto.
simple destruct s0; intros; discriminate H.
Qed.
+
+ Lemma choose_3 : forall s s' x x',
+ choose s = Some x -> choose s' = Some x' -> Equal s s' -> E.eq x x'.
+ Proof.
+ unfold choose; intros.
+ generalize (M.choose_equal H1); clear H1.
+ destruct (M.choose s) as [(?,?)|?]; destruct (M.choose s') as [(?,?)|?];
+ simpl; auto; congruence.
+ Qed.
Definition elements (s : t) : list elt := let (l, _) := elements s in l.
@@ -408,6 +468,10 @@ Module NodepOfDep (M: Sdep) <: S with Module E := M.E.
Proof.
intros; unfold elements in |- *; case (M.elements s); firstorder.
Qed.
+ Hint Resolve elements_3.
+
+ Lemma elements_3w : forall s : t, NoDupA E.eq (elements s).
+ Proof. auto. Qed.
Definition min_elt (s : t) : option elt :=
match min_elt s with
@@ -578,11 +642,11 @@ Module NodepOfDep (M: Sdep) <: S with Module E := M.E.
destruct (M.elements s); auto.
Qed.
- Definition fold (B : Set) (f : elt -> B -> B) (i : t)
+ Definition fold (B : Type) (f : elt -> B -> B) (i : t)
(s : B) : B := let (fold, _) := fold f i s in fold.
Lemma fold_1 :
- forall (s : t) (A : Set) (i : A) (f : elt -> A -> A),
+ forall (s : t) (A : Type) (i : A) (f : elt -> A -> A),
fold f s i = fold_left (fun a e => f e a) (elements s) i.
Proof.
intros; unfold fold in |- *; case (M.fold f s i); unfold elements in *;
diff --git a/theories/FSets/FSetDecide.v b/theories/FSets/FSetDecide.v
new file mode 100644
index 00000000..0639c1f1
--- /dev/null
+++ b/theories/FSets/FSetDecide.v
@@ -0,0 +1,841 @@
+(***********************************************************************)
+(* 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: FSetDecide.v 11064 2008-06-06 17:00:52Z letouzey $ *)
+
+(**************************************************************)
+(* FSetDecide.v *)
+(* *)
+(* Author: Aaron Bohannon *)
+(**************************************************************)
+
+(** This file implements a decision procedure for a certain
+ class of propositions involving finite sets. *)
+
+Require Import Decidable DecidableTypeEx FSetFacts.
+
+(** First, a version for Weak Sets *)
+
+Module WDecide (E : DecidableType)(Import M : WSfun E).
+ Module F := FSetFacts.WFacts 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 FSet-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 FSetLogicalFacts.
+ 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 FSetLogicalFacts.
+ Import FSetLogicalFacts.
+
+ (** * 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 FSetDecideAuxiliary.
+
+ (** ** Generic Tactics
+ We begin by defining a few generic, useful tactics. *)
+
+ (** [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-FSet-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 FSet_elt_Prop : Prop -> Prop :=
+ | eq_Prop : forall (S : Set) (x y : S),
+ FSet_elt_Prop (x = y)
+ | eq_elt_prop : forall x y,
+ FSet_elt_Prop (E.eq x y)
+ | In_elt_prop : forall x s,
+ FSet_elt_Prop (In x s)
+ | True_elt_prop :
+ FSet_elt_Prop True
+ | False_elt_prop :
+ FSet_elt_Prop False
+ | conj_elt_prop : forall P Q,
+ FSet_elt_Prop P ->
+ FSet_elt_Prop Q ->
+ FSet_elt_Prop (P /\ Q)
+ | disj_elt_prop : forall P Q,
+ FSet_elt_Prop P ->
+ FSet_elt_Prop Q ->
+ FSet_elt_Prop (P \/ Q)
+ | impl_elt_prop : forall P Q,
+ FSet_elt_Prop P ->
+ FSet_elt_Prop Q ->
+ FSet_elt_Prop (P -> Q)
+ | not_elt_prop : forall P,
+ FSet_elt_Prop P ->
+ FSet_elt_Prop (~ P).
+
+ Inductive FSet_Prop : Prop -> Prop :=
+ | elt_FSet_Prop : forall P,
+ FSet_elt_Prop P ->
+ FSet_Prop P
+ | Empty_FSet_Prop : forall s,
+ FSet_Prop (Empty s)
+ | Subset_FSet_Prop : forall s1 s2,
+ FSet_Prop (Subset s1 s2)
+ | Equal_FSet_Prop : forall s1 s2,
+ FSet_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 FSet_elt_Prop FSet_Prop : FSet_Prop.
+ Ltac discard_nonFSet :=
+ decompose records;
+ repeat (
+ match goal with
+ | H : ?P |- _ =>
+ if prop (FSet_Prop P) holds by
+ (auto 100 with FSet_Prop)
+ then fail
+ else clear H
+ end).
+
+ (** ** Turning Set Operators into Propositional Connectives
+ The lemmas from [FSetFacts] 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 FSet 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 [FSet_decidability] will be given to
+ the [push_neg] tactic from the module [Negation]. *)
+ Hint Resolve dec_In dec_eq : FSet_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 )
+ 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 substFSet :=
+ 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_FSet_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 substFSet; 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_FSet_hypotheses;
+ autorewrite with set_simpl in *;
+ push not in * using FSet_decidability;
+ substFSet;
+ assert_decidability;
+ auto using E.eq_refl;
+ (intuition fsetdec_rec) ||
+ fail 1
+ "because the goal is beyond the scope of this tactic".
+
+ End FSetDecideAuxiliary.
+ Import FSetDecideAuxiliary.
+
+ (** * 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;
+ (** Now we decompose conjunctions, which will allow the
+ [discard_nonFSet] and [assert_decidability] tactics to
+ do a much better job. *)
+ decompose records;
+ discard_nonFSet;
+ (** 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 [FSet_elt_Prop], noting
+ that any [FSet_elt_Prop] is decidable. *)
+ pull not using FSet_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 (FSet_elt_Prop P) holds by
+ (auto 100 with FSet_Prop)
+ then (contradict H; fsetdec_body)
+ else fsetdec_body
+ | |- _ =>
+ fsetdec_body
+ end.
+
+ (** * Examples *)
+
+ Module FSetDecideTestCases.
+
+ 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_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.
+
+ End FSetDecideTestCases.
+
+End WDecide.
+
+Require Import FSetInterface.
+
+(** Now comes a special version dedicated to full sets. For this
+ one, only one argument [(M:S)] is necessary. *)
+
+Module Decide (M : S).
+ Module D:=OT_as_DT M.E.
+ Module WD := WDecide D M.
+ Ltac fsetdec := WD.fsetdec.
+End Decide. \ No newline at end of file
diff --git a/theories/FSets/FSetEqProperties.v b/theories/FSets/FSetEqProperties.v
index d7062d5a..a397cc28 100644
--- a/theories/FSets/FSetEqProperties.v
+++ b/theories/FSets/FSetEqProperties.v
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: FSetEqProperties.v 8853 2006-05-23 18:17:38Z herbelin $ *)
+(* $Id: FSetEqProperties.v 11064 2008-06-06 17:00:52Z letouzey $ *)
(** * Finite sets library *)
@@ -17,21 +17,12 @@
[mem x s=true] instead of [In x s],
[equal s s'=true] instead of [Equal s s'], etc. *)
+Require Import FSetProperties Zerob Sumbool Omega DecidableTypeEx.
-Require Import FSetProperties.
-Require Import Zerob.
-Require Import Sumbool.
-Require Import Omega.
-
-Module EqProperties (M:S).
+Module WEqProperties (Import E:DecidableType)(M:WSfun E).
+Module Import MP := WProperties E M.
+Import FM Dec.F.
Import M.
-Import Logic. (* to unmask [eq] *)
-Import Peano. (* to unmask [lt] *)
-
-Module ME := OrderedTypeFacts E.
-Module MP := Properties M.
-Import MP.
-Import MP.FM.
Definition Add := MP.Add.
@@ -76,7 +67,7 @@ Qed.
Lemma empty_mem: mem x empty=false.
Proof.
-rewrite <- not_mem_iff; auto.
+rewrite <- not_mem_iff; auto with set.
Qed.
Lemma is_empty_equal_empty: is_empty s = equal s empty.
@@ -88,17 +79,17 @@ Qed.
Lemma choose_mem_1: choose s=Some x -> mem x s=true.
Proof.
-auto.
+auto with set.
Qed.
Lemma choose_mem_2: choose s=None -> is_empty s=true.
Proof.
-auto.
+auto with set.
Qed.
Lemma add_mem_1: mem x (add x s)=true.
Proof.
-auto.
+auto with set.
Qed.
Lemma add_mem_2: ~E.eq x y -> mem y (add x s)=mem y s.
@@ -108,7 +99,7 @@ Qed.
Lemma remove_mem_1: mem x (remove x s)=false.
Proof.
-rewrite <- not_mem_iff; auto.
+rewrite <- not_mem_iff; auto with set.
Qed.
Lemma remove_mem_2: ~E.eq x y -> mem y (remove x s)=mem y s.
@@ -216,7 +207,7 @@ Proof.
intros.
generalize (@choose_1 s) (@choose_2 s).
destruct (choose s);intros.
-exists e;auto.
+exists e;auto with set.
generalize (H1 (refl_equal None)); clear H1.
intros; rewrite (is_empty_1 H1) in H; discriminate.
Qed.
@@ -233,7 +224,7 @@ Qed.
Lemma add_mem_3:
mem y s=true -> mem y (add x s)=true.
Proof.
-auto.
+auto with set.
Qed.
Lemma add_equal:
@@ -260,7 +251,7 @@ 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.
+intros; apply equal_1; apply add_remove; auto with set.
Qed.
Lemma remove_add:
@@ -275,9 +266,9 @@ Qed.
Lemma is_empty_cardinal: is_empty s = zerob (cardinal s).
Proof.
intros; apply bool_1; split; intros.
-rewrite cardinal_1; simpl; auto.
+rewrite MP.cardinal_1; simpl; auto with set.
assert (cardinal s = 0) by (apply zerob_true_elim; auto).
-auto.
+auto with set.
Qed.
(** Properties of [singleton] *)
@@ -290,12 +281,12 @@ Qed.
Lemma singleton_mem_2: ~E.eq x y -> mem y (singleton x)=false.
Proof.
intros; rewrite singleton_b.
-unfold ME.eqb; destruct (ME.eq_dec x y); intuition.
+unfold eqb; destruct (eq_dec x y); intuition.
Qed.
Lemma singleton_mem_3: mem y (singleton x)=true -> E.eq x y.
Proof.
-auto.
+intros; apply singleton_1; auto with set.
Qed.
(** Properties of [union] *)
@@ -358,7 +349,7 @@ 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.
+intros; apply subset_1; apply union_subset_3; auto with set.
Qed.
(** Properties of [inter] *)
@@ -433,7 +424,7 @@ 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.
+intros; apply subset_1; apply inter_subset_3; auto with set.
Qed.
(** Properties of [diff] *)
@@ -478,45 +469,37 @@ Hint Resolve equal_mem_1 subset_mem_1 choose_mem_1
add_mem_3 add_equal remove_mem_3 remove_equal : set.
-(** General recursion principes based on [cardinal] *)
+(** General recursion principle *)
-Lemma cardinal_set_rec: forall (P:t->Type),
+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 n s, cardinal s=n -> P s.
+ P empty -> forall s, P s.
Proof.
intros.
-apply cardinal_induction with n; auto; 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 (H1 a); tauto.
+apply equal_1; intro a; rewrite add_iff; rewrite (H0 a); tauto.
apply X0; auto with set; apply mem_3; auto.
Qed.
-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 cardinal_set_rec with (cardinal s);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.
+ ~(In x s/\In x s') <-> mem x s && mem x s'=false.
Proof.
-intros; do 2 rewrite not_mem_iff.
+intros; do 2 rewrite mem_iff.
destruct (mem x s); destruct (mem x s'); intuition.
Qed.
Section Fold.
-Variables (A:Set)(eqA:A->A->Prop)(st:Setoid_Theory _ eqA).
+Variables (A:Type)(eqA:A->A->Prop)(st:Setoid_Theory _ eqA).
Variables (f:elt->A->A)(Comp:compat_op E.eq eqA f)(Ass:transpose eqA f).
Variables (i:A).
Variables (s s':t)(x:elt).
-Lemma fold_empty: eqA (fold f empty i) i.
+Lemma fold_empty: (fold f empty i) = i.
Proof.
apply fold_empty; auto.
Qed.
@@ -524,7 +507,7 @@ 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.
+intros; apply fold_equal with (eqA:=eqA); auto with set.
Qed.
Lemma fold_add:
@@ -537,13 +520,13 @@ 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.
+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.
+intros; apply remove_fold_1 with (eqA:=eqA); auto with set.
Qed.
Lemma remove_fold_2:
@@ -581,13 +564,13 @@ 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.
+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.
-auto with set.
+intros; apply Equal_cardinal; apply equal_2; auto with set.
Qed.
Lemma union_cardinal:
@@ -601,7 +584,7 @@ Qed.
Lemma subset_cardinal:
forall s s', subset s s'=true -> cardinal s<=cardinal s'.
Proof.
-intros; apply subset_cardinal; auto.
+intros; apply subset_cardinal; auto with set.
Qed.
Section Bool.
@@ -644,7 +627,7 @@ 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.
+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.
@@ -656,13 +639,30 @@ Qed.
Lemma partition_filter_1:
forall s, equal (fst (partition f s)) (filter f s)=true.
Proof.
-auto.
+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.
+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.
+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,
@@ -837,6 +837,8 @@ 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 := (compat_op E.eq (@Logic.eq _)).
+Notation transposeL := (transpose (@Logic.eq _)).
Lemma sum_plus :
forall f g, compat_nat E.eq f -> compat_nat E.eq g ->
@@ -844,12 +846,12 @@ Lemma sum_plus :
Proof.
unfold sum.
intros f g Hf Hg.
-assert (fc : compat_op E.eq (@eq _) (fun x:elt =>plus (f x))). auto.
-assert (ft : transpose (@eq _) (fun x:elt =>plus (f x))). red; intros; omega.
-assert (gc : compat_op E.eq (@eq _) (fun x:elt => plus (g x))). auto.
-assert (gt : transpose (@eq _) (fun x:elt =>plus (g x))). red; intros; omega.
-assert (fgc : compat_op E.eq (@eq _) (fun x:elt =>plus ((f x)+(g x)))). auto.
-assert (fgt : transpose (@eq _) (fun x:elt=>plus ((f x)+(g x)))). red; intros; omega.
+assert (fc : compat_opL (fun x:elt =>plus (f x))). auto.
+assert (ft : transposeL (fun x:elt =>plus (f x))). red; intros; omega.
+assert (gc : compat_opL (fun x:elt => plus (g x))). auto.
+assert (gt : transposeL (fun x:elt =>plus (g x))). red; intros; omega.
+assert (fgc : compat_opL (fun x:elt =>plus ((f x)+(g x)))). auto.
+assert (fgt : transposeL (fun x:elt=>plus ((f x)+(g x)))). red; intros; omega.
assert (st := gen_st nat).
intros s;pattern s; apply set_rec.
intros.
@@ -858,7 +860,7 @@ rewrite <- (fold_equal _ _ st _ gc gt 0 _ _ H).
rewrite <- (fold_equal _ _ st _ fgc fgt 0 _ _ H); auto.
intros; do 3 (rewrite (fold_add _ _ st);auto).
rewrite H0;simpl;omega.
-intros; do 3 rewrite (fold_empty _ _ st);auto.
+do 3 rewrite fold_empty;auto.
Qed.
Lemma sum_filter : forall f, (compat_bool E.eq f) ->
@@ -866,11 +868,11 @@ Lemma sum_filter : forall f, (compat_bool E.eq f) ->
Proof.
unfold sum; intros f Hf.
assert (st := gen_st nat).
-assert (cc : compat_op E.eq (@eq _) (fun x => plus (if f x then 1 else 0))).
- unfold compat_op; intros.
+assert (cc : compat_opL (fun x => plus (if f x then 1 else 0))).
+ red; intros.
rewrite (Hf x x' H); auto.
-assert (ct : transpose (@eq _) (fun x => plus (if f x then 1 else 0))).
- unfold transpose; intros; omega.
+assert (ct : transposeL (fun x => plus (if f x then 1 else 0))).
+ red; intros; omega.
intros s;pattern s; apply set_rec.
intros.
change elt with E.t.
@@ -883,14 +885,14 @@ assert (~ In x (filter f s0)).
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 _ _ st);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:Set)(eqA:A->A->Prop)(st:(Setoid_Theory _ eqA))
+ forall (A:Type)(eqA:A->A->Prop)(st:Setoid_Theory _ eqA)
(f g:elt->A->A),
(compat_op E.eq eqA f) -> (transpose eqA f) ->
(compat_op E.eq eqA g) -> (transpose eqA g) ->
@@ -903,26 +905,35 @@ trans_st (fold f s0 i).
apply fold_equal with (eqA:=eqA); auto.
rewrite equal_sym; auto.
trans_st (fold g s0 i).
-apply H0; intros; apply H1; auto.
-elim (equal_2 H x); auto; intros.
-apply fold_equal with (eqA:=eqA); auto.
+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.
trans_st (f x (fold f s0 i)).
-apply fold_add with (eqA:=eqA); auto.
-trans_st (g x (fold f s0 i)).
-trans_st (g x (fold g s0 i)).
+apply fold_add with (eqA:=eqA); auto with set.
+trans_st (g x (fold f s0 i)); auto with set.
+trans_st (g x (fold g s0 i)); auto with set.
sym_st; apply fold_add with (eqA:=eqA); auto.
-trans_st i; [idtac | sym_st ]; apply fold_empty; auto.
+do 2 rewrite fold_empty; refl_st.
Qed.
Lemma sum_compat :
forall f g, compat_nat E.eq f -> compat_nat E.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 _ (@eq nat)); auto.
-unfold transpose; intros; omega.
-unfold transpose; intros; omega.
+unfold sum; apply (fold_compat _ (@Logic.eq nat)); auto.
+red; intros; omega.
+red; intros; omega.
Qed.
End Sum.
-End EqProperties.
+End WEqProperties.
+
+
+(** Now comes a special version dedicated to full sets. For this
+ one, only one argument [(M:S)] is necessary. *)
+
+Module EqProperties (M:S).
+ Module D := OT_as_DT M.E.
+ Include WEqProperties D M.
+End EqProperties.
diff --git a/theories/FSets/FSetFacts.v b/theories/FSets/FSetFacts.v
index aa57f066..b4b834b1 100644
--- a/theories/FSets/FSetFacts.v
+++ b/theories/FSets/FSetFacts.v
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: FSetFacts.v 8882 2006-05-31 21:55:30Z letouzey $ *)
+(* $Id: FSetFacts.v 10765 2008-04-08 16:15:23Z msozeau $ *)
(** * Finite sets library *)
@@ -16,16 +16,19 @@
Moreover, we prove that [E.Eq] and [Equal] are setoid equalities.
*)
+Require Import DecidableTypeEx.
Require Export FSetInterface.
Set Implicit Arguments.
Unset Strict Implicit.
-Module Facts (M: S).
-Module ME := OrderedTypeFacts M.E.
-Import ME.
-Import M.
-Import Logic. (* to unmask [eq] *)
-Import Peano. (* to unmask [lt] *)
+(** First, a functor for Weak Sets. Since the signature [WS] includes
+ an EqualityType and not a stronger DecidableType, this functor
+ should take two arguments in order to compensate this. *)
+
+Module WFacts (Import E : DecidableType)(Import M : WSfun E).
+
+Notation eq_dec := E.eq_dec.
+Definition eqb x y := if eq_dec x y then true else false.
(** * Specifications written using equivalences *)
@@ -259,6 +262,8 @@ symmetry.
rewrite H0; intros.
destruct H1 as (_,H1).
apply H1; auto.
+rewrite H2.
+rewrite InA_alt; eauto.
Qed.
Lemma exists_b : compat_bool E.eq f ->
@@ -271,7 +276,8 @@ 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; auto.
+exists a; split; auto.
+rewrite H2; rewrite InA_alt; eauto.
symmetry.
rewrite H0.
destruct H1 as (_,H1).
@@ -289,17 +295,25 @@ End BoolSpec.
Definition E_ST : Setoid_Theory elt E.eq.
Proof.
-constructor; [apply E.eq_refl|apply E.eq_sym|apply E.eq_trans].
+constructor ; red; [apply E.eq_refl|apply E.eq_sym|apply E.eq_trans].
Qed.
-Add Setoid elt E.eq E_ST as EltSetoid.
-
Definition Equal_ST : Setoid_Theory t Equal.
Proof.
-constructor; [apply eq_refl | apply eq_sym | apply eq_trans].
+constructor ; red; [apply eq_refl | apply eq_sym | apply eq_trans].
Qed.
-Add Setoid t Equal Equal_ST as EqualSetoid.
+Add Relation elt E.eq
+ reflexivity proved by E.eq_refl
+ symmetry proved by E.eq_sym
+ transitivity proved by E.eq_trans
+ as EltSetoid.
+
+Add Relation t Equal
+ reflexivity proved by eq_refl
+ symmetry proved by eq_sym
+ transitivity proved by eq_trans
+ as EqualSetoid.
Add Morphism In with signature E.eq ==> Equal ==> iff as In_m.
Proof.
@@ -325,7 +339,7 @@ exact (H1 (refl_equal true) _ Ha).
Qed.
Add Morphism Empty with signature Equal ==> iff as Empty_m.
-Proof.
+Proof.
intros; do 2 rewrite is_empty_iff; rewrite H; intuition.
Qed.
@@ -340,7 +354,9 @@ Qed.
Add Morphism singleton : singleton_m.
Proof.
unfold Equal; intros x y H a.
-do 2 rewrite singleton_iff; split; order.
+do 2 rewrite singleton_iff; split; intros.
+apply E.eq_trans with x; auto.
+apply E.eq_trans with y; auto.
Qed.
Add Morphism add : add_m.
@@ -396,6 +412,63 @@ rewrite H in H1; rewrite H0 in H1; intuition.
rewrite H in H1; rewrite H0 in H1; intuition.
Qed.
+
+(* [Subset] is a setoid order *)
+
+Lemma Subset_refl : forall s, s[<=]s.
+Proof. red; auto. Defined.
+
+Lemma Subset_trans : forall s s' s'', s[<=]s'->s'[<=]s''->s[<=]s''.
+Proof. unfold Subset; eauto. Defined.
+
+Add Relation t Subset
+ reflexivity proved by Subset_refl
+ transitivity proved by Subset_trans
+ as SubsetSetoid.
+(* NB: for the moment, it is important to use Defined and not Qed in
+ the two previous lemmas, in order to allow conversion of
+ SubsetSetoid coming from separate Facts modules. See bug #1738. *)
+
+Instance In_s_m : Morphism (E.eq ==> Subset ++> impl) In | 1.
+Proof.
+ simpl_relation. eauto with set.
+Qed.
+
+Add Morphism Empty with signature Subset --> impl as Empty_s_m.
+Proof.
+unfold Subset, Empty, impl; firstorder.
+Qed.
+
+Add Morphism add with signature E.eq ==> Subset ++> Subset as add_s_m.
+Proof.
+unfold Subset; intros x y H s s' H0 a.
+do 2 rewrite add_iff; rewrite H; intuition.
+Qed.
+
+Add Morphism remove with signature E.eq ==> Subset ++> Subset as remove_s_m.
+Proof.
+unfold Subset; intros x y H s s' H0 a.
+do 2 rewrite remove_iff; rewrite H; intuition.
+Qed.
+
+Add Morphism union with signature Subset ++> Subset ++> Subset as union_s_m.
+Proof.
+unfold Equal; intros s s' H s'' s''' H0 a.
+do 2 rewrite union_iff; intuition.
+Qed.
+
+Add Morphism inter with signature Subset ++> Subset ++> Subset as inter_s_m.
+Proof.
+unfold Equal; intros s s' H s'' s''' H0 a.
+do 2 rewrite inter_iff; intuition.
+Qed.
+
+Add Morphism diff with signature Subset ++> Subset --> Subset as diff_s_m.
+Proof.
+unfold Subset; intros s s' H s'' s''' H0 a.
+do 2 rewrite diff_iff; intuition.
+Qed.
+
(* [fold], [filter], [for_all], [exists_] and [partition] cannot be proved morphism
without additional hypothesis on [f]. For instance: *)
@@ -405,6 +478,12 @@ Proof.
unfold Equal; intros; repeat rewrite filter_iff; auto; rewrite H0; tauto.
Qed.
+Lemma filter_subset : forall f, compat_bool E.eq f ->
+ forall s s', s[<=]s' -> filter f s [<=] filter f s'.
+Proof.
+unfold Subset; intros; rewrite filter_iff in *; intuition.
+Qed.
+
(* For [elements], [min_elt], [max_elt] and [choose], we would need setoid
structures on [list elt] and [option elt]. *)
@@ -412,4 +491,15 @@ Qed.
Add Morphism cardinal ; cardinal_m.
*)
+End WFacts.
+
+
+(** Now comes a special version dedicated to full sets. For this
+ one, only one argument [(M:S)] is necessary. *)
+
+Module Facts (Import M:S).
+ Module D:=OT_as_DT M.E.
+ Include WFacts D M.
+
End Facts.
+
diff --git a/theories/FSets/FSetFullAVL.v b/theories/FSets/FSetFullAVL.v
new file mode 100644
index 00000000..1fc109f3
--- /dev/null
+++ b/theories/FSets/FSetFullAVL.v
@@ -0,0 +1,1125 @@
+(***********************************************************************)
+(* 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 *)
+(***********************************************************************)
+
+(* Finite sets library.
+ * Authors: Pierre Letouzey and Jean-Christophe Filliâtre
+ * Institution: LRI, CNRS UMR 8623 - Université Paris Sud
+ * 91405 Orsay, France *)
+
+(* $Id: FSetFullAVL.v 10739 2008-04-01 14:45:20Z herbelin $ *)
+
+(** * FSetFullAVL
+
+ This file contains some complements to [FSetAVL].
+
+ - Functor [AvlProofs] proves that trees of [FSetAVL] are not only
+ binary search trees, but moreover well-balanced ones. This is done
+ by proving that all operations preserve the balancing.
+
+ - Functor [OcamlOps] contains variants of [union], [subset],
+ [compare] and [equal] that are faithful to the original ocaml codes,
+ while the versions in FSetAVL have been adapted to perform only
+ structural recursive code.
+
+ - Finally, we pack the previous elements in a [Make] functor
+ similar to the one of [FSetAVL], but richer.
+*)
+
+Require Import Recdef FSetInterface FSetList ZArith Int FSetAVL.
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+
+Module AvlProofs (Import I:Int)(X:OrderedType).
+Module Import Raw := Raw I X.
+Import Raw.Proofs.
+Module Import II := MoreInt I.
+Open Local Scope pair_scope.
+Open Local Scope Int_scope.
+
+(** * AVL trees *)
+
+(** [avl s] : [s] is a properly balanced AVL tree,
+ i.e. for any node the heights of the two children
+ differ by at most 2 *)
+
+Inductive avl : tree -> Prop :=
+ | RBLeaf : avl Leaf
+ | RBNode : forall x l r h, avl l -> avl r ->
+ -(2) <= height l - height r <= 2 ->
+ h = max (height l) (height r) + 1 ->
+ avl (Node l x r h).
+
+(** * Automation and dedicated tactics *)
+
+Hint Constructors avl.
+
+(** A tactic for cleaning hypothesis after use of functional induction. *)
+
+Ltac clearf :=
+ match goal with
+ | H : (@Logic.eq (Compare _ _ _ _) _ _) |- _ => clear H; clearf
+ | H : (@Logic.eq (sumbool _ _) _ _) |- _ => clear H; clearf
+ | _ => idtac
+ end.
+
+(** Tactics about [avl] *)
+
+Lemma height_non_negative : forall s : tree, avl s -> height s >= 0.
+Proof.
+ induction s; simpl; intros; auto with zarith.
+ inv avl; intuition; omega_max.
+Qed.
+Implicit Arguments height_non_negative.
+
+(** When [H:avl r], typing [avl_nn H] or [avl_nn r] adds [height r>=0] *)
+
+Ltac avl_nn_hyp H :=
+ let nz := fresh "nz" in assert (nz := height_non_negative H).
+
+Ltac avl_nn h :=
+ let t := type of h in
+ match type of t with
+ | Prop => avl_nn_hyp h
+ | _ => match goal with H : avl h |- _ => avl_nn_hyp H end
+ end.
+
+(* Repeat the previous tactic.
+ Drawback: need to clear the [avl _] hyps ... Thank you Ltac *)
+
+Ltac avl_nns :=
+ match goal with
+ | H:avl _ |- _ => avl_nn_hyp H; clear H; avl_nns
+ | _ => idtac
+ end.
+
+(** Results about [height] *)
+
+Lemma height_0 : forall s, avl s -> height s = 0 -> s = Leaf.
+Proof.
+ destruct 1; intuition; simpl in *.
+ avl_nns; simpl in *; elimtype False; omega_max.
+Qed.
+
+(** * Results about [avl] *)
+
+Lemma avl_node :
+ forall x l r, avl l -> avl r ->
+ -(2) <= height l - height r <= 2 ->
+ avl (Node l x r (max (height l) (height r) + 1)).
+Proof.
+ intros; auto.
+Qed.
+Hint Resolve avl_node.
+
+
+(** empty *)
+
+Lemma empty_avl : avl empty.
+Proof.
+ auto.
+Qed.
+
+(** singleton *)
+
+Lemma singleton_avl : forall x : elt, avl (singleton x).
+Proof.
+ unfold singleton; intro.
+ constructor; auto; try red; simpl; omega_max.
+Qed.
+
+(** create *)
+
+Lemma create_avl :
+ forall l x r, avl l -> avl r -> -(2) <= height l - height r <= 2 ->
+ avl (create l x r).
+Proof.
+ unfold create; auto.
+Qed.
+
+Lemma create_height :
+ forall l x r, avl l -> avl r -> -(2) <= height l - height r <= 2 ->
+ height (create l x r) = max (height l) (height r) + 1.
+Proof.
+ unfold create; auto.
+Qed.
+
+(** bal *)
+
+Lemma bal_avl : forall l x r, avl l -> avl r ->
+ -(3) <= height l - height r <= 3 -> avl (bal l x r).
+Proof.
+ intros l x r; functional induction bal l x r; intros; clearf;
+ inv avl; simpl in *;
+ match goal with |- avl (assert_false _ _ _) => avl_nns
+ | _ => repeat apply create_avl; simpl in *; auto
+ end; omega_max.
+Qed.
+
+Lemma bal_height_1 : forall l x r, avl l -> avl r ->
+ -(3) <= height l - height r <= 3 ->
+ 0 <= height (bal l x r) - max (height l) (height r) <= 1.
+Proof.
+ intros l x r; functional induction bal l x r; intros; clearf;
+ inv avl; avl_nns; simpl in *; omega_max.
+Qed.
+
+Lemma bal_height_2 :
+ forall l x r, avl l -> avl r -> -(2) <= height l - height r <= 2 ->
+ height (bal l x r) == max (height l) (height r) +1.
+Proof.
+ intros l x r; functional induction bal l x r; intros; clearf;
+ inv avl; simpl in *; omega_max.
+Qed.
+
+Ltac omega_bal := match goal with
+ | H:avl ?l, H':avl ?r |- context [ bal ?l ?x ?r ] =>
+ generalize (bal_height_1 x H H') (bal_height_2 x H H');
+ omega_max
+ end.
+
+(** add *)
+
+Lemma add_avl_1 : forall s x, avl s ->
+ avl (add x s) /\ 0 <= height (add x s) - height s <= 1.
+Proof.
+ intros s x; functional induction (add x s); subst;intros; inv avl; simpl in *.
+ intuition; try constructor; simpl; auto; try omega_max.
+ (* LT *)
+ destruct IHt; auto.
+ split.
+ apply bal_avl; auto; omega_max.
+ omega_bal.
+ (* EQ *)
+ intuition; omega_max.
+ (* GT *)
+ destruct IHt; auto.
+ split.
+ apply bal_avl; auto; omega_max.
+ omega_bal.
+Qed.
+
+Lemma add_avl : forall s x, avl s -> avl (add x s).
+Proof.
+ intros; destruct (add_avl_1 x H); auto.
+Qed.
+Hint Resolve add_avl.
+
+(** join *)
+
+Lemma join_avl_1 : forall l x r, avl l -> avl r -> avl (join l x r) /\
+ 0<= height (join l x r) - max (height l) (height r) <= 1.
+Proof.
+ join_tac.
+
+ split; simpl; auto.
+ destruct (add_avl_1 x H0).
+ avl_nns; omega_max.
+ set (l:=Node ll lx lr lh) in *.
+ split; auto.
+ destruct (add_avl_1 x H).
+ simpl (height Leaf).
+ avl_nns; omega_max.
+
+ inversion_clear H.
+ assert (height (Node rl rx rr rh) = rh); auto.
+ set (r := Node rl rx rr rh) in *; clearbody r.
+ destruct (Hlr x r H2 H0); clear Hrl Hlr.
+ set (j := join lr x r) in *; clearbody j.
+ simpl.
+ assert (-(3) <= height ll - height j <= 3) by omega_max.
+ split.
+ apply bal_avl; auto.
+ omega_bal.
+
+ inversion_clear H0.
+ assert (height (Node ll lx lr lh) = lh); auto.
+ set (l := Node ll lx lr lh) in *; clearbody l.
+ destruct (Hrl H H1); clear Hrl Hlr.
+ set (j := join l x rl) in *; clearbody j.
+ simpl.
+ assert (-(3) <= height j - height rr <= 3) by omega_max.
+ split.
+ apply bal_avl; auto.
+ omega_bal.
+
+ clear Hrl Hlr.
+ assert (height (Node ll lx lr lh) = lh); auto.
+ assert (height (Node rl rx rr rh) = rh); auto.
+ set (l := Node ll lx lr lh) in *; clearbody l.
+ set (r := Node rl rx rr rh) in *; clearbody r.
+ assert (-(2) <= height l - height r <= 2) by omega_max.
+ split.
+ apply create_avl; auto.
+ rewrite create_height; auto; omega_max.
+Qed.
+
+Lemma join_avl : forall l x r, avl l -> avl r -> avl (join l x r).
+Proof.
+ intros; destruct (join_avl_1 x H H0); auto.
+Qed.
+Hint Resolve join_avl.
+
+(** remove_min *)
+
+Lemma remove_min_avl_1 : forall l x r h, avl (Node l x r h) ->
+ avl (remove_min l x r)#1 /\
+ 0 <= height (Node l x r h) - height (remove_min l x r)#1 <= 1.
+Proof.
+ intros l x r; functional induction (remove_min l x r); subst;simpl in *; intros.
+ inv avl; simpl in *; split; auto.
+ avl_nns; omega_max.
+ inversion_clear H.
+ rewrite e0 in IHp;simpl in IHp;destruct (IHp _x); auto.
+ split; simpl in *.
+ apply bal_avl; auto; omega_max.
+ omega_bal.
+Qed.
+
+Lemma remove_min_avl : forall l x r h, avl (Node l x r h) ->
+ avl (remove_min l x r)#1.
+Proof.
+ intros; destruct (remove_min_avl_1 H); auto.
+Qed.
+
+(** merge *)
+
+Lemma merge_avl_1 : forall s1 s2, avl s1 -> avl s2 ->
+ -(2) <= height s1 - height s2 <= 2 ->
+ avl (merge s1 s2) /\
+ 0<= height (merge s1 s2) - max (height s1) (height s2) <=1.
+Proof.
+ intros s1 s2; functional induction (merge s1 s2); intros;
+ try factornode _x _x0 _x1 _x2 as s1.
+ simpl; split; auto; avl_nns; omega_max.
+ simpl; split; auto; avl_nns; simpl in *; omega_max.
+ generalize (remove_min_avl_1 H0).
+ rewrite e1; destruct 1.
+ split.
+ apply bal_avl; auto.
+ simpl; omega_max.
+ simpl in *; omega_bal.
+Qed.
+
+Lemma merge_avl : forall s1 s2, avl s1 -> avl s2 ->
+ -(2) <= height s1 - height s2 <= 2 -> avl (merge s1 s2).
+Proof.
+ intros; destruct (merge_avl_1 H H0 H1); auto.
+Qed.
+
+
+(** remove *)
+
+Lemma remove_avl_1 : forall s x, avl s ->
+ avl (remove x s) /\ 0 <= height s - height (remove x s) <= 1.
+Proof.
+ intros s x; functional induction (remove x s); intros.
+ intuition; omega_max.
+ (* LT *)
+ inv avl.
+ destruct (IHt H0).
+ split.
+ apply bal_avl; auto.
+ omega_max.
+ omega_bal.
+ (* EQ *)
+ inv avl.
+ generalize (merge_avl_1 H0 H1 H2).
+ intuition omega_max.
+ (* GT *)
+ inv avl.
+ destruct (IHt H1).
+ split.
+ apply bal_avl; auto.
+ omega_max.
+ omega_bal.
+Qed.
+
+Lemma remove_avl : forall s x, avl s -> avl (remove x s).
+Proof.
+ intros; destruct (remove_avl_1 x H); auto.
+Qed.
+Hint Resolve remove_avl.
+
+(** concat *)
+
+Lemma concat_avl : forall s1 s2, avl s1 -> avl s2 -> avl (concat s1 s2).
+Proof.
+ intros s1 s2; functional induction (concat s1 s2); auto.
+ intros; apply join_avl; auto.
+ generalize (remove_min_avl H0); rewrite e1; simpl; auto.
+Qed.
+Hint Resolve concat_avl.
+
+(** split *)
+
+Lemma split_avl : forall s x, avl s ->
+ avl (split x s)#l /\ avl (split x s)#r.
+Proof.
+ intros s x; functional induction (split x s); simpl; auto.
+ rewrite e1 in IHt;simpl in IHt;inversion_clear 1; intuition.
+ simpl; inversion_clear 1; auto.
+ rewrite e1 in IHt;simpl in IHt;inversion_clear 1; intuition.
+Qed.
+
+(** inter *)
+
+Lemma inter_avl : forall s1 s2, avl s1 -> avl s2 -> avl (inter s1 s2).
+Proof.
+ intros s1 s2; functional induction inter s1 s2; auto; intros A1 A2;
+ generalize (split_avl x1 A2); rewrite e1; simpl; destruct 1;
+ inv avl; auto.
+Qed.
+
+(** diff *)
+
+Lemma diff_avl : forall s1 s2, avl s1 -> avl s2 -> avl (diff s1 s2).
+Proof.
+ intros s1 s2; functional induction diff s1 s2; auto; intros A1 A2;
+ generalize (split_avl x1 A2); rewrite e1; simpl; destruct 1;
+ inv avl; auto.
+Qed.
+
+(** union *)
+
+Lemma union_avl : forall s1 s2, avl s1 -> avl s2 -> avl (union s1 s2).
+Proof.
+ intros s1 s2; functional induction union s1 s2; auto; intros A1 A2;
+ generalize (split_avl x1 A2); rewrite e1; simpl; destruct 1;
+ inv avl; auto.
+Qed.
+
+(** filter *)
+
+Lemma filter_acc_avl : forall f s acc, avl s -> avl acc ->
+ avl (filter_acc f acc s).
+Proof.
+ induction s; simpl; auto.
+ intros.
+ inv avl.
+ destruct (f t); auto.
+Qed.
+Hint Resolve filter_acc_avl.
+
+Lemma filter_avl : forall f s, avl s -> avl (filter f s).
+Proof.
+ unfold filter; intros; apply filter_acc_avl; auto.
+Qed.
+
+(** partition *)
+
+Lemma partition_acc_avl_1 : forall f s acc, avl s ->
+ avl acc#1 -> avl (partition_acc f acc s)#1.
+Proof.
+ induction s; simpl; auto.
+ destruct acc as [acct accf]; simpl in *.
+ intros.
+ inv avl.
+ apply IHs2; auto.
+ apply IHs1; auto.
+ destruct (f t); simpl; auto.
+Qed.
+
+Lemma partition_acc_avl_2 : forall f s acc, avl s ->
+ avl acc#2 -> avl (partition_acc f acc s)#2.
+Proof.
+ induction s; simpl; auto.
+ destruct acc as [acct accf]; simpl in *.
+ intros.
+ inv avl.
+ apply IHs2; auto.
+ apply IHs1; auto.
+ destruct (f t); simpl; auto.
+Qed.
+
+Lemma partition_avl_1 : forall f s, avl s -> avl (partition f s)#1.
+Proof.
+ unfold partition; intros; apply partition_acc_avl_1; auto.
+Qed.
+
+Lemma partition_avl_2 : forall f s, avl s -> avl (partition f s)#2.
+Proof.
+ unfold partition; intros; apply partition_acc_avl_2; auto.
+Qed.
+
+End AvlProofs.
+
+
+Module OcamlOps (Import I:Int)(X:OrderedType).
+Module Import AvlProofs := AvlProofs I X.
+Import Raw.
+Import Raw.Proofs.
+Import II.
+Open Local Scope pair_scope.
+Open Local Scope nat_scope.
+
+(** Properties of cardinal *)
+
+Lemma bal_cardinal : forall l x r,
+ cardinal (bal l x r) = S (cardinal l + cardinal r).
+Proof.
+ intros l x r; functional induction bal l x r; intros; clearf;
+ simpl; auto with arith; romega with *.
+Qed.
+
+Lemma add_cardinal : forall x s,
+ cardinal (add x s) <= S (cardinal s).
+Proof.
+ intros; functional induction add x s; simpl; auto with arith;
+ rewrite bal_cardinal; romega with *.
+Qed.
+
+Lemma join_cardinal : forall l x r,
+ cardinal (join l x r) <= S (cardinal l + cardinal r).
+Proof.
+ join_tac; auto with arith.
+ simpl; apply add_cardinal.
+ simpl; destruct X.compare; simpl; auto with arith.
+ generalize (bal_cardinal (add x ll) lx lr) (add_cardinal x ll);
+ romega with *.
+ generalize (bal_cardinal ll lx (add x lr)) (add_cardinal x lr);
+ romega with *.
+ generalize (bal_cardinal ll lx (join lr x (Node rl rx rr rh)))
+ (Hlr x (Node rl rx rr rh)); simpl; romega with *.
+ simpl S in *; generalize (bal_cardinal (join (Node ll lx lr lh) x rl) rx rr).
+ romega with *.
+Qed.
+
+Lemma split_cardinal_1 : forall x s,
+ (cardinal (split x s)#l <= cardinal s)%nat.
+Proof.
+ intros x s; functional induction split x s; simpl; auto.
+ rewrite e1 in IHt; simpl in *.
+ romega with *.
+ romega with *.
+ rewrite e1 in IHt; simpl in *.
+ generalize (@join_cardinal l y rl); romega with *.
+Qed.
+
+Lemma split_cardinal_2 : forall x s,
+ (cardinal (split x s)#r <= cardinal s)%nat.
+Proof.
+ intros x s; functional induction split x s; simpl; auto.
+ rewrite e1 in IHt; simpl in *.
+ generalize (@join_cardinal rl y r); romega with *.
+ romega with *.
+ rewrite e1 in IHt; simpl in *; romega with *.
+Qed.
+
+(** * [ocaml_union], an union faithful to the original ocaml code *)
+
+Definition cardinal2 (s:t*t) := (cardinal s#1 + cardinal s#2)%nat.
+
+Ltac ocaml_union_tac :=
+ intros; unfold cardinal2; simpl fst in *; simpl snd in *;
+ match goal with H: split ?x ?s = _ |- _ =>
+ generalize (split_cardinal_1 x s) (split_cardinal_2 x s);
+ rewrite H; simpl; romega with *
+ end.
+
+Import Logic. (* Unhide eq, otherwise Function complains. *)
+
+Function ocaml_union (s : t * t) { measure cardinal2 s } : t :=
+ match s with
+ | (Leaf, Leaf) => s#2
+ | (Leaf, Node _ _ _ _) => s#2
+ | (Node _ _ _ _, Leaf) => s#1
+ | (Node l1 x1 r1 h1, Node l2 x2 r2 h2) =>
+ if ge_lt_dec h1 h2 then
+ if eq_dec h2 1%I then add x2 s#1 else
+ let (l2',_,r2') := split x1 s#2 in
+ join (ocaml_union (l1,l2')) x1 (ocaml_union (r1,r2'))
+ else
+ if eq_dec h1 1%I then add x1 s#2 else
+ let (l1',_,r1') := split x2 s#1 in
+ join (ocaml_union (l1',l2)) x2 (ocaml_union (r1',r2))
+ end.
+Proof.
+abstract ocaml_union_tac.
+abstract ocaml_union_tac.
+abstract ocaml_union_tac.
+abstract ocaml_union_tac.
+Defined.
+
+Lemma ocaml_union_in : forall s y,
+ bst s#1 -> avl s#1 -> bst s#2 -> avl s#2 ->
+ (In y (ocaml_union s) <-> In y s#1 \/ In y s#2).
+Proof.
+ intros s; functional induction ocaml_union s; intros y B1 A1 B2 A2;
+ simpl fst in *; simpl snd in *; try clear e0 e1.
+ intuition_in.
+ intuition_in.
+ intuition_in.
+ (* add x2 s#1 *)
+ inv avl.
+ rewrite (height_0 H); [ | avl_nn l2; omega_max].
+ rewrite (height_0 H0); [ | avl_nn r2; omega_max].
+ rewrite add_in; intuition_in.
+ (* join (union (l1,l2')) x1 (union (r1,r2')) *)
+ generalize
+ (split_avl x1 A2) (split_bst x1 B2)
+ (split_in_1 x1 y B2) (split_in_2 x1 y B2).
+ rewrite e2; simpl.
+ destruct 1; destruct 1; inv avl; inv bst.
+ rewrite join_in, IHt, IHt0; auto.
+ do 2 (intro Eq; rewrite Eq; clear Eq).
+ case (X.compare y x1); intuition_in.
+ (* add x1 s#2 *)
+ inv avl.
+ rewrite (height_0 H3); [ | avl_nn l1; omega_max].
+ rewrite (height_0 H4); [ | avl_nn r1; omega_max].
+ rewrite add_in; auto; intuition_in.
+ (* join (union (l1',l2)) x1 (union (r1',r2)) *)
+ generalize
+ (split_avl x2 A1) (split_bst x2 B1)
+ (split_in_1 x2 y B1) (split_in_2 x2 y B1).
+ rewrite e2; simpl.
+ destruct 1; destruct 1; inv avl; inv bst.
+ rewrite join_in, IHt, IHt0; auto.
+ do 2 (intro Eq; rewrite Eq; clear Eq).
+ case (X.compare y x2); intuition_in.
+Qed.
+
+Lemma ocaml_union_bst : forall s,
+ bst s#1 -> avl s#1 -> bst s#2 -> avl s#2 -> bst (ocaml_union s).
+Proof.
+ intros s; functional induction ocaml_union s; intros B1 A1 B2 A2;
+ simpl fst in *; simpl snd in *; try clear e0 e1;
+ try apply add_bst; auto.
+ (* join (union (l1,l2')) x1 (union (r1,r2')) *)
+ clear _x _x0; factornode l2 x2 r2 h2 as s2.
+ generalize (split_avl x1 A2) (split_bst x1 B2)
+ (@split_in_1 s2 x1)(@split_in_2 s2 x1).
+ rewrite e2; simpl.
+ destruct 1; destruct 1; intros.
+ inv bst; inv avl.
+ apply join_bst; auto.
+ intro y; rewrite ocaml_union_in, H3; intuition_in.
+ intro y; rewrite ocaml_union_in, H4; intuition_in.
+ (* join (union (l1',l2)) x1 (union (r1',r2)) *)
+ clear _x _x0; factornode l1 x1 r1 h1 as s1.
+ generalize (split_avl x2 A1) (split_bst x2 B1)
+ (@split_in_1 s1 x2)(@split_in_2 s1 x2).
+ rewrite e2; simpl.
+ destruct 1; destruct 1; intros.
+ inv bst; inv avl.
+ apply join_bst; auto.
+ intro y; rewrite ocaml_union_in, H3; intuition_in.
+ intro y; rewrite ocaml_union_in, H4; intuition_in.
+Qed.
+
+Lemma ocaml_union_avl : forall s,
+ avl s#1 -> avl s#2 -> avl (ocaml_union s).
+Proof.
+ intros s; functional induction ocaml_union s;
+ simpl fst in *; simpl snd in *; auto.
+ intros A1 A2; generalize (split_avl x1 A2); rewrite e2; simpl.
+ inv avl; destruct 1; auto.
+ intros A1 A2; generalize (split_avl x2 A1); rewrite e2; simpl.
+ inv avl; destruct 1; auto.
+Qed.
+
+Lemma ocaml_union_alt : forall s, bst s#1 -> avl s#1 -> bst s#2 -> avl s#2 ->
+ Equal (ocaml_union s) (union s#1 s#2).
+Proof.
+ red; intros; rewrite ocaml_union_in, union_in; simpl; intuition.
+Qed.
+
+
+(** * [ocaml_subset], a subset faithful to the original ocaml code *)
+
+Function ocaml_subset (s:t*t) { measure cardinal2 s } : bool :=
+ match s with
+ | (Leaf, _) => true
+ | (Node _ _ _ _, Leaf) => false
+ | (Node l1 x1 r1 h1, Node l2 x2 r2 h2) =>
+ match X.compare x1 x2 with
+ | EQ _ => ocaml_subset (l1,l2) && ocaml_subset (r1,r2)
+ | LT _ => ocaml_subset (Node l1 x1 Leaf 0%I, l2) && ocaml_subset (r1,s#2)
+ | GT _ => ocaml_subset (Node Leaf x1 r1 0%I, r2) && ocaml_subset (l1,s#2)
+ end
+ end.
+
+Proof.
+ intros; unfold cardinal2; simpl; abstract romega with *.
+ intros; unfold cardinal2; simpl; abstract romega with *.
+ intros; unfold cardinal2; simpl; abstract romega with *.
+ intros; unfold cardinal2; simpl; abstract romega with *.
+ intros; unfold cardinal2; simpl; abstract romega with *.
+ intros; unfold cardinal2; simpl; abstract romega with *.
+Defined.
+
+Lemma ocaml_subset_12 : forall s,
+ bst s#1 -> bst s#2 ->
+ (ocaml_subset s = true <-> Subset s#1 s#2).
+Proof.
+ intros s; functional induction ocaml_subset s; simpl;
+ intros B1 B2; try clear e0.
+ intuition.
+ red; auto; inversion 1.
+ split; intros; try discriminate.
+ assert (H': In _x0 Leaf) by auto; inversion H'.
+ (**)
+ simpl in *; inv bst.
+ rewrite andb_true_iff, IHb, IHb0; auto; clear IHb IHb0.
+ unfold Subset; intuition_in.
+ assert (X.eq a x2) by order; intuition_in.
+ assert (In a (Node l2 x2 r2 h2)) by auto; intuition_in; order.
+ assert (In a (Node l2 x2 r2 h2)) by auto; intuition_in; order.
+ (**)
+ simpl in *; inv bst.
+ rewrite andb_true_iff, IHb, IHb0; auto; clear IHb IHb0.
+ unfold Subset; intuition_in.
+ assert (In a (Node l2 x2 r2 h2)) by auto; intuition_in; order.
+ assert (In a (Node l2 x2 r2 h2)) by auto; intuition_in; order.
+ (**)
+ simpl in *; inv bst.
+ rewrite andb_true_iff, IHb, IHb0; auto; clear IHb IHb0.
+ unfold Subset; intuition_in.
+ assert (In a (Node l2 x2 r2 h2)) by auto; intuition_in; order.
+ assert (In a (Node l2 x2 r2 h2)) by auto; intuition_in; order.
+Qed.
+
+Lemma ocaml_subset_alt : forall s, bst s#1 -> bst s#2 ->
+ ocaml_subset s = subset s#1 s#2.
+Proof.
+ intros.
+ generalize (ocaml_subset_12 H H0); rewrite <-subset_12 by auto.
+ destruct ocaml_subset; destruct subset; intuition.
+Qed.
+
+
+
+(** [ocaml_compare], a compare faithful to the original ocaml code *)
+
+(** termination of [compare_aux] *)
+
+Fixpoint cardinal_e e := match e with
+ | End => 0
+ | More _ s r => S (cardinal s + cardinal_e r)
+ end.
+
+Lemma cons_cardinal_e : forall s e,
+ cardinal_e (cons s e) = cardinal s + cardinal_e e.
+Proof.
+ induction s; simpl; intros; auto.
+ rewrite IHs1; simpl; rewrite <- plus_n_Sm; auto with arith.
+Qed.
+
+Definition cardinal_e_2 e := cardinal_e e#1 + cardinal_e e#2.
+
+Function ocaml_compare_aux
+ (e:enumeration*enumeration) { measure cardinal_e_2 e } : comparison :=
+ match e with
+ | (End,End) => Eq
+ | (End,More _ _ _) => Lt
+ | (More _ _ _, End) => Gt
+ | (More x1 r1 e1, More x2 r2 e2) =>
+ match X.compare x1 x2 with
+ | EQ _ => ocaml_compare_aux (cons r1 e1, cons r2 e2)
+ | LT _ => Lt
+ | GT _ => Gt
+ end
+ end.
+
+Proof.
+intros; unfold cardinal_e_2; simpl;
+abstract (do 2 rewrite cons_cardinal_e; romega with *).
+Defined.
+
+Definition ocaml_compare s1 s2 :=
+ ocaml_compare_aux (cons s1 End, cons s2 End).
+
+Lemma ocaml_compare_aux_Cmp : forall e,
+ Cmp (ocaml_compare_aux e) (flatten_e e#1) (flatten_e e#2).
+Proof.
+ intros e; functional induction ocaml_compare_aux e; simpl; intros;
+ auto; try discriminate.
+ apply L.eq_refl.
+ simpl in *.
+ apply cons_Cmp; auto.
+ rewrite <- 2 cons_1; auto.
+Qed.
+
+Lemma ocaml_compare_Cmp : forall s1 s2,
+ Cmp (ocaml_compare s1 s2) (elements s1) (elements s2).
+Proof.
+ unfold ocaml_compare; intros.
+ assert (H1:=cons_1 s1 End).
+ assert (H2:=cons_1 s2 End).
+ simpl in *; rewrite <- app_nil_end in *; rewrite <-H1,<-H2.
+ apply (@ocaml_compare_aux_Cmp (cons s1 End, cons s2 End)).
+Qed.
+
+Lemma ocaml_compare_alt : forall s1 s2, bst s1 -> bst s2 ->
+ ocaml_compare s1 s2 = compare s1 s2.
+Proof.
+ intros s1 s2 B1 B2.
+ generalize (ocaml_compare_Cmp s1 s2)(compare_Cmp s1 s2).
+ unfold Cmp.
+ destruct ocaml_compare; destruct compare; auto; intros; elimtype False.
+ elim (lt_not_eq B1 B2 H0); auto.
+ elim (lt_not_eq B2 B1 H0); auto.
+ apply eq_sym; auto.
+ elim (lt_not_eq B1 B2 H); auto.
+ elim (lt_not_eq B1 B1).
+ red; eapply L.lt_trans; eauto.
+ apply eq_refl.
+ elim (lt_not_eq B2 B1 H); auto.
+ apply eq_sym; auto.
+ elim (lt_not_eq B1 B2 H0); auto.
+ elim (lt_not_eq B1 B1).
+ red; eapply L.lt_trans; eauto.
+ apply eq_refl.
+Qed.
+
+
+(** * Equality test *)
+
+Definition ocaml_equal s1 s2 : bool :=
+ match ocaml_compare s1 s2 with
+ | Eq => true
+ | _ => false
+ end.
+
+Lemma ocaml_equal_1 : forall s1 s2, bst s1 -> bst s2 ->
+ Equal s1 s2 -> ocaml_equal s1 s2 = true.
+Proof.
+unfold ocaml_equal; intros s1 s2 B1 B2 E.
+generalize (ocaml_compare_Cmp s1 s2).
+destruct (ocaml_compare s1 s2); auto; intros.
+elim (lt_not_eq B1 B2 H E); auto.
+elim (lt_not_eq B2 B1 H (eq_sym E)); auto.
+Qed.
+
+Lemma ocaml_equal_2 : forall s1 s2,
+ ocaml_equal s1 s2 = true -> Equal s1 s2.
+Proof.
+unfold ocaml_equal; intros s1 s2 E.
+generalize (ocaml_compare_Cmp s1 s2);
+ destruct ocaml_compare; auto; discriminate.
+Qed.
+
+Lemma ocaml_equal_alt : forall s1 s2, bst s1 -> bst s2 ->
+ ocaml_equal s1 s2 = equal s1 s2.
+Proof.
+intros; unfold ocaml_equal, equal; rewrite ocaml_compare_alt; auto.
+Qed.
+
+End OcamlOps.
+
+
+
+(** * Encapsulation
+
+ We can implement [S] with balanced binary search trees.
+ When compared to [FSetAVL], we maintain here two invariants
+ (bst and avl) instead of only bst, which is enough for fulfilling
+ the FSet interface.
+
+ This encapsulation propose the non-structural variants
+ [ocaml_union], [ocaml_subset], [ocaml_compare], [ocaml_equal].
+*)
+
+Module IntMake (I:Int)(X: OrderedType) <: S with Module E := X.
+
+ Module E := X.
+ Module Import OcamlOps := OcamlOps I X.
+ Import AvlProofs.
+ Import Raw.
+ Import Raw.Proofs.
+
+ Record bbst := Bbst {this :> Raw.t; is_bst : bst this; is_avl : avl this}.
+ Definition t := bbst.
+ Definition elt := E.t.
+
+ Definition In (x : elt) (s : t) : Prop := In x s.
+ Definition Equal (s s':t) : Prop := forall a : elt, In a s <-> In a s'.
+ Definition Subset (s s':t) : Prop := forall a : elt, In a s -> In a s'.
+ Definition Empty (s:t) : Prop := forall a : elt, ~ In a s.
+ Definition For_all (P : elt -> Prop) (s:t) : Prop := forall x, In x s -> P x.
+ Definition Exists (P : elt -> Prop) (s:t) : Prop := exists x, In x s /\ P x.
+
+ Lemma In_1 : forall (s:t)(x y:elt), E.eq x y -> In x s -> In y s.
+ Proof. intro s; exact (@In_1 s). Qed.
+
+ Definition mem (x:elt)(s:t) : bool := mem x s.
+
+ Definition empty : t := Bbst empty_bst empty_avl.
+ Definition is_empty (s:t) : bool := is_empty s.
+ Definition singleton (x:elt) : t :=
+ Bbst (singleton_bst x) (singleton_avl x).
+ Definition add (x:elt)(s:t) : t :=
+ Bbst (add_bst x (is_bst s)) (add_avl x (is_avl s)).
+ Definition remove (x:elt)(s:t) : t :=
+ Bbst (remove_bst x (is_bst s)) (remove_avl x (is_avl s)).
+ Definition inter (s s':t) : t :=
+ Bbst (inter_bst (is_bst s) (is_bst s'))
+ (inter_avl (is_avl s) (is_avl s')).
+ Definition union (s s':t) : t :=
+ Bbst (union_bst (is_bst s) (is_bst s'))
+ (union_avl (is_avl s) (is_avl s')).
+ Definition ocaml_union (s s':t) : t :=
+ Bbst (@ocaml_union_bst (s.(this),s'.(this))
+ (is_bst s) (is_avl s) (is_bst s') (is_avl s'))
+ (@ocaml_union_avl (s.(this),s'.(this)) (is_avl s) (is_avl s')).
+ Definition diff (s s':t) : t :=
+ Bbst (diff_bst (is_bst s) (is_bst s'))
+ (diff_avl (is_avl s) (is_avl s')).
+ Definition elements (s:t) : list elt := elements s.
+ Definition min_elt (s:t) : option elt := min_elt s.
+ Definition max_elt (s:t) : option elt := max_elt s.
+ Definition choose (s:t) : option elt := choose s.
+ Definition fold (B : Type) (f : elt -> B -> B) (s:t) : B -> B := fold f s.
+ Definition cardinal (s:t) : nat := cardinal s.
+ Definition filter (f : elt -> bool) (s:t) : t :=
+ Bbst (filter_bst f (is_bst s)) (filter_avl f (is_avl s)).
+ Definition for_all (f : elt -> bool) (s:t) : bool := for_all f s.
+ Definition exists_ (f : elt -> bool) (s:t) : bool := exists_ f s.
+ Definition partition (f : elt -> bool) (s:t) : t * t :=
+ let p := partition f s in
+ (@Bbst (fst p) (partition_bst_1 f (is_bst s))
+ (partition_avl_1 f (is_avl s)),
+ @Bbst (snd p) (partition_bst_2 f (is_bst s))
+ (partition_avl_2 f (is_avl s))).
+
+ Definition equal (s s':t) : bool := equal s s'.
+ Definition ocaml_equal (s s':t) : bool := ocaml_equal s s'.
+ Definition subset (s s':t) : bool := subset s s'.
+ Definition ocaml_subset (s s':t) : bool :=
+ ocaml_subset (s.(this),s'.(this)).
+
+ Definition eq (s s':t) : Prop := Equal s s'.
+ Definition lt (s s':t) : Prop := lt s s'.
+
+ Definition compare (s s':t) : Compare lt eq s s'.
+ Proof.
+ intros (s,b,a) (s',b',a').
+ generalize (compare_Cmp s s').
+ destruct Raw.compare; intros; [apply EQ|apply LT|apply GT]; red; auto.
+ change (Raw.Equal s s'); auto.
+ Defined.
+
+ Definition ocaml_compare (s s':t) : Compare lt eq s s'.
+ Proof.
+ intros (s,b,a) (s',b',a').
+ generalize (ocaml_compare_Cmp s s').
+ destruct ocaml_compare; intros; [apply EQ|apply LT|apply GT]; red; auto.
+ change (Raw.Equal s s'); auto.
+ Defined.
+
+ (* specs *)
+ Section Specs.
+ Variable s s' s'': t.
+ Variable x y : elt.
+
+ Hint Resolve is_bst is_avl.
+
+ Lemma mem_1 : In x s -> mem x s = true.
+ Proof. exact (mem_1 (is_bst s)). Qed.
+ Lemma mem_2 : mem x s = true -> In x s.
+ Proof. exact (@mem_2 s x). Qed.
+
+ Lemma equal_1 : Equal s s' -> equal s s' = true.
+ Proof. exact (equal_1 (is_bst s) (is_bst s')). Qed.
+ Lemma equal_2 : equal s s' = true -> Equal s s'.
+ Proof. exact (@equal_2 s s'). Qed.
+
+ Lemma ocaml_equal_alt : ocaml_equal s s' = equal s s'.
+ Proof.
+ destruct s; destruct s'; unfold ocaml_equal, equal; simpl.
+ apply ocaml_equal_alt; auto.
+ Qed.
+
+ Lemma ocaml_equal_1 : Equal s s' -> ocaml_equal s s' = true.
+ Proof. exact (ocaml_equal_1 (is_bst s) (is_bst s')). Qed.
+ Lemma ocaml_equal_2 : ocaml_equal s s' = true -> Equal s s'.
+ Proof. exact (@ocaml_equal_2 s s'). Qed.
+
+ Ltac wrap t H := unfold t, In; simpl; rewrite H; auto; intuition.
+
+ Lemma subset_1 : Subset s s' -> subset s s' = true.
+ Proof. wrap subset subset_12. Qed.
+ Lemma subset_2 : subset s s' = true -> Subset s s'.
+ Proof. wrap subset subset_12. Qed.
+
+ Lemma ocaml_subset_alt : ocaml_subset s s' = subset s s'.
+ Proof.
+ destruct s; destruct s'; unfold ocaml_subset, subset; simpl.
+ rewrite ocaml_subset_alt; auto.
+ Qed.
+
+ Lemma ocaml_subset_1 : Subset s s' -> ocaml_subset s s' = true.
+ Proof. wrap ocaml_subset ocaml_subset_12; simpl; auto. Qed.
+ Lemma ocaml_subset_2 : ocaml_subset s s' = true -> Subset s s'.
+ Proof. wrap ocaml_subset ocaml_subset_12; simpl; auto. Qed.
+
+ Lemma empty_1 : Empty empty.
+ Proof. exact empty_1. Qed.
+
+ Lemma is_empty_1 : Empty s -> is_empty s = true.
+ Proof. exact (@is_empty_1 s). Qed.
+ Lemma is_empty_2 : is_empty s = true -> Empty s.
+ Proof. exact (@is_empty_2 s). Qed.
+
+ Lemma add_1 : E.eq x y -> In y (add x s).
+ Proof. wrap add add_in. Qed.
+ Lemma add_2 : In y s -> In y (add x s).
+ Proof. wrap add add_in. Qed.
+ Lemma add_3 : ~ E.eq x y -> In y (add x s) -> In y s.
+ Proof. wrap add add_in. elim H; auto. Qed.
+
+ Lemma remove_1 : E.eq x y -> ~ In y (remove x s).
+ Proof. wrap remove remove_in. Qed.
+ Lemma remove_2 : ~ E.eq x y -> In y s -> In y (remove x s).
+ Proof. wrap remove remove_in. Qed.
+ Lemma remove_3 : In y (remove x s) -> In y s.
+ Proof. wrap remove remove_in. Qed.
+
+ Lemma singleton_1 : In y (singleton x) -> E.eq x y.
+ Proof. exact (@singleton_1 x y). Qed.
+ Lemma singleton_2 : E.eq x y -> In y (singleton x).
+ Proof. exact (@singleton_2 x y). Qed.
+
+ Lemma union_1 : In x (union s s') -> In x s \/ In x s'.
+ Proof. wrap union union_in. Qed.
+ Lemma union_2 : In x s -> In x (union s s').
+ Proof. wrap union union_in. Qed.
+ Lemma union_3 : In x s' -> In x (union s s').
+ Proof. wrap union union_in. Qed.
+
+ Lemma ocaml_union_alt : Equal (ocaml_union s s') (union s s').
+ Proof.
+ unfold ocaml_union, union, Equal, In.
+ destruct s as (s0,b,a); destruct s' as (s0',b',a'); simpl.
+ exact (@ocaml_union_alt (s0,s0') b a b' a').
+ Qed.
+
+ Lemma ocaml_union_1 : In x (ocaml_union s s') -> In x s \/ In x s'.
+ Proof. wrap ocaml_union ocaml_union_in; simpl; auto. Qed.
+ Lemma ocaml_union_2 : In x s -> In x (ocaml_union s s').
+ Proof. wrap ocaml_union ocaml_union_in; simpl; auto. Qed.
+ Lemma ocaml_union_3 : In x s' -> In x (ocaml_union s s').
+ Proof. wrap ocaml_union ocaml_union_in; simpl; auto. Qed.
+
+ Lemma inter_1 : In x (inter s s') -> In x s.
+ Proof. wrap inter inter_in. Qed.
+ Lemma inter_2 : In x (inter s s') -> In x s'.
+ Proof. wrap inter inter_in. Qed.
+ Lemma inter_3 : In x s -> In x s' -> In x (inter s s').
+ Proof. wrap inter inter_in. Qed.
+
+ Lemma diff_1 : In x (diff s s') -> In x s.
+ Proof. wrap diff diff_in. Qed.
+ Lemma diff_2 : In x (diff s s') -> ~ In x s'.
+ Proof. wrap diff diff_in. Qed.
+ Lemma diff_3 : In x s -> ~ In x s' -> In x (diff s s').
+ Proof. wrap diff diff_in. Qed.
+
+ Lemma fold_1 : 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.
+ unfold fold, elements; intros; apply fold_1; auto.
+ Qed.
+
+ Lemma cardinal_1 : cardinal s = length (elements s).
+ Proof.
+ unfold cardinal, elements; intros; apply elements_cardinal; auto.
+ Qed.
+
+ Section Filter.
+ Variable f : elt -> bool.
+
+ Lemma filter_1 : compat_bool E.eq f -> In x (filter f s) -> In x s.
+ Proof. intro. wrap filter filter_in. Qed.
+ Lemma filter_2 : compat_bool E.eq f -> In x (filter f s) -> f x = true.
+ Proof. intro. wrap filter filter_in. Qed.
+ Lemma filter_3 : compat_bool E.eq f -> In x s -> f x = true -> In x (filter f s).
+ Proof. intro. wrap filter filter_in. Qed.
+
+ Lemma for_all_1 : compat_bool E.eq f -> For_all (fun x => f x = true) s -> for_all f s = true.
+ Proof. exact (@for_all_1 f s). Qed.
+ Lemma for_all_2 : compat_bool E.eq f -> for_all f s = true -> For_all (fun x => f x = true) s.
+ Proof. exact (@for_all_2 f s). Qed.
+
+ Lemma exists_1 : compat_bool E.eq f -> Exists (fun x => f x = true) s -> exists_ f s = true.
+ Proof. exact (@exists_1 f s). Qed.
+ Lemma exists_2 : compat_bool E.eq f -> exists_ f s = true -> Exists (fun x => f x = true) s.
+ Proof. exact (@exists_2 f s). Qed.
+
+ Lemma partition_1 : compat_bool E.eq f ->
+ Equal (fst (partition f s)) (filter f s).
+ Proof.
+ unfold partition, filter, Equal, In; simpl ;intros H a.
+ rewrite partition_in_1, filter_in; intuition.
+ Qed.
+
+ Lemma partition_2 : compat_bool E.eq f ->
+ Equal (snd (partition f s)) (filter (fun x => negb (f x)) s).
+ Proof.
+ unfold partition, filter, Equal, In; simpl ;intros H a.
+ rewrite partition_in_2, filter_in; intuition.
+ rewrite H2; auto.
+ destruct (f a); auto.
+ red; intros; f_equal.
+ rewrite (H _ _ H0); auto.
+ Qed.
+
+ End Filter.
+
+ Lemma elements_1 : In x s -> InA E.eq x (elements s).
+ Proof. wrap elements elements_in. Qed.
+ Lemma elements_2 : InA E.eq x (elements s) -> In x s.
+ Proof. wrap elements elements_in. Qed.
+ Lemma elements_3 : sort E.lt (elements s).
+ Proof. exact (elements_sort (is_bst s)). Qed.
+ Lemma elements_3w : NoDupA E.eq (elements s).
+ Proof. exact (elements_nodup (is_bst s)). Qed.
+
+ Lemma min_elt_1 : min_elt s = Some x -> In x s.
+ Proof. exact (@min_elt_1 s x). Qed.
+ Lemma min_elt_2 : min_elt s = Some x -> In y s -> ~ E.lt y x.
+ Proof. exact (@min_elt_2 s x y (is_bst s)). Qed.
+ Lemma min_elt_3 : min_elt s = None -> Empty s.
+ Proof. exact (@min_elt_3 s). Qed.
+
+ Lemma max_elt_1 : max_elt s = Some x -> In x s.
+ Proof. exact (@max_elt_1 s x). Qed.
+ Lemma max_elt_2 : max_elt s = Some x -> In y s -> ~ E.lt x y.
+ Proof. exact (@max_elt_2 s x y (is_bst s)). Qed.
+ Lemma max_elt_3 : max_elt s = None -> Empty s.
+ Proof. exact (@max_elt_3 s). Qed.
+
+ Lemma choose_1 : choose s = Some x -> In x s.
+ Proof. exact (@choose_1 s x). Qed.
+ Lemma choose_2 : choose s = None -> Empty s.
+ Proof. exact (@choose_2 s). Qed.
+ Lemma choose_3 : choose s = Some x -> choose s' = Some y ->
+ Equal s s' -> E.eq x y.
+ Proof. exact (@choose_3 _ _ (is_bst s) (is_bst s') x y). Qed.
+
+ Lemma eq_refl : eq s s.
+ Proof. exact (eq_refl s). Qed.
+ Lemma eq_sym : eq s s' -> eq s' s.
+ Proof. exact (@eq_sym s s'). Qed.
+ Lemma eq_trans : eq s s' -> eq s' s'' -> eq s s''.
+ Proof. exact (@eq_trans s s' s''). Qed.
+
+ Lemma lt_trans : lt s s' -> lt s' s'' -> lt s s''.
+ Proof. exact (@lt_trans s s' s''). Qed.
+ Lemma lt_not_eq : lt s s' -> ~eq s s'.
+ Proof. exact (@lt_not_eq _ _ (is_bst s) (is_bst s')). Qed.
+
+ End Specs.
+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/FSets/FSetInterface.v b/theories/FSets/FSetInterface.v
index 64ad234b..1255fcc8 100644
--- a/theories/FSets/FSetInterface.v
+++ b/theories/FSets/FSetInterface.v
@@ -6,40 +6,53 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: FSetInterface.v 8820 2006-05-15 11:44:05Z letouzey $ *)
+(* $Id: FSetInterface.v 10616 2008-03-04 17:33:35Z letouzey $ *)
(** * Finite set library *)
-(** Set interfaces *)
-
-(* begin hide *)
-Require Export Bool.
-Require Export OrderedType.
+(** 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:
+ - [WSfun] : functorial signature for weak sets, non-dependent style
+ - [WS] : self-contained version of [WSfun]
+ - [Sfun] : functorial signature for ordered sets, non-dependent style
+ - [S] : self-contained version of [Sfun]
+ - [Sdep] : analog of [S] written using dependent style
+
+ If unsure, [S] is probably what you're looking for: other signatures
+ are subsets of it, apart from [Sdep] which is isomorphic to [S] (see
+ [FSetBridge]).
+*)
+
+Require Export Bool OrderedType DecidableType.
Set Implicit Arguments.
Unset Strict Implicit.
-(* end hide *)
-(** Compatibility of a boolean function with respect to an equality. *)
-Definition compat_bool (A:Set)(eqA: A->A->Prop)(f: A-> bool) :=
- forall x y : A, eqA x y -> f x = f y.
+(** * Non-dependent signatures
-(** Compatibility of a predicate with respect to an equality. *)
-Definition compat_P (A:Set)(eqA: A->A->Prop)(P : A -> Prop) :=
- forall x y : A, eqA x y -> P x -> P y.
+ The following signatures presents sets as purely informative
+ programs together with axioms *)
-Hint Unfold compat_bool compat_P.
-(** * Non-dependent signature
- Signature [S] presents sets as purely informative programs
- together with axioms *)
+(** ** Functorial signature for weak sets
-Module Type S.
+ Weak sets are sets without ordering on base elements, only
+ a decidable equality. *)
+
+Module Type WSfun (E : EqualityType).
+
+ (** The module E of base objects is meant to be a [DecidableType]
+ (and used to be so). But requiring only an [EqualityType] here
+ allows subtyping between weak and ordered sets *)
- Declare Module E : OrderedType.
Definition elt := E.t.
- Parameter t : Set. (** the abstract type of sets *)
+ Parameter t : Type. (** the abstract type of sets *)
(** Logical predicates *)
Parameter In : elt -> t -> Prop.
@@ -82,10 +95,12 @@ Module Type S.
(** Set difference. *)
Definition eq : t -> t -> Prop := Equal.
- Parameter lt : t -> t -> Prop.
- Parameter compare : forall s s' : t, Compare lt eq s s'.
- (** Total ordering between sets. Can be used as the ordering function
- for doing sets of sets. *)
+ (** In order to have the subtyping WS < S between weak and ordered
+ sets, we do not require here an [eq_dec]. This interface is hence
+ not compatible with [DecidableType], but only with [EqualityType],
+ so in general it may not possible to form weak sets of weak sets.
+ Some particular implementations may allow this nonetheless, in
+ particular [FSetWeakList.Make]. *)
Parameter equal : t -> t -> bool.
(** [equal s1 s2] tests whether the sets [s1] and [s2] are
@@ -95,15 +110,11 @@ Module Type S.
(** [subset s1 s2] tests whether the set [s1] is a subset of
the set [s2]. *)
- (** Coq comment: [iter] is useless in a purely functional world *)
- (** iter: (elt -> unit) -> set -> unit. i*)
- (** [iter f s] applies [f] in turn to all elements of [s].
- The order in which the elements of [s] are presented to [f]
- is unspecified. *)
-
- Parameter fold : forall A : Set, (elt -> A -> A) -> t -> A -> A.
+ 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], in increasing order. *)
+ 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
@@ -125,59 +136,39 @@ Module Type S.
Parameter cardinal : t -> nat.
(** Return the number of elements of a set. *)
- (** Coq comment: nat instead of int ... *)
Parameter elements : t -> list elt.
- (** Return the list of all elements of the given set.
- The returned list is sorted in increasing order with respect
- to the ordering [Ord.compare], where [Ord] is the argument
- given to {!Set.Make}. *)
-
- Parameter min_elt : t -> option elt.
- (** Return the smallest element of the given set
- (with respect to the [Ord.compare] ordering), or raise
- [Not_found] if the set is empty. *)
- (** Coq comment: [Not_found] is represented by the option type *)
-
- Parameter max_elt : t -> option elt.
- (** Same as {!Set.S.min_elt}, but returns the largest element of the
- given set. *)
- (** Coq comment: [Not_found] is represented by the option type *)
+ (** 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 raise [Not_found] if
- the set is empty. Which element is chosen is unspecified,
- but equal elements will be chosen for equal sets. *)
- (** Coq comment: [Not_found] is represented by the option type *)
+ (** 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. *)
Section Spec.
- Variable s s' s'' : t.
+ Variable s s' s'': t.
Variable x y : elt.
(** Specification of [In] *)
Parameter In_1 : E.eq x y -> In x s -> In y s.
-
+
(** Specification of [eq] *)
Parameter eq_refl : eq s s.
Parameter eq_sym : eq s s' -> eq s' s.
Parameter eq_trans : eq s s' -> eq s' s'' -> eq s s''.
-
- (** Specification of [lt] *)
- Parameter lt_trans : lt s s' -> lt s' s'' -> lt s s''.
- Parameter lt_not_eq : lt s s' -> ~ eq s s'.
(** Specification of [mem] *)
Parameter mem_1 : In x s -> mem x s = true.
Parameter mem_2 : mem x s = true -> In x s.
(** Specification of [equal] *)
- Parameter equal_1 : s[=]s' -> equal s s' = true.
- Parameter equal_2 : equal s s' = true ->s[=]s'.
+ Parameter equal_1 : Equal s s' -> equal s s' = true.
+ Parameter equal_2 : equal s s' = true -> Equal s s'.
(** Specification of [subset] *)
- Parameter subset_1 : s[<=]s' -> subset s s' = true.
- Parameter subset_2 : subset s s' = true -> s[<=]s'.
+ Parameter subset_1 : Subset s s' -> subset s s' = true.
+ Parameter subset_2 : subset s s' = true -> Subset s s'.
(** Specification of [empty] *)
Parameter empty_1 : Empty empty.
@@ -216,7 +207,7 @@ Module Type S.
Parameter diff_3 : In x s -> ~ In x s' -> In x (diff s s').
(** Specification of [fold] *)
- Parameter fold_1 : forall (A : Set) (i : A) (f : elt -> A -> A),
+ Parameter fold_1 : forall (A : Type) (i : A) (f : elt -> A -> A),
fold f s i = fold_left (fun a e => f e a) (elements s) i.
(** Specification of [cardinal] *)
@@ -249,18 +240,93 @@ Module Type S.
exists_ f s = true -> Exists (fun x => f x = true) s.
(** Specification of [partition] *)
- Parameter partition_1 : compat_bool E.eq f ->
- fst (partition f s) [=] filter f s.
- Parameter partition_2 : compat_bool E.eq f ->
- snd (partition f s) [=] filter (fun x => negb (f x)) s.
+ Parameter partition_1 :
+ compat_bool E.eq f -> Equal (fst (partition f s)) (filter f s).
+ Parameter partition_2 :
+ compat_bool E.eq f ->
+ Equal (snd (partition f s)) (filter (fun x => negb (f x)) s).
End Filter.
(** Specification of [elements] *)
Parameter elements_1 : In x s -> InA E.eq x (elements s).
Parameter elements_2 : 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_3w : NoDupA E.eq (elements s).
+
+ (** Specification of [choose] *)
+ Parameter choose_1 : choose s = Some x -> In x s.
+ Parameter choose_2 : choose s = None -> Empty s.
+
+ End Spec.
+
+ 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.
+
+End WSfun.
+
+
+
+(** ** Static signature for weak sets
+
+ Similar to the functorial signature [SW], except that the
+ module [E] of base elements is incorporated in the signature. *)
+
+Module Type WS.
+ Declare Module E : EqualityType.
+ Include Type WSfun E.
+End WS.
+
+
+
+(** ** Functorial signature for sets on ordered elements
+
+ Based on [WSfun], plus ordering on sets and [min_elt] and [max_elt]
+ and some stronger specifications for other functions. *)
+
+Module Type Sfun (E : OrderedType).
+ Include Type WSfun E.
+
+ Parameter lt : t -> t -> Prop.
+ Parameter compare : forall s s' : t, Compare lt eq s s'.
+ (** 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. *)
+
+ Section Spec.
+
+ Variable s s' s'' : t.
+ Variable x y : elt.
+
+ (** Specification of [lt] *)
+ Parameter lt_trans : lt s s' -> lt s' s'' -> lt s s''.
+ Parameter lt_not_eq : lt s s' -> ~ eq s s'.
+
+ (** Additional specification of [elements] *)
Parameter elements_3 : 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.
+ *)
+
(** Specification of [min_elt] *)
Parameter min_elt_1 : min_elt s = Some x -> In x s.
Parameter min_elt_2 : min_elt s = Some x -> In y s -> ~ E.lt y x.
@@ -271,37 +337,56 @@ Module Type S.
Parameter max_elt_2 : max_elt s = Some x -> In y s -> ~ E.lt x y.
Parameter max_elt_3 : max_elt s = None -> Empty s.
- (** Specification of [choose] *)
- Parameter choose_1 : choose s = Some x -> In x s.
- Parameter choose_2 : choose s = None -> Empty s.
-(* Parameter choose_equal:
- (equal s s')=true -> E.eq (choose s) (choose s'). *)
+ (** Additional specification of [choose] *)
+ Parameter choose_3 : choose s = Some x -> choose s' = Some y ->
+ Equal s s' -> E.eq x y.
End Spec.
- (* begin hide *)
- Hint Immediate In_1.
-
- Hint Resolve mem_1 mem_2 equal_1 equal_2 subset_1 subset_2 empty_1
- is_empty_1 is_empty_2 choose_1 choose_2 add_1 add_2 add_3 remove_1
- remove_2 remove_3 singleton_1 singleton_2 union_1 union_2 union_3 inter_1
- inter_2 inter_3 diff_1 diff_2 diff_3 filter_1 filter_2 filter_3 for_all_1
- for_all_2 exists_1 exists_2 partition_1 partition_2 elements_1 elements_2
- elements_3 min_elt_1 min_elt_2 min_elt_3 max_elt_1 max_elt_2 max_elt_3.
- (* end hide *)
+ Hint Resolve elements_3 : set.
+ Hint Immediate
+ min_elt_1 min_elt_2 min_elt_3 max_elt_1 max_elt_2 max_elt_3 : set.
+
+End Sfun.
+
+(** ** Static signature for sets on ordered elements
+
+ Similar to the functorial signature [Sfun], except that the
+ module [E] of base elements is incorporated in the signature. *)
+
+Module Type S.
+ Declare Module E : OrderedType.
+ Include Type Sfun E.
End S.
+
+(** ** Some subtyping tests
+<<
+WSfun ---> WS
+ | |
+ | |
+ V V
+Sfun ---> S
+
+
+Module S_WS (M : S) <: SW := M.
+Module Sfun_WSfun (E:OrderedType)(M : Sfun E) <: WSfun E := M.
+Module S_Sfun (E:OrderedType)(M : S with Module E:=E) <: Sfun E := M.
+Module WS_WSfun (E:EqualityType)(M : WS with Module E:=E) <: WSfun E := M.
+>>
+*)
+
(** * Dependent signature
- Signature [Sdep] presents sets using dependent types *)
+ Signature [Sdep] presents ordered sets using dependent types *)
Module Type Sdep.
Declare Module E : OrderedType.
Definition elt := E.t.
- Parameter t : Set.
+ Parameter t : Type.
Parameter In : elt -> t -> Prop.
Definition Equal s s' := forall a : elt, In a s <-> In a s'.
@@ -397,7 +482,7 @@ Module Type Sdep.
Parameter
fold :
- forall (A : Set) (f : elt -> A -> A) (s : t) (i : A),
+ forall (A : Type) (f : elt -> A -> A) (s : t) (i : A),
{r : A | let (l,_) := elements s in
r = fold_left (fun a e => f e a) l i}.
@@ -418,4 +503,14 @@ Module Type Sdep.
Parameter choose : forall s : t, {x : elt | In x s} + {Empty s}.
+ (** The [choose_3] specification of [S] cannot be packed
+ in the dependent version of [choose], so we leave it separate. *)
+ Parameter choose_equal : forall s s', Equal s s' ->
+ match choose s, choose s' with
+ | inleft (exist x _), inleft (exist x' _) => E.eq x x'
+ | inright _, inright _ => True
+ | _, _ => False
+ end.
+
End Sdep.
+
diff --git a/theories/FSets/FSetList.v b/theories/FSets/FSetList.v
index f6205542..a205d5b0 100644
--- a/theories/FSets/FSetList.v
+++ b/theories/FSets/FSetList.v
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: FSetList.v 8834 2006-05-20 00:41:35Z letouzey $ *)
+(* $Id: FSetList.v 10616 2008-03-04 17:33:35Z letouzey $ *)
(** * Finite sets library *)
@@ -25,7 +25,6 @@ Unset Strict Implicit.
Module Raw (X: OrderedType).
- Module E := X.
Module MX := OrderedTypeFacts X.
Import MX.
@@ -145,7 +144,7 @@ Module Raw (X: OrderedType).
| _, _ => false
end.
- Fixpoint fold (B : Set) (f : elt -> B -> B) (s : t) {struct s} :
+ Fixpoint fold (B : Type) (f : elt -> B -> B) (s : t) {struct s} :
B -> B := fun i => match s with
| nil => i
| x :: l => fold f l (f x i)
@@ -649,6 +648,11 @@ Module Raw (X: OrderedType).
unfold elements; auto.
Qed.
+ Lemma elements_3w : forall (s : t) (Hs : Sort s), NoDupA X.eq (elements s).
+ Proof.
+ unfold elements; auto.
+ Qed.
+
Lemma min_elt_1 : forall (s : t) (x : elt), min_elt s = Some x -> In x s.
Proof.
intro s; case s; simpl; intros; inversion H; auto.
@@ -718,8 +722,21 @@ Module Raw (X: OrderedType).
Definition choose_2 : forall s : t, choose s = None -> Empty s := min_elt_3.
+ Lemma choose_3: forall s s', Sort s -> Sort s' -> forall x x',
+ choose s = Some x -> choose s' = Some x' -> Equal s s' -> X.eq x x'.
+ Proof.
+ unfold choose, Equal; intros s s' Hs Hs' x x' Hx Hx' H.
+ assert (~X.lt x x').
+ apply min_elt_2 with s'; auto.
+ rewrite <-H; auto using min_elt_1.
+ assert (~X.lt x' x).
+ apply min_elt_2 with s; auto.
+ rewrite H; auto using min_elt_1.
+ destruct (X.compare x x'); intuition.
+ Qed.
+
Lemma fold_1 :
- forall (s : t) (Hs : Sort s) (A : Set) (i : A) (f : elt -> A -> A),
+ forall (s : t) (Hs : Sort 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.
induction s.
@@ -1037,7 +1054,7 @@ Module Make (X: OrderedType) <: S with Module E := X.
Module Raw := Raw X.
Module E := X.
- Record slist : Set := {this :> Raw.t; sorted : sort E.lt this}.
+ Record slist := {this :> Raw.t; sorted : sort E.lt this}.
Definition t := slist.
Definition elt := E.t.
@@ -1066,7 +1083,7 @@ Module Make (X: OrderedType) <: S with Module E := X.
Definition min_elt (s : t) : option elt := Raw.min_elt s.
Definition max_elt (s : t) : option elt := Raw.max_elt s.
Definition choose (s : t) : option elt := Raw.choose s.
- Definition fold (B : Set) (f : elt -> B -> B) (s : t) : B -> B := Raw.fold (B:=B) f s.
+ Definition fold (B : Type) (f : elt -> B -> B) (s : t) : B -> B := Raw.fold (B:=B) f s.
Definition cardinal (s : t) : nat := Raw.cardinal s.
Definition filter (f : elt -> bool) (s : t) : t :=
Build_slist (Raw.filter_sort (sorted s) f).
@@ -1149,7 +1166,7 @@ Module Make (X: OrderedType) <: S with Module E := X.
Lemma diff_3 : In x s -> ~ In x s' -> In x (diff s s').
Proof. exact (fun H => Raw.diff_3 s.(sorted) s'.(sorted) H). Qed.
- Lemma fold_1 : forall (A : Set) (i : A) (f : elt -> A -> A),
+ Lemma fold_1 : 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 (Raw.fold_1 s.(sorted)). Qed.
@@ -1202,6 +1219,8 @@ Module Make (X: OrderedType) <: S with Module E := X.
Proof. exact (fun H => Raw.elements_2 H). Qed.
Lemma elements_3 : sort E.lt (elements s).
Proof. exact (Raw.elements_3 s.(sorted)). Qed.
+ Lemma elements_3w : NoDupA E.eq (elements s).
+ Proof. exact (Raw.elements_3w s.(sorted)). Qed.
Lemma min_elt_1 : min_elt s = Some x -> In x s.
Proof. exact (fun H => Raw.min_elt_1 H). Qed.
@@ -1221,6 +1240,9 @@ Module Make (X: OrderedType) <: S with Module E := X.
Proof. exact (fun H => Raw.choose_1 H). Qed.
Lemma choose_2 : choose s = None -> Empty s.
Proof. exact (fun H => Raw.choose_2 H). Qed.
+ Lemma choose_3 : choose s = Some x -> choose s' = Some y ->
+ Equal s s' -> E.eq x y.
+ Proof. exact (@Raw.choose_3 _ _ s.(sorted) s'.(sorted) x y). Qed.
Lemma eq_refl : eq s s.
Proof. exact (Raw.eq_refl s). Qed.
diff --git a/theories/FSets/FSetProperties.v b/theories/FSets/FSetProperties.v
index 6e93a546..7413b06b 100644
--- a/theories/FSets/FSetProperties.v
+++ b/theories/FSets/FSetProperties.v
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: FSetProperties.v 8853 2006-05-23 18:17:38Z herbelin $ *)
+(* $Id: FSetProperties.v 11064 2008-06-06 17:00:52Z letouzey $ *)
(** * Finite sets library *)
@@ -16,414 +16,259 @@
[In x s] instead of [mem x s=true],
[Equal s s'] instead of [equal s s'=true], etc. *)
-Require Export FSetInterface.
-Require Import FSetFacts.
+Require Export FSetInterface.
+Require Import DecidableTypeEx FSetFacts FSetDecide.
Set Implicit Arguments.
Unset Strict Implicit.
-Hint Unfold transpose compat_op compat_nat.
+Hint Unfold transpose compat_op.
Hint Extern 1 (Setoid_Theory _ _) => constructor; congruence.
-Module Properties (M: S).
- Module ME:=OrderedTypeFacts(M.E).
- Import ME. (* for ME.eq_dec *)
- Import M.E.
- Import M.
- Import Logic. (* to unmask [eq] *)
- Import Peano. (* to unmask [lt] *)
-
- (** Results about lists without duplicates *)
+(** First, a functor for Weak Sets. Since the signature [WS] includes
+ an EqualityType and not a stronger DecidableType, this functor
+ should take two arguments in order to compensate this. *)
- Module FM := Facts M.
- Import FM.
-
- Definition Add (x : elt) (s s' : t) :=
- forall y : elt, In y s' <-> E.eq x y \/ In y s.
+Module WProperties (Import E : DecidableType)(M : WSfun E).
+ Module Import Dec := WDecide E M.
+ Module Import FM := Dec.F (* FSetFacts.WFacts 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.
- Section BasicProperties.
-
- (** properties of [Equal] *)
+ Definition Add x s s' := forall y, In y s' <-> E.eq x y \/ In y s.
- Lemma equal_refl : forall s, s[=]s.
+ Lemma Add_Equal : forall x s s', Add x s s' <-> s' [=] add x s.
Proof.
- unfold Equal; intuition.
- Qed.
-
- Lemma equal_sym : forall s s', s[=]s' -> s'[=]s.
- Proof.
- unfold Equal; intros.
- rewrite H; intuition.
+ unfold Add.
+ split; intros.
+ red; intros.
+ rewrite H; clear H.
+ fsetdec.
+ fsetdec.
Qed.
+
+ Ltac expAdd := repeat rewrite Add_Equal.
- Lemma equal_trans : forall s1 s2 s3, s1[=]s2 -> s2[=]s3 -> s1[=]s3.
- Proof.
- unfold Equal; intros.
- rewrite H; exact (H0 a).
- Qed.
+ Section BasicProperties.
Variable s s' s'' s1 s2 s3 : t.
Variable x x' : elt.
- (** properties of [Subset] *)
-
- Lemma subset_refl : s[<=]s.
- Proof.
- unfold Subset; intuition.
- Qed.
+ Lemma equal_refl : s[=]s.
+ Proof. fsetdec. Qed.
- Lemma subset_antisym : s[<=]s' -> s'[<=]s -> s[=]s'.
- Proof.
- unfold Subset, Equal; intuition.
- 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.
- unfold Subset; intuition.
- Qed.
+ Proof. fsetdec. Qed.
+
+ Lemma subset_antisym : s[<=]s' -> s'[<=]s -> s[=]s'.
+ Proof. fsetdec. Qed.
Lemma subset_equal : s[=]s' -> s[<=]s'.
- Proof.
- unfold Subset, Equal; firstorder.
- Qed.
+ Proof. fsetdec. Qed.
Lemma subset_empty : empty[<=]s.
- Proof.
- unfold Subset; intros a; set_iff; intuition.
- Qed.
+ Proof. fsetdec. Qed.
Lemma subset_remove_3 : s1[<=]s2 -> remove x s1 [<=] s2.
- Proof.
- unfold Subset; intros H a; set_iff; intuition.
- Qed.
+ Proof. fsetdec. Qed.
Lemma subset_diff : s1[<=]s3 -> diff s1 s2 [<=] s3.
- Proof.
- unfold Subset; intros H a; set_iff; intuition.
- Qed.
-
+ Proof. fsetdec. Qed.
+
Lemma subset_add_3 : In x s2 -> s1[<=]s2 -> add x s1 [<=] s2.
- Proof.
- unfold Subset; intros H H0 a; set_iff; intuition.
- rewrite <- H2; auto.
- Qed.
+ Proof. fsetdec. Qed.
Lemma subset_add_2 : s1[<=]s2 -> s1[<=] add x s2.
- Proof.
- unfold Subset; intuition.
- Qed.
+ Proof. fsetdec. Qed.
Lemma in_subset : In x s1 -> s1[<=]s2 -> In x s2.
- Proof.
- unfold Subset; intuition.
- Qed.
+ Proof. fsetdec. Qed.
Lemma double_inclusion : s1[=]s2 <-> s1[<=]s2 /\ s2[<=]s1.
- Proof.
- unfold Subset, Equal; split; intros; intuition; generalize (H a); intuition.
- Qed.
-
- (** properties of [empty] *)
+ Proof. intuition fsetdec. Qed.
Lemma empty_is_empty_1 : Empty s -> s[=]empty.
- Proof.
- unfold Empty, Equal; intros; generalize (H a); set_iff; tauto.
- Qed.
+ Proof. fsetdec. Qed.
Lemma empty_is_empty_2 : s[=]empty -> Empty s.
- Proof.
- unfold Empty, Equal; intros; generalize (H a); set_iff; tauto.
- Qed.
-
- (** properties of [add] *)
+ Proof. fsetdec. Qed.
Lemma add_equal : In x s -> add x s [=] s.
- Proof.
- unfold Equal; intros; set_iff; intuition.
- rewrite <- H1; auto.
- Qed.
+ Proof. fsetdec. Qed.
Lemma add_add : add x (add x' s) [=] add x' (add x s).
- Proof.
- unfold Equal; intros; set_iff; tauto.
- Qed.
-
- (** properties of [remove] *)
+ Proof. fsetdec. Qed.
Lemma remove_equal : ~ In x s -> remove x s [=] s.
- Proof.
- unfold Equal; intros; set_iff; intuition.
- rewrite H1 in H; auto.
- Qed.
+ Proof. fsetdec. Qed.
Lemma Equal_remove : s[=]s' -> remove x s [=] remove x s'.
- Proof.
- intros; rewrite H; apply equal_refl.
- Qed.
-
- (** properties of [add] and [remove] *)
+ Proof. fsetdec. Qed.
Lemma add_remove : In x s -> add x (remove x s) [=] s.
- Proof.
- unfold Equal; intros; set_iff; elim (eq_dec x a); intuition.
- rewrite <- H1; auto.
- Qed.
+ Proof. fsetdec. Qed.
Lemma remove_add : ~In x s -> remove x (add x s) [=] s.
- Proof.
- unfold Equal; intros; set_iff; elim (eq_dec x a); intuition.
- rewrite H1 in H; auto.
- Qed.
-
- (** properties of [singleton] *)
+ Proof. fsetdec. Qed.
Lemma singleton_equal_add : singleton x [=] add x empty.
- Proof.
- unfold Equal; intros; set_iff; intuition.
- Qed.
-
- (** properties of [union] *)
+ Proof. fsetdec. Qed.
Lemma union_sym : union s s' [=] union s' s.
- Proof.
- unfold Equal; intros; set_iff; tauto.
- Qed.
+ Proof. fsetdec. Qed.
Lemma union_subset_equal : s[<=]s' -> union s s' [=] s'.
- Proof.
- unfold Subset, Equal; intros; set_iff; intuition.
- Qed.
+ Proof. fsetdec. Qed.
Lemma union_equal_1 : s[=]s' -> union s s'' [=] union s' s''.
- Proof.
- intros; rewrite H; apply equal_refl.
- Qed.
+ Proof. fsetdec. Qed.
Lemma union_equal_2 : s'[=]s'' -> union s s' [=] union s s''.
- Proof.
- intros; rewrite H; apply equal_refl.
- Qed.
+ Proof. fsetdec. Qed.
Lemma union_assoc : union (union s s') s'' [=] union s (union s' s'').
- Proof.
- unfold Equal; intros; set_iff; tauto.
- Qed.
+ Proof. fsetdec. Qed.
Lemma add_union_singleton : add x s [=] union (singleton x) s.
- Proof.
- unfold Equal; intros; set_iff; tauto.
- Qed.
+ Proof. fsetdec. Qed.
Lemma union_add : union (add x s) s' [=] add x (union s s').
- Proof.
- unfold Equal; intros; set_iff; tauto.
- Qed.
+ 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.
- unfold Subset; intuition.
- Qed.
+ Proof. fsetdec. Qed.
Lemma union_subset_2 : s' [<=] union s s'.
- Proof.
- unfold Subset; intuition.
- Qed.
+ Proof. fsetdec. Qed.
Lemma union_subset_3 : s[<=]s'' -> s'[<=]s'' -> union s s' [<=] s''.
- Proof.
- unfold Subset; intros H H0 a; set_iff; intuition.
- Qed.
+ Proof. fsetdec. Qed.
Lemma union_subset_4 : s[<=]s' -> union s s'' [<=] union s' s''.
- Proof.
- unfold Subset; intros H a; set_iff; intuition.
- Qed.
+ Proof. fsetdec. Qed.
Lemma union_subset_5 : s[<=]s' -> union s'' s [<=] union s'' s'.
- Proof.
- unfold Subset; intros H a; set_iff; intuition.
- Qed.
+ Proof. fsetdec. Qed.
Lemma empty_union_1 : Empty s -> union s s' [=] s'.
- Proof.
- unfold Equal, Empty; intros; set_iff; firstorder.
- Qed.
+ Proof. fsetdec. Qed.
Lemma empty_union_2 : Empty s -> union s' s [=] s'.
- Proof.
- unfold Equal, Empty; intros; set_iff; firstorder.
- Qed.
+ Proof. fsetdec. Qed.
Lemma not_in_union : ~In x s -> ~In x s' -> ~In x (union s s').
- Proof.
- intros; set_iff; intuition.
- Qed.
-
- (** properties of [inter] *)
+ Proof. fsetdec. Qed.
Lemma inter_sym : inter s s' [=] inter s' s.
- Proof.
- unfold Equal; intros; set_iff; tauto.
- Qed.
+ Proof. fsetdec. Qed.
Lemma inter_subset_equal : s[<=]s' -> inter s s' [=] s.
- Proof.
- unfold Equal; intros; set_iff; intuition.
- Qed.
+ Proof. fsetdec. Qed.
Lemma inter_equal_1 : s[=]s' -> inter s s'' [=] inter s' s''.
- Proof.
- intros; rewrite H; apply equal_refl.
- Qed.
+ Proof. fsetdec. Qed.
Lemma inter_equal_2 : s'[=]s'' -> inter s s' [=] inter s s''.
- Proof.
- intros; rewrite H; apply equal_refl.
- Qed.
+ Proof. fsetdec. Qed.
Lemma inter_assoc : inter (inter s s') s'' [=] inter s (inter s' s'').
- Proof.
- unfold Equal; intros; set_iff; tauto.
- Qed.
+ Proof. fsetdec. Qed.
Lemma union_inter_1 : inter (union s s') s'' [=] union (inter s s'') (inter s' s'').
- Proof.
- unfold Equal; intros; set_iff; tauto.
- Qed.
+ Proof. fsetdec. Qed.
Lemma union_inter_2 : union (inter s s') s'' [=] inter (union s s'') (union s' s'').
- Proof.
- unfold Equal; intros; set_iff; tauto.
- Qed.
+ Proof. fsetdec. Qed.
Lemma inter_add_1 : In x s' -> inter (add x s) s' [=] add x (inter s s').
- Proof.
- unfold Equal; intros; set_iff; intuition.
- rewrite <- H1; auto.
- Qed.
+ Proof. fsetdec. Qed.
Lemma inter_add_2 : ~ In x s' -> inter (add x s) s' [=] inter s s'.
- Proof.
- unfold Equal; intros; set_iff; intuition.
- destruct H; rewrite H0; auto.
- Qed.
+ Proof. fsetdec. Qed.
Lemma empty_inter_1 : Empty s -> Empty (inter s s').
- Proof.
- unfold Empty; intros; set_iff; firstorder.
- Qed.
+ Proof. fsetdec. Qed.
Lemma empty_inter_2 : Empty s' -> Empty (inter s s').
- Proof.
- unfold Empty; intros; set_iff; firstorder.
- Qed.
+ Proof. fsetdec. Qed.
Lemma inter_subset_1 : inter s s' [<=] s.
- Proof.
- unfold Subset; intro a; set_iff; tauto.
- Qed.
+ Proof. fsetdec. Qed.
Lemma inter_subset_2 : inter s s' [<=] s'.
- Proof.
- unfold Subset; intro a; set_iff; tauto.
- Qed.
+ Proof. fsetdec. Qed.
Lemma inter_subset_3 :
s''[<=]s -> s''[<=]s' -> s''[<=] inter s s'.
- Proof.
- unfold Subset; intros H H' a; set_iff; intuition.
- Qed.
-
- (** properties of [diff] *)
+ Proof. fsetdec. Qed.
Lemma empty_diff_1 : Empty s -> Empty (diff s s').
- Proof.
- unfold Empty, Equal; intros; set_iff; firstorder.
- Qed.
+ Proof. fsetdec. Qed.
Lemma empty_diff_2 : Empty s -> diff s' s [=] s'.
- Proof.
- unfold Empty, Equal; intros; set_iff; firstorder.
- Qed.
+ Proof. fsetdec. Qed.
Lemma diff_subset : diff s s' [<=] s.
- Proof.
- unfold Subset; intros a; set_iff; tauto.
- Qed.
+ Proof. fsetdec. Qed.
Lemma diff_subset_equal : s[<=]s' -> diff s s' [=] empty.
- Proof.
- unfold Subset, Equal; intros; set_iff; intuition; absurd (In a empty); auto.
- Qed.
+ Proof. fsetdec. Qed.
Lemma remove_diff_singleton :
remove x s [=] diff s (singleton x).
- Proof.
- unfold Equal; intros; set_iff; intuition.
- Qed.
+ Proof. fsetdec. Qed.
Lemma diff_inter_empty : inter (diff s s') (inter s s') [=] empty.
- Proof.
- unfold Equal; intros; set_iff; intuition; absurd (In a empty); auto.
- Qed.
+ Proof. fsetdec. Qed.
Lemma diff_inter_all : union (diff s s') (inter s s') [=] s.
- Proof.
- unfold Equal; intros; set_iff; intuition.
- elim (In_dec a s'); auto.
- Qed.
-
- (** properties of [Add] *)
+ Proof. fsetdec. Qed.
Lemma Add_add : Add x s (add x s).
- Proof.
- unfold Add; intros; set_iff; intuition.
- Qed.
+ Proof. expAdd; fsetdec. Qed.
Lemma Add_remove : In x s -> Add x (remove x s) s.
- Proof.
- unfold Add; intros; set_iff; intuition.
- elim (eq_dec x y); auto.
- rewrite <- H1; auto.
- Qed.
+ Proof. expAdd; fsetdec. Qed.
Lemma union_Add : Add x s s' -> Add x (union s s'') (union s' s'').
- Proof.
- unfold Add; intros; set_iff; rewrite H; tauto.
- Qed.
+ Proof. expAdd; fsetdec. Qed.
Lemma inter_Add :
In x s'' -> Add x s s' -> Add x (inter s s'') (inter s' s'').
- Proof.
- unfold Add; intros; set_iff; rewrite H0; intuition.
- rewrite <- H2; auto.
- Qed.
+ Proof. expAdd; fsetdec. Qed.
Lemma union_Equal :
In x s'' -> Add x s s' -> union s s'' [=] union s' s''.
- Proof.
- unfold Add, Equal; intros; set_iff; rewrite H0; intuition.
- rewrite <- H1; auto.
- Qed.
+ Proof. expAdd; fsetdec. Qed.
Lemma inter_Add_2 :
~In x s'' -> Add x s s' -> inter s s'' [=] inter s' s''.
- Proof.
- unfold Add, Equal; intros; set_iff; rewrite H0; intuition.
- destruct H; rewrite H1; auto.
- Qed.
+ Proof. expAdd; fsetdec. Qed.
End BasicProperties.
- Hint Immediate equal_sym: set.
- Hint Resolve equal_refl equal_trans : set.
-
- Hint Immediate add_remove remove_add union_sym inter_sym: set.
- Hint Resolve subset_refl subset_equal subset_antisym
+ 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
@@ -436,6 +281,31 @@ Module Properties (M: S).
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.
+ 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.
+
(** * Alternative (weaker) specifications for [fold] *)
Section Old_Spec_Now_Properties.
@@ -447,14 +317,14 @@ Module Properties (M: S).
*)
Lemma fold_0 :
- forall s (A : Set) (i : A) (f : elt -> A -> A),
+ forall s (A : Type) (i : A) (f : elt -> A -> A),
exists l : list elt,
NoDup l /\
(forall x : elt, In x s <-> InA E.eq x l) /\
fold f s i = fold_right f i l.
Proof.
intros; exists (rev (elements s)); split.
- apply NoDupA_rev; auto.
+ apply NoDupA_rev; auto with set.
exact E.eq_trans.
split; intros.
rewrite elements_iff; do 2 rewrite InA_alt.
@@ -468,7 +338,7 @@ Module Properties (M: S).
[fold_2]. *)
Lemma fold_1 :
- forall s (A : Set) (eqA : A -> A -> Prop)
+ forall s (A : Type) (eqA : A -> A -> Prop)
(st : Setoid_Theory A eqA) (i : A) (f : elt -> A -> A),
Empty s -> eqA (fold f s i) i.
Proof.
@@ -481,7 +351,7 @@ Module Properties (M: S).
Qed.
Lemma fold_2 :
- forall s s' x (A : Set) (eqA : A -> A -> Prop)
+ forall s s' x (A : Type) (eqA : A -> A -> Prop)
(st : Setoid_Theory A eqA) (i : A) (f : elt -> A -> A),
compat_op E.eq eqA f ->
transpose eqA f ->
@@ -492,9 +362,21 @@ Module Properties (M: S).
rewrite Hl2; rewrite Hl'2; clear Hl2 Hl'2.
apply fold_right_add with (eqA:=E.eq)(eqB:=eqA); auto.
eauto.
- exact eq_dec.
rewrite <- Hl1; auto.
- intros; rewrite <- Hl1; rewrite <- Hl'1; 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 M.fold_1.
+ rewrite elements_Empty in H; rewrite H; simpl; auto.
Qed.
(** Similar specifications for [cardinal]. *)
@@ -531,41 +413,46 @@ Module Properties (M: S).
(** * Induction principle over sets *)
+ Lemma cardinal_Empty : forall s, Empty s <-> cardinal s = 0.
+ Proof.
+ intros.
+ rewrite elements_Empty, M.cardinal_1.
+ destruct (elements s); intuition; discriminate.
+ Qed.
+
Lemma cardinal_inv_1 : forall s, cardinal s = 0 -> Empty s.
- Proof.
- intros s; rewrite M.cardinal_1; intros H a; red.
- rewrite elements_iff.
- destruct (elements s); simpl in *; discriminate || inversion 1.
+ 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 M.cardinal_1 in H.
- generalize (elements_2 (s:=s)).
- destruct (elements s); try discriminate.
- exists e; auto.
+ intros; rewrite M.cardinal_1 in H.
+ generalize (elements_2 (s:=s)).
+ destruct (elements s); try discriminate.
+ exists e; auto.
Qed.
- Lemma Equal_cardinal_aux :
- forall n s s', cardinal s = n -> s[=]s' -> cardinal s = cardinal s'.
+ Lemma cardinal_inv_2b :
+ forall s, cardinal s <> 0 -> { x : elt | In x s }.
Proof.
- simple induction n; intros.
- rewrite H; symmetry .
- apply cardinal_1.
- rewrite <- H0; auto.
- destruct (cardinal_inv_2 H0) as (x,H2).
- revert H0.
- rewrite (cardinal_2 (s:=remove x s) (s':=s) (x:=x)); auto with set.
- rewrite (cardinal_2 (s:=remove x s') (s':=s') (x:=x)); auto with set.
- rewrite H1 in H2; auto with set.
+ intro; generalize (@cardinal_inv_2 s); destruct cardinal;
+ [intuition|eauto].
Qed.
Lemma Equal_cardinal : forall s s', s[=]s' -> cardinal s = cardinal s'.
Proof.
- intros; apply Equal_cardinal_aux with (cardinal s); auto.
- Qed.
+ 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.
+ rewrite (cardinal_2 (s:=remove x s') (s':=s') (x:=x)); eauto with set.
+ Qed.
Add Morphism cardinal : cardinal_m.
Proof.
@@ -574,40 +461,33 @@ Module Properties (M: S).
Hint Resolve Add_add Add_remove Equal_remove cardinal_inv_1 Equal_cardinal.
- Lemma cardinal_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 n s, cardinal s = n -> P s.
- Proof.
- simple induction n; intros; auto.
- destruct (cardinal_inv_2 H) as (x,H0).
- apply X0 with (remove x s) x; auto.
- apply X1; auto.
- rewrite (cardinal_2 (x:=x)(s:=remove x s)(s':=s)) in H; auto.
- Qed.
-
Lemma set_induction :
forall P : t -> Type,
(forall s : t, Empty s -> P s) ->
(forall s s' : t, P s -> forall x : elt, ~In x s -> Add x s s' -> P s') ->
forall s : t, P s.
Proof.
- intros; apply cardinal_induction with (cardinal s); auto.
- Qed.
+ intros; remember (cardinal s) as n; revert s Heqn; induction n; intros; auto.
+ destruct (cardinal_inv_2 (sym_eq Heqn)) as (x,H0).
+ apply X0 with (remove x s) x; auto with set.
+ apply IHn; auto.
+ assert (S n = S (cardinal (remove x s))).
+ rewrite Heqn; apply cardinal_2 with x; auto with set.
+ inversion H; auto.
+ Qed.
(** Other properties of [fold]. *)
Section Fold.
- Variables (A:Set)(eqA:A->A->Prop)(st:Setoid_Theory _ eqA).
+ Variables (A:Type)(eqA:A->A->Prop)(st:Setoid_Theory _ eqA).
Variables (f:elt->A->A)(Comp:compat_op E.eq eqA f)(Ass:transpose eqA f).
Section Fold_1.
Variable i i':A.
- Lemma fold_empty : eqA (fold f empty i) i.
+ Lemma fold_empty : (fold f empty i) = i.
Proof.
- apply fold_1; auto.
+ apply fold_1b; auto with set.
Qed.
Lemma fold_equal :
@@ -642,7 +522,7 @@ Module Properties (M: S).
Proof.
intros.
sym_st.
- apply fold_2 with (eqA:=eqA); auto.
+ apply fold_2 with (eqA:=eqA); auto with set.
Qed.
Lemma remove_fold_2: forall s x, ~In x s ->
@@ -742,7 +622,8 @@ Module Properties (M: S).
apply fold_1; auto with set.
Qed.
- Lemma fold_union: forall s s', (forall x, ~In x s\/~In x s') ->
+ Lemma fold_union: forall 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.
@@ -760,8 +641,8 @@ Module Properties (M: S).
forall s p, fold (fun _ => S) s p = fold (fun _ => S) s 0 + p.
Proof.
assert (st := gen_st nat).
- assert (fe : compat_op E.eq (@eq _) (fun _ => S)) by (unfold compat_op; auto).
- assert (fp : transpose (@eq _) (fun _:elt => S)) by (unfold transpose; auto).
+ assert (fe : compat_op E.eq (@Logic.eq _) (fun _ => S)) by (unfold compat_op; auto).
+ assert (fp : transpose (@Logic.eq _) (fun _:elt => S)) by (unfold transpose; auto).
intros s p; pattern s; apply set_induction; clear s; intros.
rewrite (fold_1 st p (fun _ => S) H).
rewrite (fold_1 st 0 (fun _ => S) H); trivial.
@@ -774,11 +655,11 @@ Module Properties (M: S).
simpl; auto.
Qed.
- (** properties of [cardinal] *)
+ (** more properties of [cardinal] *)
Lemma empty_cardinal : cardinal empty = 0.
Proof.
- rewrite cardinal_fold; apply fold_1; auto.
+ rewrite cardinal_fold; apply fold_1; auto with set.
Qed.
Hint Immediate empty_cardinal cardinal_1 : set.
@@ -798,11 +679,11 @@ Module Properties (M: S).
Proof.
intros; do 3 rewrite cardinal_fold.
rewrite <- fold_plus.
- apply fold_diff_inter with (eqA:=@eq nat); auto.
+ apply fold_diff_inter with (eqA:=@Logic.eq nat); auto.
Qed.
Lemma union_cardinal:
- forall s s', (forall x, ~In x s\/~In x s') ->
+ 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.
@@ -841,7 +722,7 @@ Module Properties (M: S).
intros.
do 4 rewrite cardinal_fold.
do 2 rewrite <- fold_plus.
- apply fold_union_inter with (eqA:=@eq nat); auto.
+ apply fold_union_inter with (eqA:=@Logic.eq nat); auto.
Qed.
Lemma union_cardinal_inter :
@@ -872,7 +753,7 @@ Module Properties (M: S).
intros.
do 2 rewrite cardinal_fold.
change S with ((fun _ => S) x);
- apply fold_add with (eqA:=@eq nat); auto.
+ apply fold_add with (eqA:=@Logic.eq nat); auto.
Qed.
Lemma remove_cardinal_1 :
@@ -881,7 +762,7 @@ Module Properties (M: S).
intros.
do 2 rewrite cardinal_fold.
change S with ((fun _ =>S) x).
- apply remove_fold_1 with (eqA:=@eq nat); auto.
+ apply remove_fold_1 with (eqA:=@Logic.eq nat); auto.
Qed.
Lemma remove_cardinal_2 :
@@ -892,4 +773,295 @@ Module Properties (M: S).
Hint Resolve subset_cardinal union_cardinal add_cardinal_1 add_cardinal_2.
+End WProperties.
+
+
+(** A clone of [WProperties] working on full sets. *)
+
+Module Properties (M:S).
+ Module D := OT_as_DT M.E.
+ Include WProperties D M.
End Properties.
+
+
+(** Now comes some properties specific to the element ordering,
+ invalid for Weak Sets. *)
+
+Module OrdProperties (M:S).
+ Module ME:=OrderedTypeFacts(M.E).
+ Module Import P := Properties M.
+ Import FM.
+ Import M.E.
+ Import M.
+
+ (** 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.
+ 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; unfold gtb; destruct (E.compare x y); intuition; try discriminate; ME.order.
+ Qed.
+
+ Lemma leb_1 : forall x y, leb x y = true <-> ~E.lt y x.
+ Proof.
+ intros; unfold leb, gtb; destruct (E.compare x y); intuition; try discriminate; ME.order.
+ Qed.
+
+ Lemma gtb_compat : forall x, compat_bool E.eq (gtb x).
+ Proof.
+ red; intros x a b H.
+ generalize (gtb_1 x a)(gtb_1 x b); destruct (gtb x a); destruct (gtb x b); auto.
+ intros.
+ symmetry; rewrite H1.
+ apply ME.eq_lt with a; auto.
+ rewrite <- H0; auto.
+ intros.
+ rewrite H0.
+ apply ME.eq_lt with b; auto.
+ rewrite <- H1; auto.
+ Qed.
+
+ Lemma leb_compat : forall x, compat_bool E.eq (leb x).
+ Proof.
+ red; intros x a b H; unfold leb.
+ f_equal; apply gtb_compat; 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 set. ME.order. ME.order. ME.order.
+ intros.
+ rewrite gtb_1 in H.
+ assert (~E.lt y x).
+ unfold gtb in *; destruct (E.compare x y); intuition; try discriminate; ME.order.
+ ME.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.
+ apply (@filter_sort _ E.eq); auto with set; eauto with set.
+ constructor; auto.
+ apply (@filter_sort _ E.eq); auto with set; eauto with set.
+ rewrite ME.Inf_alt by (apply (@filter_sort _ E.eq); eauto with set).
+ intros.
+ rewrite filter_InA in H1; auto; destruct H1.
+ rewrite leb_1 in H2.
+ rewrite <- elements_iff in H1.
+ assert (~E.eq x y).
+ contradict H; rewrite H; auto.
+ ME.order.
+ intros.
+ rewrite filter_InA in H1; auto; destruct H1.
+ rewrite gtb_1 in H3.
+ inversion_clear H2.
+ ME.order.
+ rewrite filter_InA in H4; auto; destruct H4.
+ rewrite leb_1 in H4.
+ ME.order.
+ red; intros a.
+ rewrite InA_app_iff; rewrite InA_cons.
+ do 2 (rewrite filter_InA; auto).
+ do 2 rewrite <- elements_iff.
+ rewrite leb_1; rewrite gtb_1.
+ rewrite (H0 a); intuition.
+ destruct (E.compare a x); intuition.
+ right; right; split; auto.
+ ME.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 set.
+ intros.
+ inversion_clear H2.
+ rewrite <- elements_iff in H1.
+ apply ME.lt_eq with x; auto.
+ inversion H3.
+ red; intros a.
+ rewrite InA_app_iff; rewrite InA_cons; rewrite InA_nil.
+ do 2 rewrite <- elements_iff; rewrite (H0 a); 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 set.
+ intros.
+ inversion_clear H1.
+ rewrite <- elements_iff in H2.
+ apply ME.eq_lt with x; auto.
+ inversion H3.
+ red; intros a.
+ rewrite InA_cons.
+ do 2 rewrite <- elements_iff; rewrite (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.
+ inversion H0; auto.
+ red; intros.
+ rewrite remove_iff in H0; destruct H0.
+ generalize (@max_elt_2 s e y H H0); ME.order.
+
+ assert (H0:=max_elt_3 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.
+ inversion H0; auto.
+ red; intros.
+ rewrite remove_iff in H0; destruct H0.
+ generalize (@min_elt_2 s e y H H0); ME.order.
+
+ assert (H0:=min_elt_3 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 : Setoid_Theory A eqA) (i : A) (f : elt -> A -> A),
+ compat_op E.eq eqA f ->
+ Above x s -> Add x s s' -> eqA (fold f s' i) (f x (fold f s i)).
+ Proof.
+ intros.
+ do 2 rewrite M.fold_1.
+ do 2 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.
+ 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 : Setoid_Theory A eqA) (i : A) (f : elt -> A -> A),
+ compat_op E.eq eqA f ->
+ Below x s -> Add x s s' -> eqA (fold f s' i) (fold f s (f x i)).
+ Proof.
+ intros.
+ do 2 rewrite M.fold_1.
+ set (g:=fun (a : A) (e : elt) => f e a).
+ change (eqA (fold_left g (elements s') i) (fold_left g (x::elements s) i)).
+ unfold g.
+ do 2 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:Setoid_Theory _ eqA).
+ Variables (f:elt->A->A)(Comp:compat_op E.eq eqA f).
+
+ Lemma fold_equal :
+ forall i s s', s[=]s' -> eqA (fold f s i) (fold f s' i).
+ Proof.
+ intros; do 2 rewrite M.fold_1.
+ do 2 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 fold_init : forall i i' s, eqA i i' ->
+ eqA (fold f s i) (fold f s i').
+ Proof.
+ intros; do 2 rewrite M.fold_1.
+ do 2 rewrite <- fold_left_rev_right.
+ induction (rev (elements s)); simpl; 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_1 s)(@choose_2 s)
+ (@choose_1 s')(@choose_2 s')(@choose_3 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/FSets/FSetToFiniteSet.v b/theories/FSets/FSetToFiniteSet.v
index 8cf85efe..ae51d905 100644
--- a/theories/FSets/FSetToFiniteSet.v
+++ b/theories/FSets/FSetToFiniteSet.v
@@ -11,16 +11,16 @@
* Institution: LRI, CNRS UMR 8623 - Université Paris Sud
* 91405 Orsay, France *)
-(* $Id: FSetToFiniteSet.v 8876 2006-05-30 13:43:15Z letouzey $ *)
+(* $Id: FSetToFiniteSet.v 10739 2008-04-01 14:45:20Z herbelin $ *)
Require Import Ensembles Finite_sets.
-Require Import FSetInterface FSetProperties OrderedTypeEx.
+Require Import FSetInterface FSetProperties OrderedTypeEx DecidableTypeEx.
-(** * Going from [FSets] with usual equality
- to the old [Ensembles] and [Finite_sets] theory. *)
+(** * Going from [FSets] with usual Leibniz equality
+ to the good old [Ensembles] and [Finite_sets] theory. *)
-Module S_to_Finite_set (U:UsualOrderedType)(M:S with Module E := U).
- Module MP:= Properties(M).
+Module WS_to_Finite_set (U:UsualDecidableType)(M: WSfun U).
+ Module MP:= WProperties U M.
Import M MP FM Ensembles Finite_sets.
Definition mkEns : M.t -> Ensemble M.elt :=
@@ -82,7 +82,7 @@ Module S_to_Finite_set (U:UsualOrderedType)(M:S with Module E := U).
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; unfold E.eq; auto with sets.
+ split; intro; set_iff; inversion 1; auto with sets.
inversion H0.
constructor 2; constructor.
constructor 1; auto.
@@ -97,7 +97,7 @@ Module S_to_Finite_set (U:UsualOrderedType)(M:S with Module E := U).
inversion H0.
constructor 2; constructor.
constructor 1; auto.
- red in H; rewrite H; unfold E.eq in *.
+ red in H; rewrite H.
inversion H0; auto.
inversion H1; auto.
Qed.
@@ -105,10 +105,10 @@ Module S_to_Finite_set (U:UsualOrderedType)(M:S with Module E := U).
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; unfold E.eq in *; auto with sets.
+ split; intro; set_iff; inversion 1; auto with sets.
split; auto.
- swap H1.
- inversion H2; auto.
+ contradict H1.
+ inversion H1; auto.
Qed.
Lemma mkEns_Finite : forall s, Finite _ (!!s).
@@ -136,4 +136,28 @@ Module S_to_Finite_set (U:UsualOrderedType)(M:S with Module E := U).
apply Add_Add; auto.
Qed.
+ (** we can even build a function from Finite Ensemble to FSet
+ ... at least in Prop. *)
+
+ Lemma Ens_to_FSet : 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: Sfun U).
+ Module D := OT_as_DT U.
+ Include WS_to_Finite_set D M.
End S_to_Finite_set.
+
+
diff --git a/theories/FSets/FSetWeak.v b/theories/FSets/FSetWeak.v
deleted file mode 100644
index c88a7869..00000000
--- a/theories/FSets/FSetWeak.v
+++ /dev/null
@@ -1,16 +0,0 @@
-(***********************************************************************)
-(* 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: FSetWeak.v 9278 2006-10-25 13:43:17Z letouzey $ *)
-
-Require Export DecidableType.
-Require Export DecidableTypeEx.
-Require Export FSetWeakInterface.
-Require Export FSetWeakFacts.
-Require Export FSetWeakProperties.
-Require Export FSetWeakList.
diff --git a/theories/FSets/FSetWeakFacts.v b/theories/FSets/FSetWeakFacts.v
deleted file mode 100644
index 61797a95..00000000
--- a/theories/FSets/FSetWeakFacts.v
+++ /dev/null
@@ -1,421 +0,0 @@
-(***********************************************************************)
-(* 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: FSetWeakFacts.v 8882 2006-05-31 21:55:30Z letouzey $ *)
-
-(** * Finite sets library *)
-
-(** This functor derives additional facts from [FSetInterface.S]. These
- facts are mainly the specifications of [FSetInterface.S] written using
- different styles: equivalence and boolean equalities.
- Moreover, we prove that [E.Eq] and [Equal] are setoid equalities.
-*)
-
-Require Export FSetWeakInterface.
-Set Implicit Arguments.
-Unset Strict Implicit.
-
-Module Facts (M: S).
-Import M.E.
-Import M.
-Import Logic. (* to unmask [eq] *)
-
-(** * Specifications written using equivalences *)
-
-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.
-split; apply In_1; auto.
-Qed.
-
-Lemma mem_iff : In x s <-> mem x s = true.
-Proof.
-split; [apply mem_1|apply mem_2].
-Qed.
-
-Lemma not_mem_iff : ~In x s <-> mem x s = false.
-Proof.
-rewrite mem_iff; destruct (mem x s); intuition.
-Qed.
-
-Lemma equal_iff : s[=]s' <-> equal s s' = true.
-Proof.
-split; [apply equal_1|apply equal_2].
-Qed.
-
-Lemma subset_iff : s[<=]s' <-> subset s s' = true.
-Proof.
-split; [apply subset_1|apply subset_2].
-Qed.
-
-Lemma empty_iff : In x empty <-> False.
-Proof.
-intuition; apply (empty_1 H).
-Qed.
-
-Lemma is_empty_iff : Empty s <-> is_empty s = true.
-Proof.
-split; [apply is_empty_1|apply is_empty_2].
-Qed.
-
-Lemma singleton_iff : In y (singleton x) <-> E.eq x y.
-Proof.
-split; [apply singleton_1|apply singleton_2].
-Qed.
-
-Lemma add_iff : In y (add x s) <-> E.eq x y \/ In y s.
-Proof.
-split; [ | destruct 1; [apply add_1|apply add_2]]; auto.
-destruct (eq_dec x y) as [E|E]; auto.
-intro H; right; exact (add_3 E H).
-Qed.
-
-Lemma add_neq_iff : ~ E.eq x y -> (In y (add x s) <-> In y s).
-Proof.
-split; [apply add_3|apply add_2]; auto.
-Qed.
-
-Lemma remove_iff : In y (remove x s) <-> In y s /\ ~E.eq x y.
-Proof.
-split; [split; [apply remove_3 with x |] | destruct 1; apply remove_2]; auto.
-intro.
-apply (remove_1 H0 H).
-Qed.
-
-Lemma remove_neq_iff : ~ E.eq x y -> (In y (remove x s) <-> In y s).
-Proof.
-split; [apply remove_3|apply remove_2]; auto.
-Qed.
-
-Lemma union_iff : In x (union s s') <-> In x s \/ In x s'.
-Proof.
-split; [apply union_1 | destruct 1; [apply union_2|apply union_3]]; auto.
-Qed.
-
-Lemma inter_iff : In x (inter s s') <-> In x s /\ In x s'.
-Proof.
-split; [split; [apply inter_1 with s' | apply inter_2 with s] | destruct 1; apply inter_3]; auto.
-Qed.
-
-Lemma diff_iff : In x (diff s s') <-> In x s /\ ~ In x s'.
-Proof.
-split; [split; [apply diff_1 with s' | apply diff_2 with s] | destruct 1; apply diff_3]; auto.
-Qed.
-
-Variable f : elt->bool.
-
-Lemma filter_iff : compat_bool E.eq f -> (In x (filter f s) <-> In x s /\ f x = true).
-Proof.
-split; [split; [apply filter_1 with f | apply filter_2 with s] | destruct 1; apply filter_3]; auto.
-Qed.
-
-Lemma for_all_iff : compat_bool E.eq f ->
- (For_all (fun x => f x = true) s <-> for_all f s = true).
-Proof.
-split; [apply for_all_1 | apply for_all_2]; auto.
-Qed.
-
-Lemma exists_iff : compat_bool E.eq f ->
- (Exists (fun x => f x = true) s <-> exists_ f s = true).
-Proof.
-split; [apply exists_1 | apply exists_2]; auto.
-Qed.
-
-Lemma elements_iff : In x s <-> InA E.eq x (elements s).
-Proof.
-split; [apply elements_1 | apply elements_2].
-Qed.
-
-End IffSpec.
-
-(** 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 *)
-
-Definition eqb x y := if eq_dec x y then true else false.
-
-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 : compat_bool E.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 : compat_bool E.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; eauto.
-Qed.
-
-Lemma exists_b : compat_bool E.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; auto.
-split; auto.
-rewrite H2; rewrite InA_alt; eauto.
-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.
-
-(** * [E.eq] and [Equal] are setoid equalities *)
-
-Definition E_ST : Setoid_Theory elt E.eq.
-Proof.
-constructor; [apply E.eq_refl|apply E.eq_sym|apply E.eq_trans].
-Qed.
-
-Add Setoid elt E.eq E_ST as EltSetoid.
-
-Definition Equal_ST : Setoid_Theory t Equal.
-Proof.
-constructor; unfold Equal; firstorder.
-Qed.
-
-Add Setoid t Equal Equal_ST as EqualSetoid.
-
-Add Morphism In with signature E.eq ==> Equal ==> iff as In_m.
-Proof.
-unfold Equal; intros x y H s s' H0.
-rewrite (In_eq_iff s H); auto.
-Qed.
-
-Add Morphism is_empty : is_empty_m.
-Proof.
-unfold Equal; intros s s' H.
-generalize (is_empty_iff s)(is_empty_iff s').
-destruct (is_empty s); destruct (is_empty s');
- unfold Empty; auto; intros.
-symmetry.
-rewrite <- H1; intros a Ha.
-rewrite <- (H a) in Ha.
-destruct H0 as (_,H0).
-exact (H0 (refl_equal true) _ Ha).
-rewrite <- H0; intros a Ha.
-rewrite (H a) in Ha.
-destruct H1 as (_,H1).
-exact (H1 (refl_equal true) _ Ha).
-Qed.
-
-Add Morphism Empty with signature Equal ==> iff as Empty_m.
-Proof.
-intros; do 2 rewrite is_empty_iff; rewrite H; intuition.
-Qed.
-
-Add Morphism mem : mem_m.
-Proof.
-unfold Equal; intros x y H s s' H0.
-generalize (H0 x); clear H0; rewrite (In_eq_iff s' H).
-generalize (mem_iff s x)(mem_iff s' y).
-destruct (mem x s); destruct (mem y s'); intuition.
-Qed.
-
-Add Morphism singleton : singleton_m.
-Proof.
-unfold Equal; intros x y H a.
-do 2 rewrite singleton_iff; split.
-intros; apply E.eq_trans with x; auto.
-intros; apply E.eq_trans with y; auto.
-Qed.
-
-Add Morphism add : add_m.
-Proof.
-unfold Equal; intros x y H s s' H0 a.
-do 2 rewrite add_iff; rewrite H; rewrite H0; intuition.
-Qed.
-
-Add Morphism remove : remove_m.
-Proof.
-unfold Equal; intros x y H s s' H0 a.
-do 2 rewrite remove_iff; rewrite H; rewrite H0; intuition.
-Qed.
-
-Add Morphism union : union_m.
-Proof.
-unfold Equal; intros s s' H s'' s''' H0 a.
-do 2 rewrite union_iff; rewrite H; rewrite H0; intuition.
-Qed.
-
-Add Morphism inter : inter_m.
-Proof.
-unfold Equal; intros s s' H s'' s''' H0 a.
-do 2 rewrite inter_iff; rewrite H; rewrite H0; intuition.
-Qed.
-
-Add Morphism diff : diff_m.
-Proof.
-unfold Equal; intros s s' H s'' s''' H0 a.
-do 2 rewrite diff_iff; rewrite H; rewrite H0; intuition.
-Qed.
-
-Add Morphism Subset with signature Equal ==> Equal ==> iff as Subset_m.
-Proof.
-unfold Equal, Subset; firstorder.
-Qed.
-
-Add Morphism subset : subset_m.
-Proof.
-intros s s' H s'' s''' H0.
-generalize (subset_iff s s'') (subset_iff s' s''').
-destruct (subset s s''); destruct (subset s' s'''); auto; intros.
-rewrite H in H1; rewrite H0 in H1; intuition.
-rewrite H in H1; rewrite H0 in H1; intuition.
-Qed.
-
-Add Morphism equal : equal_m.
-Proof.
-intros s s' H s'' s''' H0.
-generalize (equal_iff s s'') (equal_iff s' s''').
-destruct (equal s s''); destruct (equal s' s'''); auto; intros.
-rewrite H in H1; rewrite H0 in H1; intuition.
-rewrite H in H1; rewrite H0 in H1; intuition.
-Qed.
-
-(* [fold], [filter], [for_all], [exists_] and [partition] cannot be proved morphism
- without additional hypothesis on [f]. For instance: *)
-
-Lemma filter_equal : forall f, compat_bool E.eq f ->
- forall s s', s[=]s' -> filter f s [=] filter f s'.
-Proof.
-unfold Equal; intros; repeat rewrite filter_iff; auto; rewrite H0; tauto.
-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 Facts.
diff --git a/theories/FSets/FSetWeakInterface.v b/theories/FSets/FSetWeakInterface.v
deleted file mode 100644
index a281ce22..00000000
--- a/theories/FSets/FSetWeakInterface.v
+++ /dev/null
@@ -1,251 +0,0 @@
-(***********************************************************************)
-(* 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: FSetWeakInterface.v 8820 2006-05-15 11:44:05Z letouzey $ *)
-
-(** * Finite sets library *)
-
-(** Set interfaces for types with only a decidable equality, but no ordering *)
-
-Require Export Bool.
-Require Export DecidableType.
-Set Implicit Arguments.
-Unset Strict Implicit.
-
-(** Compatibility of a boolean function with respect to an equality. *)
-Definition compat_bool (A:Set)(eqA: A->A->Prop)(f: A-> bool) :=
- forall x y : A, eqA x y -> f x = f y.
-
-(** Compatibility of a predicate with respect to an equality. *)
-Definition compat_P (A:Set)(eqA: A->A->Prop)(P : A -> Prop) :=
- forall x y : A, eqA x y -> P x -> P y.
-
-Hint Unfold compat_bool compat_P.
-
-(** * Non-dependent signature
-
- Signature [S] presents sets as purely informative programs
- together with axioms *)
-
-Module Type S.
-
- Declare Module E : DecidableType.
- Definition elt := E.t.
-
- Parameter t : Set. (** the abstract type of sets *)
-
- (** Logical predicates *)
- Parameter In : elt -> t -> Prop.
- 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).
-
- 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]. *)
-
- (** Coq comment: [iter] is useless in a purely functional world *)
- (** iter: (elt -> unit) -> set -> unit. i*)
- (** [iter f s] applies [f] in turn to all elements of [s].
- The order in which the elements of [s] are presented to [f]
- is unspecified. *)
-
- Parameter fold : forall A : Set, (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. *)
- (** Coq comment: nat instead of int ... *)
-
- 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 raise [Not_found] if
- the set is empty. Which element is chosen is unspecified.
- Equal sets could return different elements. *)
- (** Coq comment: [Not_found] is represented by the option type *)
-
- Section Spec.
-
- Variable s s' : t.
- Variable x y : elt.
-
- (** Specification of [In] *)
- Parameter In_1 : E.eq x y -> In x s -> In y s.
-
- (** Specification of [mem] *)
- Parameter mem_1 : In x s -> mem x s = true.
- Parameter mem_2 : mem x s = true -> In x s.
-
- (** Specification of [equal] *)
- Parameter equal_1 : Equal s s' -> equal s s' = true.
- Parameter equal_2 : equal s s' = true -> Equal s s'.
-
- (** Specification of [subset] *)
- Parameter subset_1 : Subset s s' -> subset s s' = true.
- Parameter subset_2 : subset s s' = true -> Subset s s'.
-
- (** Specification of [empty] *)
- Parameter empty_1 : Empty empty.
-
- (** Specification of [is_empty] *)
- Parameter is_empty_1 : Empty s -> is_empty s = true.
- Parameter is_empty_2 : is_empty s = true -> Empty s.
-
- (** Specification of [add] *)
- Parameter add_1 : E.eq x y -> In y (add x s).
- Parameter add_2 : In y s -> In y (add x s).
- Parameter add_3 : ~ E.eq x y -> In y (add x s) -> In y s.
-
- (** Specification of [remove] *)
- Parameter remove_1 : E.eq x y -> ~ In y (remove x s).
- Parameter remove_2 : ~ E.eq x y -> In y s -> In y (remove x s).
- Parameter remove_3 : In y (remove x s) -> In y s.
-
- (** Specification of [singleton] *)
- Parameter singleton_1 : In y (singleton x) -> E.eq x y.
- Parameter singleton_2 : E.eq x y -> In y (singleton x).
-
- (** Specification of [union] *)
- Parameter union_1 : In x (union s s') -> In x s \/ In x s'.
- Parameter union_2 : In x s -> In x (union s s').
- Parameter union_3 : In x s' -> In x (union s s').
-
- (** Specification of [inter] *)
- Parameter inter_1 : In x (inter s s') -> In x s.
- Parameter inter_2 : In x (inter s s') -> In x s'.
- Parameter inter_3 : In x s -> In x s' -> In x (inter s s').
-
- (** Specification of [diff] *)
- Parameter diff_1 : In x (diff s s') -> In x s.
- Parameter diff_2 : In x (diff s s') -> ~ In x s'.
- Parameter diff_3 : In x s -> ~ In x s' -> In x (diff s s').
-
- (** Specification of [fold] *)
- Parameter fold_1 : forall (A : Set) (i : A) (f : elt -> A -> A),
- fold f s i = fold_left (fun a e => f e a) (elements s) i.
-
- (** Specification of [cardinal] *)
- Parameter cardinal_1 : cardinal s = length (elements s).
-
- Section Filter.
-
- Variable f : elt -> bool.
-
- (** Specification of [filter] *)
- Parameter filter_1 : compat_bool E.eq f -> In x (filter f s) -> In x s.
- Parameter filter_2 : compat_bool E.eq f -> In x (filter f s) -> f x = true.
- Parameter filter_3 :
- compat_bool E.eq f -> In x s -> f x = true -> In x (filter f s).
-
- (** Specification of [for_all] *)
- Parameter for_all_1 :
- compat_bool E.eq f ->
- For_all (fun x => f x = true) s -> for_all f s = true.
- Parameter for_all_2 :
- compat_bool E.eq f ->
- for_all f s = true -> For_all (fun x => f x = true) s.
-
- (** Specification of [exists] *)
- Parameter exists_1 :
- compat_bool E.eq f ->
- Exists (fun x => f x = true) s -> exists_ f s = true.
- Parameter exists_2 :
- compat_bool E.eq f ->
- exists_ f s = true -> Exists (fun x => f x = true) s.
-
- (** Specification of [partition] *)
- Parameter partition_1 :
- compat_bool E.eq f -> Equal (fst (partition f s)) (filter f s).
- Parameter partition_2 :
- compat_bool E.eq f ->
- Equal (snd (partition f s)) (filter (fun x => negb (f x)) s).
-
- End Filter.
-
- (** Specification of [elements] *)
- Parameter elements_1 : In x s -> InA E.eq x (elements s).
- Parameter elements_2 : InA E.eq x (elements s) -> In x s.
- Parameter elements_3 : NoDupA E.eq (elements s).
-
- (** Specification of [choose] *)
- Parameter choose_1 : choose s = Some x -> In x s.
- Parameter choose_2 : choose s = None -> Empty s.
-
- End Spec.
-
- Hint Immediate In_1.
-
- Hint Resolve mem_1 mem_2 equal_1 equal_2 subset_1 subset_2 empty_1
- is_empty_1 is_empty_2 choose_1 choose_2 add_1 add_2 add_3 remove_1
- remove_2 remove_3 singleton_1 singleton_2 union_1 union_2 union_3 inter_1
- inter_2 inter_3 diff_1 diff_2 diff_3 filter_1 filter_2 filter_3 for_all_1
- for_all_2 exists_1 exists_2 partition_1 partition_2 elements_1 elements_2
- elements_3.
-
-End S.
diff --git a/theories/FSets/FSetWeakList.v b/theories/FSets/FSetWeakList.v
index 97080b7a..71a0d584 100644
--- a/theories/FSets/FSetWeakList.v
+++ b/theories/FSets/FSetWeakList.v
@@ -6,14 +6,14 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: FSetWeakList.v 8834 2006-05-20 00:41:35Z letouzey $ *)
+(* $Id: FSetWeakList.v 10631 2008-03-06 18:17:24Z msozeau $ *)
(** * Finite sets library *)
(** This file proposes an implementation of the non-dependant
interface [FSetWeakInterface.S] using lists without redundancy. *)
-Require Import FSetWeakInterface.
+Require Import FSetInterface.
Set Implicit Arguments.
Unset Strict Implicit.
@@ -24,8 +24,6 @@ Unset Strict Implicit.
And the functions returning sets are proved to preserve this invariant. *)
Module Raw (X: DecidableType).
-
- Module E := X.
Definition elt := X.t.
Definition t := list elt.
@@ -59,7 +57,7 @@ Module Raw (X: DecidableType).
if X.eq_dec x y then l else y :: remove x l
end.
- Fixpoint fold (B : Set) (f : elt -> B -> B) (s : t) {struct s} :
+ Fixpoint fold (B : Type) (f : elt -> B -> B) (s : t) {struct s} :
B -> B := fun i => match s with
| nil => i
| x :: l => fold f l (f x i)
@@ -127,7 +125,7 @@ Module Raw (X: DecidableType).
Lemma In_eq :
forall (s : t) (x y : elt), X.eq x y -> In x s -> In y s.
Proof.
- intros s x y; do 2 setoid_rewrite InA_alt; firstorder eauto.
+ intros s x y; setoid_rewrite InA_alt; firstorder eauto.
Qed.
Hint Immediate In_eq.
@@ -287,13 +285,13 @@ Module Raw (X: DecidableType).
unfold elements; auto.
Qed.
- Lemma elements_3 : forall (s : t) (Hs : NoDup s), NoDup (elements s).
+ Lemma elements_3w : forall (s : t) (Hs : NoDup s), NoDup (elements s).
Proof.
unfold elements; auto.
Qed.
Lemma fold_1 :
- forall (s : t) (Hs : NoDup s) (A : Set) (i : A) (f : elt -> A -> A),
+ forall (s : t) (Hs : NoDup 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.
induction s; simpl; auto; intros.
@@ -732,22 +730,68 @@ Module Raw (X: DecidableType).
generalize (Hrec H0 f).
case (f x); case (partition f l); simpl; auto.
Qed.
-
+
Definition eq : t -> t -> Prop := Equal.
- Lemma eq_refl : forall s : t, eq s s.
- Proof.
- unfold eq, Equal; intuition.
- Qed.
+ Lemma eq_refl : forall s, eq s s.
+ Proof. firstorder. Qed.
- Lemma eq_sym : forall s s' : t, eq s s' -> eq s' s.
- Proof.
- unfold eq, Equal; firstorder.
- Qed.
+ Lemma eq_sym : forall s s', eq s s' -> eq s' s.
+ Proof. firstorder. Qed.
- Lemma eq_trans : forall s s' s'' : t, eq s s' -> eq s' s'' -> eq s s''.
- Proof.
- unfold eq, Equal; firstorder.
+ Lemma eq_trans :
+ forall s s' s'', eq s s' -> eq s' s'' -> eq s s''.
+ Proof. firstorder. Qed.
+
+ Definition eq_dec : forall (s s':t)(Hs:NoDup s)(Hs':NoDup s'),
+ { eq s s' }+{ ~eq s s' }.
+ Proof.
+ unfold eq.
+ induction s; intros s'.
+ (* nil *)
+ destruct s'; [left|right].
+ firstorder.
+ unfold not, Equal.
+ intros H; generalize (H e); clear H.
+ rewrite InA_nil, InA_cons; intuition.
+ (* cons *)
+ intros.
+ case_eq (mem a s'); intros H;
+ [ destruct (IHs (remove a s')) as [H'|H'];
+ [ | | left|right]|right];
+ clear IHs.
+ inversion_clear Hs; auto.
+ apply remove_unique; auto.
+ (* In a s' /\ s [=] remove a s' *)
+ generalize (mem_2 H); clear H; intro H.
+ unfold Equal in *; intros b.
+ rewrite InA_cons; split.
+ destruct 1.
+ apply In_eq with a; auto.
+ rewrite H' in H0.
+ apply remove_3 with a; auto.
+ destruct (X.eq_dec b a); [left|right]; auto.
+ rewrite H'.
+ apply remove_2; auto.
+ (* In a s' /\ ~ s [=] remove a s' *)
+ generalize (mem_2 H); clear H; intro H.
+ contradict H'.
+ unfold Equal in *; intros b.
+ split; intros.
+ apply remove_2; auto.
+ inversion_clear Hs.
+ contradict H1; apply In_eq with b; auto.
+ rewrite <- H'; rewrite InA_cons; auto.
+ assert (In b s') by (apply remove_3 with a; auto).
+ rewrite <- H', InA_cons in H1; destruct H1; auto.
+ elim (remove_1 Hs' (X.eq_sym H1) H0).
+ (* ~ In a s' *)
+ assert (~In a s').
+ red; intro H'; rewrite (mem_1 H') in H; discriminate.
+ contradict H0.
+ unfold Equal in *.
+ rewrite <- H0.
+ rewrite InA_cons; auto.
Qed.
End ForNotations.
@@ -758,12 +802,12 @@ End Raw.
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) <: S with Module E := X.
+Module Make (X: DecidableType) <: WS with Module E := X.
Module Raw := Raw X.
Module E := X.
- Record slist : Set := {this :> Raw.t; unique : NoDupA E.eq this}.
+ Record slist := {this :> Raw.t; unique : NoDupA E.eq this}.
Definition t := slist.
Definition elt := E.t.
@@ -791,7 +835,7 @@ Module Make (X: DecidableType) <: S with Module E := X.
Definition is_empty (s : t) : bool := Raw.is_empty s.
Definition elements (s : t) : list elt := Raw.elements s.
Definition choose (s:t) : option elt := Raw.choose s.
- Definition fold (B : Set) (f : elt -> B -> B) (s : t) : B -> B := Raw.fold (B:=B) f s.
+ Definition fold (B : Type) (f : elt -> B -> B) (s : t) : B -> B := Raw.fold (B:=B) f s.
Definition cardinal (s : t) : nat := Raw.cardinal s.
Definition filter (f : elt -> bool) (s : t) : t :=
Build_slist (Raw.filter_unique (unique s) f).
@@ -872,7 +916,7 @@ Module Make (X: DecidableType) <: S with Module E := X.
Lemma diff_3 : In x s -> ~ In x s' -> In x (diff s s').
Proof. exact (fun H => Raw.diff_3 s.(unique) s'.(unique) H). Qed.
- Lemma fold_1 : forall (A : Set) (i : A) (f : elt -> A -> A),
+ Lemma fold_1 : 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 (Raw.fold_1 s.(unique)). Qed.
@@ -923,8 +967,8 @@ Module Make (X: DecidableType) <: S with Module E := X.
Proof. exact (fun H => Raw.elements_1 H). Qed.
Lemma elements_2 : InA E.eq x (elements s) -> In x s.
Proof. exact (fun H => Raw.elements_2 H). Qed.
- Lemma elements_3 : NoDupA E.eq (elements s).
- Proof. exact (Raw.elements_3 s.(unique)). Qed.
+ Lemma elements_3w : NoDupA E.eq (elements s).
+ Proof. exact (Raw.elements_3w s.(unique)). Qed.
Lemma choose_1 : choose s = Some x -> In x s.
Proof. exact (fun H => Raw.choose_1 H). Qed.
@@ -933,4 +977,22 @@ Module Make (X: DecidableType) <: S with Module E := X.
End Spec.
+ Definition eq : t -> t -> Prop := Equal.
+
+ Lemma eq_refl : forall s, eq s s.
+ Proof. firstorder. Qed.
+
+ Lemma eq_sym : forall s s', eq s s' -> eq s' s.
+ Proof. firstorder. Qed.
+
+ Lemma eq_trans :
+ forall s s' s'', eq s s' -> eq s' s'' -> eq s s''.
+ Proof. firstorder. Qed.
+
+ Definition eq_dec : forall (s s':t),
+ { eq s s' }+{ ~eq s s' }.
+ Proof.
+ intros s s'; exact (Raw.eq_dec s.(unique) s'.(unique)).
+ Qed.
+
End Make.
diff --git a/theories/FSets/FSetWeakProperties.v b/theories/FSets/FSetWeakProperties.v
deleted file mode 100644
index a0054d36..00000000
--- a/theories/FSets/FSetWeakProperties.v
+++ /dev/null
@@ -1,896 +0,0 @@
-(***********************************************************************)
-(* 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: FSetWeakProperties.v 8853 2006-05-23 18:17:38Z herbelin $ *)
-
-(** * Finite sets library *)
-
-(** NB: this file is a clone of [FSetProperties] for weak sets
- and should remain so until we find a way to share the two. *)
-
-(** This functor derives additional properties from [FSetWeakInterface.S].
- Contrary to the functor in [FSetWeakEqProperties] 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 FSetWeakInterface.
-Require Import FSetWeakFacts.
-Set Implicit Arguments.
-Unset Strict Implicit.
-
-Hint Unfold transpose compat_op.
-Hint Extern 1 (Setoid_Theory _ _) => constructor; congruence.
-
-Module Properties (M: S).
- Import M.E.
- Import M.
- Import Logic. (* to unmask [eq] *)
- Import Peano. (* to unmask [lt] *)
-
- (** Results about lists without duplicates *)
-
- Module FM := Facts M.
- Import FM.
-
- Definition Add (x : elt) (s s' : t) :=
- forall y : elt, In y s' <-> E.eq x y \/ In y s.
-
- 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.
-
- Section BasicProperties.
-
- (** properties of [Equal] *)
-
- Lemma equal_refl : forall s, s[=]s.
- Proof.
- unfold Equal; intuition.
- Qed.
-
- Lemma equal_sym : forall s s', s[=]s' -> s'[=]s.
- Proof.
- unfold Equal; intros.
- rewrite H; intuition.
- Qed.
-
- Lemma equal_trans : forall s1 s2 s3, s1[=]s2 -> s2[=]s3 -> s1[=]s3.
- Proof.
- unfold Equal; intros.
- rewrite H; exact (H0 a).
- Qed.
-
- Variable s s' s'' s1 s2 s3 : t.
- Variable x x' : elt.
-
- (** properties of [Subset] *)
-
- Lemma subset_refl : s[<=]s.
- Proof.
- unfold Subset; intuition.
- Qed.
-
- Lemma subset_antisym : s[<=]s' -> s'[<=]s -> s[=]s'.
- Proof.
- unfold Subset, Equal; intuition.
- Qed.
-
- Lemma subset_trans : s1[<=]s2 -> s2[<=]s3 -> s1[<=]s3.
- Proof.
- unfold Subset; intuition.
- Qed.
-
- Lemma subset_equal : s[=]s' -> s[<=]s'.
- Proof.
- unfold Subset, Equal; firstorder.
- Qed.
-
- Lemma subset_empty : empty[<=]s.
- Proof.
- unfold Subset; intros a; set_iff; intuition.
- Qed.
-
- Lemma subset_remove_3 : s1[<=]s2 -> remove x s1 [<=] s2.
- Proof.
- unfold Subset; intros H a; set_iff; intuition.
- Qed.
-
- Lemma subset_diff : s1[<=]s3 -> diff s1 s2 [<=] s3.
- Proof.
- unfold Subset; intros H a; set_iff; intuition.
- Qed.
-
- Lemma subset_add_3 : In x s2 -> s1[<=]s2 -> add x s1 [<=] s2.
- Proof.
- unfold Subset; intros H H0 a; set_iff; intuition.
- rewrite <- H2; auto.
- Qed.
-
- Lemma subset_add_2 : s1[<=]s2 -> s1[<=] add x s2.
- Proof.
- unfold Subset; intuition.
- Qed.
-
- Lemma in_subset : In x s1 -> s1[<=]s2 -> In x s2.
- Proof.
- unfold Subset; intuition.
- Qed.
-
- Lemma double_inclusion : s1[=]s2 <-> s1[<=]s2 /\ s2[<=]s1.
- Proof.
- unfold Subset, Equal; split; intros; intuition; generalize (H a); intuition.
- Qed.
-
- (** properties of [empty] *)
-
- Lemma empty_is_empty_1 : Empty s -> s[=]empty.
- Proof.
- unfold Empty, Equal; intros; generalize (H a); set_iff; tauto.
- Qed.
-
- Lemma empty_is_empty_2 : s[=]empty -> Empty s.
- Proof.
- unfold Empty, Equal; intros; generalize (H a); set_iff; tauto.
- Qed.
-
- (** properties of [add] *)
-
- Lemma add_equal : In x s -> add x s [=] s.
- Proof.
- unfold Equal; intros; set_iff; intuition.
- rewrite <- H1; auto.
- Qed.
-
- Lemma add_add : add x (add x' s) [=] add x' (add x s).
- Proof.
- unfold Equal; intros; set_iff; tauto.
- Qed.
-
- (** properties of [remove] *)
-
- Lemma remove_equal : ~ In x s -> remove x s [=] s.
- Proof.
- unfold Equal; intros; set_iff; intuition.
- rewrite H1 in H; auto.
- Qed.
-
- Lemma Equal_remove : s[=]s' -> remove x s [=] remove x s'.
- Proof.
- intros; rewrite H; apply equal_refl.
- Qed.
-
- (** properties of [add] and [remove] *)
-
- Lemma add_remove : In x s -> add x (remove x s) [=] s.
- Proof.
- unfold Equal; intros; set_iff; elim (eq_dec x a); intuition.
- rewrite <- H1; auto.
- Qed.
-
- Lemma remove_add : ~In x s -> remove x (add x s) [=] s.
- Proof.
- unfold Equal; intros; set_iff; elim (eq_dec x a); intuition.
- rewrite H1 in H; auto.
- Qed.
-
- (** properties of [singleton] *)
-
- Lemma singleton_equal_add : singleton x [=] add x empty.
- Proof.
- unfold Equal; intros; set_iff; intuition.
- Qed.
-
- (** properties of [union] *)
-
- Lemma union_sym : union s s' [=] union s' s.
- Proof.
- unfold Equal; intros; set_iff; tauto.
- Qed.
-
- Lemma union_subset_equal : s[<=]s' -> union s s' [=] s'.
- Proof.
- unfold Subset, Equal; intros; set_iff; intuition.
- Qed.
-
- Lemma union_equal_1 : s[=]s' -> union s s'' [=] union s' s''.
- Proof.
- intros; rewrite H; apply equal_refl.
- Qed.
-
- Lemma union_equal_2 : s'[=]s'' -> union s s' [=] union s s''.
- Proof.
- intros; rewrite H; apply equal_refl.
- Qed.
-
- Lemma union_assoc : union (union s s') s'' [=] union s (union s' s'').
- Proof.
- unfold Equal; intros; set_iff; tauto.
- Qed.
-
- Lemma add_union_singleton : add x s [=] union (singleton x) s.
- Proof.
- unfold Equal; intros; set_iff; tauto.
- Qed.
-
- Lemma union_add : union (add x s) s' [=] add x (union s s').
- Proof.
- unfold Equal; intros; set_iff; tauto.
- Qed.
-
- Lemma union_subset_1 : s [<=] union s s'.
- Proof.
- unfold Subset; intuition.
- Qed.
-
- Lemma union_subset_2 : s' [<=] union s s'.
- Proof.
- unfold Subset; intuition.
- Qed.
-
- Lemma union_subset_3 : s[<=]s'' -> s'[<=]s'' -> union s s' [<=] s''.
- Proof.
- unfold Subset; intros H H0 a; set_iff; intuition.
- Qed.
-
- Lemma union_subset_4 : s[<=]s' -> union s s'' [<=] union s' s''.
- Proof.
- unfold Subset; intros H a; set_iff; intuition.
- Qed.
-
- Lemma union_subset_5 : s[<=]s' -> union s'' s [<=] union s'' s'.
- Proof.
- unfold Subset; intros H a; set_iff; intuition.
- Qed.
-
- Lemma empty_union_1 : Empty s -> union s s' [=] s'.
- Proof.
- unfold Equal, Empty; intros; set_iff; firstorder.
- Qed.
-
- Lemma empty_union_2 : Empty s -> union s' s [=] s'.
- Proof.
- unfold Equal, Empty; intros; set_iff; firstorder.
- Qed.
-
- Lemma not_in_union : ~In x s -> ~In x s' -> ~In x (union s s').
- Proof.
- intros; set_iff; intuition.
- Qed.
-
- (** properties of [inter] *)
-
- Lemma inter_sym : inter s s' [=] inter s' s.
- Proof.
- unfold Equal; intros; set_iff; tauto.
- Qed.
-
- Lemma inter_subset_equal : s[<=]s' -> inter s s' [=] s.
- Proof.
- unfold Equal; intros; set_iff; intuition.
- Qed.
-
- Lemma inter_equal_1 : s[=]s' -> inter s s'' [=] inter s' s''.
- Proof.
- intros; rewrite H; apply equal_refl.
- Qed.
-
- Lemma inter_equal_2 : s'[=]s'' -> inter s s' [=] inter s s''.
- Proof.
- intros; rewrite H; apply equal_refl.
- Qed.
-
- Lemma inter_assoc : inter (inter s s') s'' [=] inter s (inter s' s'').
- Proof.
- unfold Equal; intros; set_iff; tauto.
- Qed.
-
- Lemma union_inter_1 : inter (union s s') s'' [=] union (inter s s'') (inter s' s'').
- Proof.
- unfold Equal; intros; set_iff; tauto.
- Qed.
-
- Lemma union_inter_2 : union (inter s s') s'' [=] inter (union s s'') (union s' s'').
- Proof.
- unfold Equal; intros; set_iff; tauto.
- Qed.
-
- Lemma inter_add_1 : In x s' -> inter (add x s) s' [=] add x (inter s s').
- Proof.
- unfold Equal; intros; set_iff; intuition.
- rewrite <- H1; auto.
- Qed.
-
- Lemma inter_add_2 : ~ In x s' -> inter (add x s) s' [=] inter s s'.
- Proof.
- unfold Equal; intros; set_iff; intuition.
- destruct H; rewrite H0; auto.
- Qed.
-
- Lemma empty_inter_1 : Empty s -> Empty (inter s s').
- Proof.
- unfold Empty; intros; set_iff; firstorder.
- Qed.
-
- Lemma empty_inter_2 : Empty s' -> Empty (inter s s').
- Proof.
- unfold Empty; intros; set_iff; firstorder.
- Qed.
-
- Lemma inter_subset_1 : inter s s' [<=] s.
- Proof.
- unfold Subset; intro a; set_iff; tauto.
- Qed.
-
- Lemma inter_subset_2 : inter s s' [<=] s'.
- Proof.
- unfold Subset; intro a; set_iff; tauto.
- Qed.
-
- Lemma inter_subset_3 :
- s''[<=]s -> s''[<=]s' -> s''[<=] inter s s'.
- Proof.
- unfold Subset; intros H H' a; set_iff; intuition.
- Qed.
-
- (** properties of [diff] *)
-
- Lemma empty_diff_1 : Empty s -> Empty (diff s s').
- Proof.
- unfold Empty, Equal; intros; set_iff; firstorder.
- Qed.
-
- Lemma empty_diff_2 : Empty s -> diff s' s [=] s'.
- Proof.
- unfold Empty, Equal; intros; set_iff; firstorder.
- Qed.
-
- Lemma diff_subset : diff s s' [<=] s.
- Proof.
- unfold Subset; intros a; set_iff; tauto.
- Qed.
-
- Lemma diff_subset_equal : s[<=]s' -> diff s s' [=] empty.
- Proof.
- unfold Subset, Equal; intros; set_iff; intuition; absurd (In a empty); auto.
- Qed.
-
- Lemma remove_diff_singleton :
- remove x s [=] diff s (singleton x).
- Proof.
- unfold Equal; intros; set_iff; intuition.
- Qed.
-
- Lemma diff_inter_empty : inter (diff s s') (inter s s') [=] empty.
- Proof.
- unfold Equal; intros; set_iff; intuition; absurd (In a empty); auto.
- Qed.
-
- Lemma diff_inter_all : union (diff s s') (inter s s') [=] s.
- Proof.
- unfold Equal; intros; set_iff; intuition.
- elim (In_dec a s'); auto.
- Qed.
-
- (** properties of [Add] *)
-
- Lemma Add_add : Add x s (add x s).
- Proof.
- unfold Add; intros; set_iff; intuition.
- Qed.
-
- Lemma Add_remove : In x s -> Add x (remove x s) s.
- Proof.
- unfold Add; intros; set_iff; intuition.
- elim (eq_dec x y); auto.
- rewrite <- H1; auto.
- Qed.
-
- Lemma union_Add : Add x s s' -> Add x (union s s'') (union s' s'').
- Proof.
- unfold Add; intros; set_iff; rewrite H; tauto.
- Qed.
-
- Lemma inter_Add :
- In x s'' -> Add x s s' -> Add x (inter s s'') (inter s' s'').
- Proof.
- unfold Add; intros; set_iff; rewrite H0; intuition.
- rewrite <- H2; auto.
- Qed.
-
- Lemma union_Equal :
- In x s'' -> Add x s s' -> union s s'' [=] union s' s''.
- Proof.
- unfold Add, Equal; intros; set_iff; rewrite H0; intuition.
- rewrite <- H1; auto.
- Qed.
-
- Lemma inter_Add_2 :
- ~In x s'' -> Add x s s' -> inter s s'' [=] inter s' s''.
- Proof.
- unfold Add, Equal; intros; set_iff; rewrite H0; intuition.
- destruct H; rewrite H1; auto.
- Qed.
-
- End BasicProperties.
-
- Hint Immediate equal_sym: set.
- Hint Resolve equal_refl equal_trans : set.
-
- Hint Immediate add_remove remove_add union_sym inter_sym: set.
- Hint Resolve 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.
-
- (** * Alternative (weaker) specifications for [fold] *)
-
- Section Old_Spec_Now_Properties.
-
- Notation NoDup := (NoDupA E.eq).
-
- (** When [FSets] 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 : Set) (i : A) (f : elt -> A -> A),
- exists l : list elt,
- NoDup l /\
- (forall x : elt, In x s <-> InA E.eq x l) /\
- fold f s i = fold_right f i l.
- Proof.
- intros; exists (rev (elements s)); split.
- apply NoDupA_rev; auto.
- exact E.eq_trans.
- 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 : Set) (eqA : A -> A -> Prop)
- (st : Setoid_Theory A 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.
- refl_st.
- elim (H e).
- elim (H2 e); intuition.
- Qed.
-
- Lemma fold_2 :
- forall s s' x (A : Set) (eqA : A -> A -> Prop)
- (st : Setoid_Theory A eqA) (i : A) (f : elt -> A -> A),
- compat_op E.eq 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.
- exact eq_dec.
- rewrite <- Hl1; auto.
- intros; rewrite <- Hl1; rewrite <- Hl'1; auto.
- Qed.
-
- (** Similar specifications for [cardinal]. *)
-
- Lemma cardinal_fold : forall s, cardinal s = fold (fun _ => S) s 0.
- Proof.
- intros; rewrite cardinal_1; rewrite M.fold_1.
- symmetry; apply fold_left_length; auto.
- Qed.
-
- 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.
- 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.
- Qed.
-
- End Old_Spec_Now_Properties.
-
- (** * Induction principle over sets *)
-
- Lemma cardinal_inv_1 : forall s, cardinal s = 0 -> Empty s.
- Proof.
- intros s; rewrite M.cardinal_1; intros H a; red.
- rewrite elements_iff.
- destruct (elements s); simpl in *; discriminate || inversion 1.
- 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 M.cardinal_1 in H.
- generalize (elements_2 (s:=s)).
- destruct (elements s); try discriminate.
- exists e; auto.
- Qed.
-
- Lemma Equal_cardinal_aux :
- forall n s s', cardinal s = n -> s[=]s' -> cardinal s = cardinal s'.
- Proof.
- simple induction n; intros.
- rewrite H; symmetry .
- apply cardinal_1.
- rewrite <- H0; auto.
- destruct (cardinal_inv_2 H0) as (x,H2).
- revert H0.
- rewrite (cardinal_2 (s:=remove x s) (s':=s) (x:=x)); auto with set.
- rewrite (cardinal_2 (s:=remove x s') (s':=s') (x:=x)); auto with set.
- rewrite H1 in H2; auto with set.
- Qed.
-
- Lemma Equal_cardinal : forall s s', s[=]s' -> cardinal s = cardinal s'.
- Proof.
- intros; apply Equal_cardinal_aux with (cardinal s); auto.
- Qed.
-
- Add Morphism cardinal : cardinal_m.
- Proof.
- exact Equal_cardinal.
- Qed.
-
- Hint Resolve Add_add Add_remove Equal_remove cardinal_inv_1 Equal_cardinal.
-
- Lemma cardinal_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 n s, cardinal s = n -> P s.
- Proof.
- simple induction n; intros; auto.
- destruct (cardinal_inv_2 H) as (x,H0).
- apply X0 with (remove x s) x; auto.
- apply X1; auto.
- rewrite (cardinal_2 (x:=x)(s:=remove x s)(s':=s)) in H; auto.
- Qed.
-
- Lemma set_induction :
- forall P : t -> Type,
- (forall s : t, Empty s -> P s) ->
- (forall s s' : t, P s -> forall x : elt, ~In x s -> Add x s s' -> P s') ->
- forall s : t, P s.
- Proof.
- intros; apply cardinal_induction with (cardinal s); auto.
- Qed.
-
- (** Other properties of [fold]. *)
-
- Section Fold.
- Variables (A:Set)(eqA:A->A->Prop)(st:Setoid_Theory _ eqA).
- Variables (f:elt->A->A)(Comp:compat_op E.eq eqA f)(Ass:transpose eqA f).
-
- Section Fold_1.
- Variable i i':A.
-
- Lemma fold_empty : eqA (fold f empty i) i.
- Proof.
- apply fold_1; auto.
- Qed.
-
- Lemma fold_equal :
- forall s s', s[=]s' -> eqA (fold f s i) (fold f s' i).
- Proof.
- intros s; pattern s; apply set_induction; clear s; intros.
- trans_st i.
- apply fold_1; auto.
- sym_st; apply fold_1; auto.
- rewrite <- H0; auto.
- trans_st (f x (fold f s i)).
- apply fold_2 with (eqA := eqA); auto.
- sym_st; apply fold_2 with (eqA := eqA); auto.
- unfold Add in *; intros.
- rewrite <- H2; auto.
- Qed.
-
- Lemma fold_add : forall 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.
- Qed.
-
- Lemma add_fold : forall 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 s x, In x s ->
- eqA (f x (fold f (remove x s) i)) (fold f s i).
- Proof.
- intros.
- sym_st.
- apply fold_2 with (eqA:=eqA); auto.
- Qed.
-
- Lemma remove_fold_2: forall 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_commutes : forall s x,
- eqA (fold f s (f x i)) (f x (fold f s i)).
- Proof.
- intros; pattern s; apply set_induction; clear s; intros.
- trans_st (f x i).
- apply fold_1; auto.
- sym_st.
- apply Comp; auto.
- apply fold_1; auto.
- trans_st (f x0 (fold f s (f x i))).
- apply fold_2 with (eqA:=eqA); auto.
- trans_st (f x0 (f x (fold f s i))).
- trans_st (f x (f x0 (fold f s i))).
- apply Comp; auto.
- sym_st.
- apply fold_2 with (eqA:=eqA); auto.
- Qed.
-
- Lemma fold_init : forall s, eqA i i' ->
- eqA (fold f s i) (fold f s i').
- Proof.
- intros; pattern s; apply set_induction; clear s; intros.
- trans_st i.
- apply fold_1; auto.
- trans_st i'.
- sym_st; apply fold_1; auto.
- trans_st (f x (fold f s i)).
- apply fold_2 with (eqA:=eqA); auto.
- trans_st (f x (fold f s i')).
- sym_st; apply fold_2 with (eqA:=eqA); auto.
- Qed.
-
- End Fold_1.
- Section Fold_2.
- Variable i:A.
-
- Lemma fold_union_inter : forall 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.
- trans_st (fold f s' (fold f (inter s s') i)).
- apply fold_equal; auto with set.
- trans_st (fold f s' i).
- apply fold_init; auto.
- apply fold_1; auto with set.
- sym_st; apply fold_1; auto.
- rename s'0 into s''.
- destruct (In_dec x s').
- (* In x s' *)
- trans_st (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.
- trans_st (f x (fold f s (fold f s' i))).
- trans_st (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.
- trans_st (f x (fold f (union s s') (fold f (inter s s') i))).
- apply fold_commutes; auto.
- sym_st; apply fold_2 with (eqA:=eqA); auto.
- (* ~(In x s') *)
- trans_st (f x (fold f (union s s') (fold f (inter s'' s') i))).
- apply fold_2 with (eqA:=eqA); auto with set.
- trans_st (f x (fold f (union s s') (fold f (inter s s') i))).
- apply Comp;auto.
- apply fold_init;auto.
- apply fold_equal;auto.
- apply equal_sym; apply inter_Add_2 with x; auto with set.
- trans_st (f x (fold f s (fold f s' i))).
- sym_st; apply fold_2 with (eqA:=eqA); auto.
- Qed.
-
- End Fold_2.
- Section Fold_3.
- Variable i:A.
-
- Lemma fold_diff_inter : forall s s',
- eqA (fold f (diff s s') (fold f (inter s s') i)) (fold f s i).
- Proof.
- intros.
- trans_st (fold f (union (diff s s') (inter s s'))
- (fold f (inter (diff s s') (inter s s')) i)).
- sym_st; apply fold_union_inter; auto.
- trans_st (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 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.
- trans_st (fold f (union s s') (fold f (inter s s') i)).
- apply fold_init; auto.
- sym_st; apply fold_1; auto with set.
- unfold Empty; intro a; generalize (H a); set_iff; tauto.
- apply fold_union_inter; auto.
- Qed.
-
- End Fold_3.
- End Fold.
-
- Lemma fold_plus :
- forall s p, fold (fun _ => S) s p = fold (fun _ => S) s 0 + p.
- Proof.
- assert (st := gen_st nat).
- assert (fe : compat_op E.eq (@eq _) (fun _ => S)) by (unfold compat_op; auto).
- assert (fp : transpose (@eq _) (fun _:elt => S)) by (unfold transpose; auto).
- intros s p; pattern s; apply set_induction; clear s; intros.
- rewrite (fold_1 st p (fun _ => S) H).
- rewrite (fold_1 st 0 (fun _ => S) H); trivial.
- assert (forall p s', Add x s s' -> fold (fun _ => S) s' p = S (fold (fun _ => S) s p)).
- change S with ((fun _ => S) x).
- intros; apply fold_2; auto.
- rewrite H2; auto.
- rewrite (H2 0); auto.
- rewrite H.
- simpl; auto.
- Qed.
-
- (** properties of [cardinal] *)
-
- Lemma empty_cardinal : cardinal empty = 0.
- Proof.
- rewrite cardinal_fold; apply fold_1; auto.
- 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:=@eq nat); auto.
- 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.
- 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:=@eq nat); auto.
- 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:=@eq nat); auto.
- 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:=@eq nat); auto.
- 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 Properties.
diff --git a/theories/FSets/FSets.v b/theories/FSets/FSets.v
index b0402db6..a73c1da7 100644
--- a/theories/FSets/FSets.v
+++ b/theories/FSets/FSets.v
@@ -6,13 +6,19 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: FSets.v 8897 2006-06-05 21:04:10Z letouzey $ *)
+(* $Id: FSets.v 10699 2008-03-19 20:56:43Z letouzey $ *)
Require Export OrderedType.
Require Export OrderedTypeEx.
Require Export OrderedTypeAlt.
+Require Export DecidableType.
+Require Export DecidableTypeEx.
Require Export FSetInterface.
Require Export FSetBridge.
+Require Export FSetFacts.
+Require Export FSetDecide.
Require Export FSetProperties.
Require Export FSetEqProperties.
+Require Export FSetWeakList.
Require Export FSetList.
+Require Export FSetAVL. \ No newline at end of file
diff --git a/theories/FSets/OrderedType.v b/theories/FSets/OrderedType.v
index f966cd4d..c56a24cf 100644
--- a/theories/FSets/OrderedType.v
+++ b/theories/FSets/OrderedType.v
@@ -6,32 +6,25 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: OrderedType.v 8834 2006-05-20 00:41:35Z letouzey $ *)
+(* $Id: OrderedType.v 10616 2008-03-04 17:33:35Z letouzey $ *)
Require Export SetoidList.
Set Implicit Arguments.
Unset Strict Implicit.
-(* TODO concernant la tactique order:
- * propagate_lt n'est sans doute pas complet
- * un propagate_le
- * exploiter les hypotheses negatives restant a la fin
- * faire que ca marche meme quand une hypothese depend d'un eq ou lt.
-*)
-
(** * Ordered types *)
-Inductive Compare (X : Set) (lt eq : X -> X -> Prop) (x y : X) : Set :=
+Inductive Compare (X : Type) (lt eq : X -> X -> Prop) (x y : X) : Type :=
| LT : lt x y -> Compare lt eq x y
| EQ : eq x y -> Compare lt eq x y
| GT : lt y x -> Compare lt eq x y.
Module Type OrderedType.
- Parameter t : Set.
+ Parameter Inline t : Type.
- Parameter eq : t -> t -> Prop.
- Parameter lt : t -> t -> Prop.
+ Parameter Inline eq : t -> t -> Prop.
+ Parameter Inline lt : t -> t -> Prop.
Axiom eq_refl : forall x : t, eq x x.
Axiom eq_sym : forall x y : t, eq x y -> eq y x.
@@ -122,6 +115,13 @@ Module OrderedTypeFacts (O: OrderedType).
intuition.
Qed.
+(* TODO concernant la tactique order:
+ * propagate_lt n'est sans doute pas complet
+ * un propagate_le
+ * exploiter les hypotheses negatives restant a la fin
+ * faire que ca marche meme quand une hypothese depend d'un eq ou lt.
+*)
+
Ltac abstraction := match goal with
(* First, some obvious simplifications *)
| H : False |- _ => elim H
@@ -137,9 +137,9 @@ Ltac abstraction := match goal with
| H1: ~lt ?x ?y, H2: ~eq ?y ?x |- _ =>
generalize (le_neq H1 (neq_sym H2)); clear H1 H2; intro; abstraction
(* Then, we generalize all interesting facts *)
- | H : lt ?x ?y |- _ => revert H; abstraction
- | H : ~lt ?x ?y |- _ => revert H; abstraction
| H : ~eq ?x ?y |- _ => revert H; abstraction
+ | H : ~lt ?x ?y |- _ => revert H; abstraction
+ | H : lt ?x ?y |- _ => revert H; abstraction
| H : eq ?x ?y |- _ => revert H; abstraction
| _ => idtac
end.
@@ -192,7 +192,7 @@ Ltac do_lt x y LT := match goal with
| |- lt ?z x -> _ => let H := fresh "H" in
(intro H; generalize (lt_trans H LT); intro); do_lt x y LT
| |- lt _ _ -> _ => intro; do_lt x y LT
- (* Ge *)
+ (* GE *)
| |- ~lt y x -> _ => intros _; do_lt x y LT
| |- ~lt x ?z -> _ => let H := fresh "H" in
(intro H; generalize (le_lt_trans H LT); intro); do_lt x y LT
@@ -296,12 +296,12 @@ Ltac false_order := elimtype False; order.
Lemma eq_dec : forall x y : t, {eq x y} + {~ eq x y}.
Proof.
intros; elim (compare x y); [ right | left | right ]; auto.
- Qed.
+ Defined.
Lemma lt_dec : forall x y : t, {lt x y} + {~ lt x y}.
Proof.
intros; elim (compare x y); [ left | right | right ]; auto.
- Qed.
+ Defined.
Definition eqb x y : bool := if eq_dec x y then true else false.
@@ -361,7 +361,7 @@ Module KeyOrderedType(O:OrderedType).
Import MO.
Section Elt.
- Variable elt : Set.
+ Variable elt : Type.
Notation key:=t.
Definition eqk (p p':key*elt) := eq (fst p) (fst p').
diff --git a/theories/FSets/OrderedTypeAlt.v b/theories/FSets/OrderedTypeAlt.v
index 9bcfbfc7..516df0f0 100644
--- a/theories/FSets/OrderedTypeAlt.v
+++ b/theories/FSets/OrderedTypeAlt.v
@@ -11,19 +11,19 @@
* Institution: LRI, CNRS UMR 8623 - Université Paris Sud
* 91405 Orsay, France *)
-(* $Id: OrderedTypeAlt.v 8773 2006-04-29 14:31:32Z letouzey $ *)
+(* $Id: OrderedTypeAlt.v 10739 2008-04-01 14:45:20Z herbelin $ *)
Require Import OrderedType.
(** * An alternative (but equivalent) presentation for an Ordered Type inferface. *)
-(** NB: [comparison], defined in [theories/Init/datatypes.v] is [Eq|Lt|Gt]
-whereas [compare], defined in [theories/FSets/OrderedType.v] is [EQ _ | LT _ | GT _ ]
+(** NB: [comparison], defined in [Datatypes.v] is [Eq|Lt|Gt]
+whereas [compare], defined in [OrderedType.v] is [EQ _ | LT _ | GT _ ]
*)
Module Type OrderedTypeAlt.
- Parameter t : Set.
+ Parameter t : Type.
Parameter compare : t -> t -> comparison.
@@ -103,24 +103,16 @@ Module OrderedType_to_Alt (O:OrderedType) <: OrderedTypeAlt.
Lemma compare_sym :
forall x y, (y?=x) = CompOpp (x?=y).
Proof.
- intros x y.
- unfold compare.
- destruct (O.compare y x); elim_comp; simpl; auto.
+ intros x y; unfold compare.
+ destruct O.compare; elim_comp; simpl; auto.
Qed.
Lemma compare_trans :
forall c x y z, (x?=y) = c -> (y?=z) = c -> (x?=z) = c.
Proof.
intros c x y z.
- destruct c; unfold compare.
- destruct (O.compare x y); intros; try discriminate.
- destruct (O.compare y z); intros; try discriminate.
- elim_comp; auto.
- destruct (O.compare x y); intros; try discriminate.
- destruct (O.compare y z); intros; try discriminate.
- elim_comp; auto.
- destruct (O.compare x y); intros; try discriminate.
- destruct (O.compare y z); intros; try discriminate.
+ destruct c; unfold compare;
+ do 2 (destruct O.compare; intros; try discriminate);
elim_comp; auto.
Qed.
diff --git a/theories/FSets/OrderedTypeEx.v b/theories/FSets/OrderedTypeEx.v
index 28a5705d..03171396 100644
--- a/theories/FSets/OrderedTypeEx.v
+++ b/theories/FSets/OrderedTypeEx.v
@@ -11,7 +11,7 @@
* Institution: LRI, CNRS UMR 8623 - Université Paris Sud
* 91405 Orsay, France *)
-(* $Id: OrderedTypeEx.v 9940 2007-07-05 12:32:47Z letouzey $ *)
+(* $Id: OrderedTypeEx.v 10739 2008-04-01 14:45:20Z herbelin $ *)
Require Import OrderedType.
Require Import ZArith.
@@ -25,9 +25,9 @@ Require Import Compare_dec.
the equality is the usual one of Coq. *)
Module Type UsualOrderedType.
- Parameter t : Set.
+ Parameter Inline t : Type.
Definition eq := @eq t.
- Parameter lt : t -> t -> Prop.
+ Parameter Inline lt : t -> t -> Prop.
Definition eq_refl := @refl_equal t.
Definition eq_sym := @sym_eq t.
Definition eq_trans := @trans_eq t.
@@ -154,16 +154,16 @@ Module N_as_OT <: UsualOrderedType.
Definition eq_sym := @sym_eq t.
Definition eq_trans := @trans_eq t.
- Definition lt p q:= Nle q p = false.
+ Definition lt p q:= Nleb q p = false.
- Definition lt_trans := Nlt_trans.
+ Definition lt_trans := Nltb_trans.
Lemma lt_not_eq : forall x y : t, lt x y -> ~ eq x y.
Proof.
intros; intro.
rewrite H0 in H.
unfold lt in H.
- rewrite Nle_refl in H; discriminate.
+ rewrite Nleb_refl in H; discriminate.
Qed.
Definition compare : forall x y : t, Compare lt eq x y.
@@ -172,16 +172,15 @@ Module N_as_OT <: UsualOrderedType.
case_eq ((x ?= y)%N); intros.
apply EQ; apply Ncompare_Eq_eq; auto.
apply LT; unfold lt; auto.
- generalize (Nle_Ncompare y x).
- destruct (Nle y x); auto.
- rewrite <- Ncompare_antisym.
+ generalize (Nleb_Nle y x).
+ unfold Nle; rewrite <- Ncompare_antisym.
destruct (x ?= y)%N; simpl; try discriminate.
- intros (H0,_); elim H0; auto.
+ clear H; intros H.
+ destruct (Nleb y x); intuition.
apply GT; unfold lt.
- generalize (Nle_Ncompare x y).
- destruct (Nle x y); auto.
- destruct (x ?= y)%N; simpl; try discriminate.
- intros (H0,_); elim H0; auto.
+ generalize (Nleb_Nle x y).
+ unfold Nle; destruct (x ?= y)%N; simpl; try discriminate.
+ destruct (Nleb x y); intuition.
Defined.
End N_as_OT.