summaryrefslogtreecommitdiff
path: root/theories/FSets
diff options
context:
space:
mode:
Diffstat (limited to 'theories/FSets')
-rw-r--r--theories/FSets/DecidableType.v151
-rw-r--r--theories/FSets/FMapAVL.v2058
-rw-r--r--theories/FSets/FMapFacts.v557
-rw-r--r--theories/FSets/FMapIntMap.v622
-rw-r--r--theories/FSets/FMapList.v416
-rw-r--r--theories/FSets/FMapPositive.v1153
-rw-r--r--theories/FSets/FMapWeak.v5
-rw-r--r--theories/FSets/FMapWeakFacts.v599
-rw-r--r--theories/FSets/FMapWeakList.v250
-rw-r--r--theories/FSets/FMaps.v8
-rw-r--r--theories/FSets/FSetAVL.v2900
-rw-r--r--theories/FSets/FSetBridge.v10
-rw-r--r--theories/FSets/FSetEqProperties.v11
-rw-r--r--theories/FSets/FSetFacts.v10
-rw-r--r--theories/FSets/FSetInterface.v7
-rw-r--r--theories/FSets/FSetList.v315
-rw-r--r--theories/FSets/FSetProperties.v260
-rw-r--r--theories/FSets/FSetToFiniteSet.v139
-rw-r--r--theories/FSets/FSetWeak.v4
-rw-r--r--theories/FSets/FSetWeakFacts.v8
-rw-r--r--theories/FSets/FSetWeakInterface.v13
-rw-r--r--theories/FSets/FSetWeakList.v257
-rw-r--r--theories/FSets/FSetWeakProperties.v896
-rw-r--r--theories/FSets/FSets.v4
-rw-r--r--theories/FSets/OrderedType.v10
-rw-r--r--theories/FSets/OrderedTypeAlt.v129
-rw-r--r--theories/FSets/OrderedTypeEx.v248
27 files changed, 10188 insertions, 852 deletions
diff --git a/theories/FSets/DecidableType.v b/theories/FSets/DecidableType.v
deleted file mode 100644
index 635f6bdb..00000000
--- a/theories/FSets/DecidableType.v
+++ /dev/null
@@ -1,151 +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: DecidableType.v 8639 2006-03-16 19:21:55Z letouzey $ *)
-
-Require Export SetoidList.
-Set Implicit Arguments.
-Unset Strict Implicit.
-
-(** * Types with decidable Equalities (but no ordering) *)
-
-Module Type DecidableType.
-
- Parameter t : Set.
-
- Parameter eq : 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.
- Axiom eq_trans : forall x y z : t, eq x y -> eq y z -> eq x z.
-
- Parameter eq_dec : forall x y : t, { eq x y } + { ~ eq x y }.
-
- Hint Immediate eq_sym.
- Hint Resolve eq_refl eq_trans.
-
-End DecidableType.
-
-
-Module PairDecidableType(D:DecidableType).
- Import D.
-
- Section Elt.
- Variable elt : Set.
- Notation key:=t.
-
- Definition eqk (p p':key*elt) := eq (fst p) (fst p').
- Definition eqke (p p':key*elt) :=
- eq (fst p) (fst p') /\ (snd p) = (snd p').
-
- Hint Unfold eqk eqke.
- Hint Extern 2 (eqke ?a ?b) => split.
-
- (* eqke is stricter than eqk *)
-
- Lemma eqke_eqk : forall x x', eqke x x' -> eqk x x'.
- Proof.
- unfold eqk, eqke; intuition.
- Qed.
-
- (* eqk, eqke are equalities *)
-
- Lemma eqk_refl : forall e, eqk e e.
- Proof. auto. Qed.
-
- Lemma eqke_refl : forall e, eqke e e.
- Proof. auto. Qed.
-
- Lemma eqk_sym : forall e e', eqk e e' -> eqk e' e.
- Proof. auto. Qed.
-
- Lemma eqke_sym : forall e e', eqke e e' -> eqke e' e.
- Proof. unfold eqke; intuition. Qed.
-
- Lemma eqk_trans : forall e e' e'', eqk e e' -> eqk e' e'' -> eqk e e''.
- Proof. eauto. Qed.
-
- Lemma eqke_trans : forall e e' e'', eqke e e' -> eqke e' e'' -> eqke e e''.
- Proof.
- unfold eqke; intuition; [ eauto | congruence ].
- Qed.
-
- Hint Resolve eqk_trans eqke_trans eqk_refl eqke_refl.
- Hint Immediate eqk_sym eqke_sym.
-
- Lemma InA_eqke_eqk :
- forall x m, InA eqke x m -> InA eqk x m.
- Proof.
- unfold eqke; induction 1; intuition.
- Qed.
- Hint Resolve InA_eqke_eqk.
-
- Lemma InA_eqk : forall p q m, eqk p q -> InA eqk p m -> InA eqk q m.
- Proof.
- intros; apply InA_eqA with p; auto; apply eqk_trans; auto.
- Qed.
-
- Definition MapsTo (k:key)(e:elt):= InA eqke (k,e).
- Definition In k m := exists e:elt, MapsTo k e m.
-
- Hint Unfold MapsTo In.
-
- (* An alternative formulation for [In k l] is [exists e, InA eqk (k,e) l] *)
-
- Lemma In_alt : forall k l, In k l <-> exists e, InA eqk (k,e) l.
- Proof.
- firstorder.
- exists x; auto.
- induction H.
- destruct y.
- exists e; auto.
- destruct IHInA as [e H0].
- exists e; auto.
- Qed.
-
- Lemma MapsTo_eq : forall l x y e, eq x y -> MapsTo x e l -> MapsTo y e l.
- Proof.
- intros; unfold MapsTo in *; apply InA_eqA with (x,e); eauto.
- Qed.
-
- Lemma In_eq : forall l x y, eq x y -> In x l -> In y l.
- Proof.
- destruct 2 as (e,E); exists e; eapply MapsTo_eq; eauto.
- Qed.
-
- Lemma In_inv : forall k k' e l, In k ((k',e) :: l) -> eq k k' \/ In k l.
- Proof.
- inversion 1.
- inversion_clear H0; eauto.
- destruct H1; simpl in *; intuition.
- Qed.
-
- Lemma In_inv_2 : forall k k' e e' l,
- InA eqk (k, e) ((k', e') :: l) -> ~ eq k k' -> InA eqk (k, e) l.
- Proof.
- inversion_clear 1; compute in H0; intuition.
- Qed.
-
- Lemma In_inv_3 : forall x x' l,
- InA eqke x (x' :: l) -> ~ eqk x x' -> InA eqke x l.
- Proof.
- inversion_clear 1; compute in H0; intuition.
- Qed.
-
- End Elt.
-
- Hint Unfold eqk eqke.
- Hint Extern 2 (eqke ?a ?b) => split.
- Hint Resolve eqk_trans eqke_trans eqk_refl eqke_refl.
- Hint Immediate eqk_sym eqke_sym.
- Hint Resolve InA_eqke_eqk.
- Hint Unfold MapsTo In.
- Hint Resolve In_inv_2 In_inv_3.
-
-
-End PairDecidableType.
diff --git a/theories/FSets/FMapAVL.v b/theories/FSets/FMapAVL.v
new file mode 100644
index 00000000..425528ee
--- /dev/null
+++ b/theories/FSets/FMapAVL.v
@@ -0,0 +1,2058 @@
+
+(***********************************************************************)
+(* 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: FMapAVL.v 8899 2006-06-06 11:09:43Z jforest $ *)
+
+(** This module implements map using AVL trees.
+ It follows the implementation from Ocaml's standard library. *)
+
+Require Import FSetInterface.
+Require Import FMapInterface.
+Require Import FMapList.
+
+Require Import ZArith.
+Require Import 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 Scope Int_scope.
+
+Module E := X.
+Module MX := OrderedTypeFacts X.
+Module PX := KeyOrderedType X.
+Module L := FMapList.Raw X.
+Import MX.
+Import PX.
+
+Definition key := X.t.
+
+(** * Trees *)
+
+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').
+*)
+
+Notation eqk := (eqk (elt:= elt)).
+Notation eqke := (eqke (elt:= elt)).
+Notation ltk := (ltk (elt:= elt)).
+
+Inductive tree : Set :=
+ | Leaf : tree
+ | Node : tree -> key -> elt -> tree -> int -> tree.
+
+Notation t := tree.
+
+(** The Sixth field of [Node] is the height of the tree *)
+
+(** * Occurrence in a tree *)
+
+Inductive MapsTo (x : key)(e : elt) : tree -> 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',
+ MapsTo x e l -> MapsTo x e (Node l y e' r h)
+ | 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 :=
+ | 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',
+ In x l -> In x (Node l y e' r h)
+ | 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.
+
+(** * 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.
+
+(** [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).
+
+(** * 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 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).
+
+(* We should end this section before the big proofs that follows,
+ otherwise the discharge takes a lot of time. *)
+End Elt.
+
+(** Some helpful hints and tactics. *)
+
+Notation t := tree.
+Hint Constructors tree.
+Hint Constructors MapsTo.
+Hint Constructors In.
+Hint Constructors bst.
+Hint Constructors avl.
+Hint Unfold lt_tree gt_tree.
+
+Ltac inv f :=
+ match goal with
+ | H:f (Leaf _) |- _ => inversion_clear H; inv f
+ | H:f _ (Leaf _) |- _ => inversion_clear H; inv f
+ | H:f _ _ (Leaf _) |- _ => inversion_clear H; inv f
+ | H:f _ _ _ (Leaf _) |- _ => inversion_clear H; inv f
+ | H:f (Node _ _ _ _ _) |- _ => inversion_clear H; inv f
+ | H:f _ (Node _ _ _ _ _) |- _ => inversion_clear H; inv f
+ | H:f _ _ (Node _ _ _ _ _) |- _ => inversion_clear H; inv f
+ | H:f _ _ _ (Node _ _ _ _ _) |- _ => inversion_clear H; 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
+ | H: f _ _ |- _ => inversion_clear H; inv f
+ | H: f _ _ _ |- _ => inversion_clear H; inv f
+ | H: f _ _ _ _ |- _ => inversion_clear H; inv f
+ | _ => idtac
+ end.
+
+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.
+
+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 *)
+
+Ltac avl_nns :=
+ match goal with
+ | H:avl _ |- _ => avl_nn_hyp H; clear H; avl_nns
+ | _ => idtac
+ end.
+
+
+(** Facts about [MapsTo] and [In]. *)
+
+Lemma MapsTo_In : forall elt k e (m:t elt), 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.
+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.
+Proof.
+ split.
+ intros (e,H); eauto.
+ unfold In0; apply In_MapsTo; auto.
+Qed.
+
+Lemma MapsTo_1 :
+ forall elt (m:t elt) 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.
+Proof.
+ intros elt m x y; induction m; simpl; intuition_in; eauto.
+Qed.
+
+
+(** Results about [lt_tree] and [gt_tree] *)
+
+Lemma lt_leaf : forall elt 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).
+Proof.
+ unfold gt_tree in |- *; intros; intuition_in.
+Qed.
+
+Lemma lt_tree_node : forall elt x y (l:t elt) 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.
+Qed.
+
+Lemma gt_tree_node : forall elt x y (l:t elt) 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.
+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,
+ 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,
+ 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,
+ 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,
+ gt_tree x (Node l y e r h) -> gt_tree x r.
+Proof.
+ intuition_in.
+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.
+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.
+Proof.
+ firstorder eauto.
+Qed.
+
+Lemma gt_tree_not_in :
+ forall elt x (t : t elt), gt_tree x t -> ~ In x t.
+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.
+Proof.
+ firstorder eauto.
+Qed.
+
+Hint Resolve lt_tree_not_in lt_tree_trans gt_tree_not_in gt_tree_trans.
+
+(** Results about [avl] *)
+
+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)).
+Proof.
+ intros; 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]. *)
+
+Definition create elt (l:t elt) x e r :=
+ Node l x e r (max (height l) (height r) + 1).
+
+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).
+Proof.
+ unfold create; 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.
+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.
+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.
+Qed.
+
+(** trick for emulating [assert false] in Coq *)
+
+Notation assert_false := Leaf.
+
+(** [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 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 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).
+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.
+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).
+Proof.
+ bal_tac; inv avl; repeat apply create_avl; simpl in *; auto; omega_max.
+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.
+Proof.
+ bal_tac; inv avl; avl_nns; simpl in *; omega_max.
+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.
+Proof.
+ bal_tac; inv avl; simpl in *; omega_max.
+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).
+Proof.
+ bal_tac; bal_tac_imp; repeat rewrite create_in; intuition_in.
+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)).
+Proof.
+ bal_tac; bal_tac_imp; unfold create; intuition_in.
+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 *)
+
+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 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.
+Qed.
+
+Lemma add_avl : forall elt (m:t elt) x e, avl m -> avl (add x e m).
+Proof.
+ intros; generalize (add_avl_1 x e H); intuition.
+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).
+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 H1); intuition_in.
+ (* EQ *)
+ inv avl.
+ firstorder_in.
+ eapply In_1; eauto.
+ (* GT *)
+ inv avl.
+ rewrite bal_in; auto.
+ rewrite (IHt H2); 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 H1.
+ intuition.
+ eauto.
+ (* gt_tree -> gt_tree (add ...) *)
+ red; red in H5.
+ intros.
+ rewrite (add_in x y0 e H6) in H1.
+ intuition.
+ apply lt_eq with x; auto.
+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).
+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.
+
+Lemma add_2 : forall elt (m:t elt) x y e e', avl m -> ~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.
+ destruct (X.compare x k);
+ intros; inv bst; inv avl; 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 ->
+ 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;
+ try rewrite bal_mapsto; auto; unfold create; intros; inv MapsTo; auto;
+ 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.
+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 H0. simpl in *.
+ apply bal_avl; subst;auto; omega_max.
+ rewrite_all H0;simpl in *;omega_bal.
+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.
+
+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))).
+Proof.
+ intros elt 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 H1).
+
+ rewrite_all H0; simpl; intros.
+ rewrite bal_in; auto.
+ generalize (IHp lh y H1).
+ intuition.
+ inversion_clear H8; 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)))).
+Proof.
+ intros elt 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 H1).
+ rewrite_all H0; simpl; intros.
+ rewrite bal_mapsto; auto; unfold create.
+ simpl in *;destruct (IHp lh y e').
+ auto.
+ intuition.
+ inversion_clear H3; intuition.
+ inversion_clear H10; 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)).
+Proof.
+ intros elt l x e r; functional induction (remove_min l x e r); simpl in *; intros.
+ inv bst; auto.
+ inversion_clear H; inversion_clear H1.
+ apply bal_bst; auto.
+ rewrite_all H0;simpl in *;firstorder.
+ intro; intros.
+ generalize (remove_min_in y H).
+ rewrite_all H0; simpl in *.
+ destruct 1.
+ apply H4; intuition.
+Qed.
+
+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)).
+Proof.
+ intros elt l x e r; functional induction (remove_min l x e r); simpl in *; intros.
+ inv bst; auto.
+ inversion_clear H; inversion_clear H1.
+ intro; intro.
+ rewrite_all H0;simpl in *.
+ generalize (IHp lh H2 H); clear H7 H8 IHp.
+ generalize (remove_min_avl H).
+ generalize (remove_min_in (fst m) H).
+ rewrite H0; simpl; intros.
+ rewrite (bal_in x e y H8 H6) in H1.
+ destruct H7.
+ 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 H1.
+ split; auto; avl_nns; simpl in *; omega_max.
+ destruct s1;try contradiction;clear H1.
+ generalize (remove_min_avl_1 H0).
+ rewrite H2; simpl;destruct 1.
+ split.
+ apply bal_avl; auto.
+ simpl; omega_max.
+ omega_bal.
+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.
+
+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).
+Proof.
+ intros elt s1 s2; functional induction (merge s1 s2);intros.
+ intuition_in.
+ intuition_in.
+ destruct s1;try contradiction;clear H1.
+(* 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 H2; auto].
+ rewrite bal_in; auto.
+ generalize (remove_min_avl H4); rewrite H2; simpl; auto.
+ generalize (remove_min_in y H4); rewrite H2; simpl; intro.
+ rewrite H1; 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).
+Proof.
+ intros elt s1 s2; functional induction (@merge elt s1 s2); intros.
+ intuition_in.
+ intuition_in.
+ destruct s1;try contradiction;clear H1.
+ replace s2' with (fst (remove_min l2 x2 e2 r2)); [|rewrite H2; auto].
+ rewrite bal_mapsto; auto; unfold create.
+ generalize (remove_min_avl H4); rewrite H2; simpl; auto.
+ generalize (remove_min_mapsto y e0 H4); rewrite H2; simpl; intro.
+ rewrite H1; intuition (try 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).
+Proof.
+ intros elt s1 s2; functional induction (@merge elt s1 s2); intros; auto.
+
+ apply bal_bst; auto.
+ destruct s1;try contradiction.
+ generalize (remove_min_bst H3); rewrite H2; simpl in *; auto.
+ destruct s1;try contradiction.
+ intro; intro.
+ apply H5; auto.
+ generalize (remove_min_in x H4); rewrite H2; simpl; intuition.
+ destruct s1;try contradiction.
+ generalize (remove_min_gt_tree H3); rewrite H2; 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 H1).
+ split.
+ apply bal_avl; auto.
+ omega_max.
+ omega_bal.
+ (* EQ *)
+ inv avl.
+ generalize (merge_avl_1 H1 H2 H3).
+ intuition omega_max.
+ (* GT *)
+ inv avl.
+ destruct (IHt H2).
+ split.
+ apply bal_avl; auto.
+ omega_max.
+ omega_bal.
+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.
+
+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).
+Proof.
+ intros elt s x; functional induction (@remove elt x s); simpl; intros.
+ intuition_in.
+ (* LT *)
+ inv avl; inv bst; clear H0.
+ rewrite bal_in; auto.
+ generalize (IHt y0 H1); intuition; [ order | order | intuition_in ].
+ (* EQ *)
+ inv avl; inv bst; clear H0.
+ rewrite merge_in; intuition; [ order | order | intuition_in ].
+ elim H9; eauto.
+ (* GT *)
+ inv avl; inv bst; clear H0.
+ rewrite bal_in; auto.
+ generalize (IHt y0 H6); intuition; [ order | order | intuition_in ].
+Qed.
+
+Lemma remove_bst : forall elt (s:t elt) x, bst s -> avl s -> bst (remove x s).
+Proof.
+ intros elt s x; functional induction (@remove elt x s); simpl; intros.
+ auto.
+ (* LT *)
+ inv avl; inv bst.
+ apply bal_bst; auto.
+ intro; intro.
+ rewrite (remove_in x y0 H1) in H; auto.
+ destruct H; eauto.
+ (* EQ *)
+ inv avl; inv bst.
+ apply merge_bst; eauto.
+ (* GT *)
+ inv avl; inv bst.
+ apply bal_bst; auto.
+ intro; intro.
+ rewrite (remove_in x y0 H6) 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).
+Proof.
+ intros; rewrite remove_in; intuition.
+Qed.
+
+Lemma remove_2 : forall elt (m:t elt) x y e, bst m -> avl 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.
+ destruct (X.compare x k);
+ intros; inv bst; inv avl; 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 ->
+ 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;
+ try rewrite bal_mapsto; auto; unfold create.
+ intros; inv MapsTo; auto.
+ rewrite merge_mapsto; intuition.
+ intros; inv MapsTo; auto.
+Qed.
+
+Section Elt2.
+
+Variable elt:Set.
+
+Notation eqk := (eqk (elt:= elt)).
+Notation eqke := (eqke (elt:= elt)).
+Notation ltk := (ltk (elt:= elt)).
+
+(** * Empty map *)
+
+Definition Empty m := forall (a : key)(e:elt) , ~ MapsTo a e m.
+
+Definition empty := (Leaf elt).
+
+Lemma empty_bst : bst empty.
+Proof.
+ unfold empty; auto.
+Qed.
+
+Lemma empty_avl : avl empty.
+Proof.
+ unfold empty; auto.
+Qed.
+
+Lemma empty_1 : Empty empty.
+Proof.
+ unfold empty, Empty; intuition_in.
+Qed.
+
+(** * Emptyness test *)
+
+Definition is_empty (s:t elt) := 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 e l h]; simpl; auto.
+ intro H; elim (H x e); auto.
+Qed.
+
+Lemma is_empty_2 : forall s, is_empty s = true -> Empty s.
+Proof.
+ destruct s; simpl; intros; try discriminate; red; intuition_in.
+Qed.
+
+(** * Appartness *)
+
+(** The [mem] function is deciding appartness. It exploits the [bst] property
+ to achieve logarithmic complexity. *)
+
+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.
+
+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.
+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.
+
+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.
+ 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.
+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.
+Qed.
+
+(** * 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)) (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.
+
+(** then [elements] is an instanciation with an empty [acc] *)
+
+Definition elements := elements_aux nil.
+
+Lemma elements_aux_mapsto : forall s 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.
+ intuition.
+ inversion H0.
+ intros.
+ rewrite Hl.
+ destruct (Hr acc x0 e0); clear Hl Hr.
+ intuition; inversion_clear H3; intuition.
+ destruct H0; simpl in *; subst; intuition.
+Qed.
+
+Lemma elements_mapsto : forall s 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.
+Proof.
+ intros.
+ unfold L.PX.In.
+ rewrite <- In_alt; unfold In0.
+ firstorder.
+ exists x0.
+ rewrite <- elements_mapsto; auto.
+ exists x0.
+ unfold L.PX.MapsTo; rewrite elements_mapsto; auto.
+Qed.
+
+Lemma elements_aux_sort : forall s 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.
+ induction s as [ | l Hl y e r Hr h]; simpl; intuition.
+ inv bst.
+ apply Hl; auto.
+ constructor.
+ apply Hr; eauto.
+ apply (InA_InfA (eqke_refl (elt:=elt))); intros (y',e') H6.
+ destruct (elements_aux_mapsto r acc y' e'); intuition.
+ red; simpl; eauto.
+ red; simpl; eauto.
+ intros.
+ inversion_clear H.
+ destruct H7; simpl in *.
+ order.
+ destruct (elements_aux_mapsto r acc x e0); intuition eauto.
+Qed.
+
+Lemma elements_sort : forall s : t elt, bst s -> sort ltk (elements s).
+Proof.
+ intros; unfold elements; apply elements_aux_sort; auto.
+ intros; inversion H0.
+Qed.
+Hint Resolve elements_sort.
+
+
+(** * Fold *)
+
+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.
+
+Definition fold' (A : Set) (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,
+ L.fold f (elements_aux acc s) a = L.fold f acc (fold f s a).
+Proof.
+ simple induction s.
+ simpl in |- *; intuition.
+ simpl in |- *; intros.
+ rewrite H.
+ simpl.
+ apply H0.
+Qed.
+
+Lemma fold_equiv :
+ forall (A : Set) (s : t elt) (f : key -> elt -> A -> A) (a : A),
+ fold f s a = fold' f s a.
+Proof.
+ unfold fold', elements in |- *.
+ simple induction s; simpl in |- *; auto; intros.
+ rewrite fold_equiv_aux.
+ rewrite H0.
+ simpl in |- *; auto.
+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.
+Proof.
+ intros.
+ rewrite fold_equiv.
+ unfold fold'.
+ rewrite L.fold_1.
+ unfold L.elements; auto.
+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 *)
+
+Fixpoint flatten_e (e : enumeration) : 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.
+Qed.
+
+Lemma sorted_flatten_e :
+ forall e : enumeration, sorted_e e -> sort ltk (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.
+Qed.
+
+Lemma elements_app :
+ forall s acc, 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 t1 t2 x e z l,
+ elements t1 ++ (x,e) :: elements t2 ++ l =
+ elements (Node t1 x e t2 z) ++ l.
+Proof.
+ simpl in |- *; unfold elements in |- *; simpl in |- *; intuition.
+ repeat rewrite elements_app.
+ repeat rewrite <- app_nil_end.
+ repeat rewrite app_ass; 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.
+Proof.
+ intros; simpl.
+ apply compare_flatten_1.
+Qed.
+
+Open 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.
+Proof.
+ simple induction s.
+ simpl in |- *; omega.
+ intros.
+ Measure_e_t; omega.
+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.
+Proof.
+ simple induction e.
+ simpl in |- *; omega.
+ intros.
+ Measure_e; Measure_e_t_0 t; omega.
+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'.
+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.
+do 2 rewrite elements_in; firstorder.
+destruct H.
+apply (H2 k); rewrite <- elements_mapsto; auto.
+do 2 rewrite <- elements_in; firstorder.
+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'}.
+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.
+Qed.
+
+End Elt2.
+
+Section Elts.
+
+Variable elt elt' elt'' : Set.
+
+Section Map.
+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).
+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.
+Proof.
+induction m; simpl; inversion_clear 1; auto.
+Qed.
+
+Lemma map_bst : forall m, bst m -> bst (map 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.
+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.
+
+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).
+Proof.
+induction m; simpl; inversion_clear 1; auto.
+exists k; auto.
+destruct (IHm1 _ _ H0).
+exists x0; intuition.
+destruct (IHm2 _ _ H0).
+exists x0; intuition.
+Qed.
+
+Lemma mapi_2 : forall (m: t elt)(x:key),
+ In x (mapi m) -> In x m.
+Proof.
+induction m; simpl; inversion_clear 1; auto.
+Qed.
+
+Lemma mapi_bst : forall m, bst m -> bst (mapi 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.
+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 _).
+
+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).
+Proof.
+unfold anti_elements; induction l.
+simpl; auto.
+simpl; destruct a; intros.
+apply IHl.
+apply add_avl; auto.
+Qed.
+
+Lemma anti_elements_avl : forall l, avl (anti_elements l).
+Proof.
+unfold anti_elements, empty; intros; apply anti_elements_avl_aux; auto.
+Qed.
+
+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.
+
+Lemma anti_elements_bst : forall l, bst (anti_elements l).
+Proof.
+unfold anti_elements, empty; intros; apply anti_elements_bst_aux; auto.
+Qed.
+
+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.
+
+Lemma map2_avl : forall (m: t elt)(m': t elt'), avl (map2 m m').
+Proof.
+unfold map2; intros; apply anti_elements_avl; auto.
+Qed.
+
+Lemma map2_bst : forall (m: t elt)(m': t elt'), bst (map2 m 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.
+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'.
+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.
+Qed.
+
+End Map2.
+End Elts.
+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. *)
+
+Module IntMake (I:Int)(X: OrderedType) <: S with Module E := X.
+
+ Module E := X.
+ Module Raw := Raw I X.
+
+ Record bbst (elt:Set) : Set :=
+ Bbst {this :> Raw.tree elt; is_bst : Raw.bst this; is_avl: Raw.avl this}.
+
+ Definition t := bbst.
+ Definition key := E.t.
+
+ Section Elt.
+ Variable elt elt' elt'': Set.
+
+ 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 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 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 mapi (f:key->elt->elt') m : t elt' :=
+ Bbst (Raw.mapi_bst f m.(is_bst)) (Raw.mapi_avl f m.(is_avl)).
+ Definition map2 f m (m':t elt') : t elt'' :=
+ Bbst (Raw.map2_bst f m m') (Raw.map2_avl f m m').
+ 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 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 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.
+
+ 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.
+
+ 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.
+ 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.
+ Qed.
+
+ Lemma empty_1 : Empty empty.
+ Proof. exact (@Raw.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.
+ Lemma is_empty_2 : forall m, is_empty m = true -> Empty m.
+ Proof. intros m; exact (@Raw.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.
+ 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.
+ 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.
+
+ 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.
+ 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.
+ 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.
+
+
+ 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.
+ 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.
+
+ Lemma fold_1 : forall m (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.
+ Proof. intros m; exact (@Raw.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.
+ 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.
+ 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.
+
+ 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).
+
+ Lemma Equal_Equal : forall cmp m m', Equal cmp m m' <-> Raw.Equal 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.
+
+ Lemma equal_1 : forall m m' cmp,
+ Equal 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.
+ Qed.
+
+ Lemma equal_2 : forall m m' cmp,
+ equal cmp m m' = true -> Equal cmp m m'.
+ Proof.
+ unfold equal; intros; rewrite Equal_Equal.
+ destruct (@Raw.equal _ cmp m m'); auto; try discriminate.
+ Qed.
+
+ End Elt.
+
+ 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. intros elt elt' m x e f; exact (@Raw.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.
+ Proof.
+ intros elt elt' m x f; do 2 unfold In in *; do 2 rewrite Raw.In_alt; simpl.
+ apply Raw.map_2; 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 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)
+ (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.
+
+ Lemma 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').
+ 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.
+ apply m.(is_bst).
+ apply m'.(is_bst).
+ Qed.
+
+ Lemma 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'.
+ 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.
+ 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 MapS := IntMake(I)(X).
+ Import MapS.
+
+ Module MD := OrderedTypeFacts(D).
+ Import MD.
+
+ 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.elements_sort m.(is_bst)).
+
+ Definition eq : t -> t -> Prop :=
+ fun m1 m2 => LO.eq (elements m1) (elements m2).
+
+ Definition lt : t -> t -> Prop :=
+ fun m1 m2 => LO.lt (elements m1) (elements m2).
+
+ Lemma eq_1 : forall m m', Equal 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.
+ Qed.
+
+ Lemma eq_2 : forall m m', eq m m' -> Equal cmp m m'.
+ Proof.
+ intros m m'.
+ unfold eq.
+ rewrite Equal_Equal.
+ rewrite Raw.Equal_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.
+ Qed.
+
+ Lemma eq_sym : forall m1 m2 : t, eq m1 m2 -> eq m2 m1.
+ Proof.
+ unfold eq; 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.
+ 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.
+ 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 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.
+ 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/FMapFacts.v b/theories/FSets/FMapFacts.v
new file mode 100644
index 00000000..0105095a
--- /dev/null
+++ b/theories/FSets/FMapFacts.v
@@ -0,0 +1,557 @@
+(***********************************************************************)
+(* 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: FMapFacts.v 8882 2006-05-31 21:55:30Z letouzey $ *)
+
+(** * Finite maps library *)
+
+(** This functor derives additional facts from [FMapInterface.S]. These
+ facts are mainly the specifications of [FMapInterface.S] written using
+ different styles: equivalence and boolean equalities.
+*)
+
+Require Import Bool.
+Require Import 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] *)
+
+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 (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 (eq_dec x y) as [E|E]; auto.
+right; exists e'; auto.
+apply (add_3 E H).
+destruct (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.
+
+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 eq_dec x y then Some e else find y m.
+Proof.
+intros; destruct (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 (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 eq_dec x y then None else find y m.
+Proof.
+intros; destruct (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 (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.
+
+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)).
+ 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).
+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));
+ 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 (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 (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/FMapIntMap.v b/theories/FSets/FMapIntMap.v
new file mode 100644
index 00000000..c7681bd4
--- /dev/null
+++ b/theories/FSets/FMapIntMap.v
@@ -0,0 +1,622 @@
+(***********************************************************************)
+(* 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/FMapList.v b/theories/FSets/FMapList.v
index 2d083d5b..c671ba82 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 8667 2006-03-28 11:59:44Z letouzey $ *)
+(* $Id: FMapList.v 8899 2006-06-06 11:09:43Z jforest $ *)
(** * Finite map library *)
@@ -26,7 +26,7 @@ Module Raw (X:OrderedType).
Module E := X.
Module MX := OrderedTypeFacts X.
-Module PX := PairOrderedType X.
+Module PX := KeyOrderedType X.
Import MX.
Import PX.
@@ -36,7 +36,7 @@ Definition t (elt:Set) := list (X.t * elt).
Section Elt.
Variable elt : Set.
-(* Now in PairOrderedtype:
+(* 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').
@@ -96,7 +96,7 @@ Qed.
(** * [mem] *)
-Fixpoint mem (k : key) (s : t elt) {struct s} : bool :=
+Function mem (k : key) (s : t elt) {struct s} : bool :=
match s with
| nil => false
| (k',_) :: l =>
@@ -110,33 +110,33 @@ Fixpoint mem (k : key) (s : t elt) {struct s} : bool :=
Lemma mem_1 : forall m (Hm:Sort m) x, In x m -> mem x m = true.
Proof.
intros m Hm x; generalize Hm; clear Hm.
- functional induction mem x m;intros sorted belong1;trivial.
+ functional induction (mem x m);intros sorted belong1;trivial.
inversion belong1. inversion H.
- absurd (In k ((k', e) :: l));try assumption.
- apply Sort_Inf_NotIn with e;auto.
+ absurd (In x ((k', _x) :: l));try assumption.
+ apply Sort_Inf_NotIn with _x;auto.
- apply H.
+ apply IHb.
elim (sort_inv sorted);auto.
elim (In_inv belong1);auto.
intro abs.
- absurd (X.eq k k');auto.
+ absurd (X.eq x k');auto.
Qed.
Lemma mem_2 : forall m (Hm:Sort m) x, mem x m = true -> In x m.
Proof.
intros m Hm x; generalize Hm; clear Hm; unfold PX.In,PX.MapsTo.
- functional induction mem x m; intros sorted hyp;try ((inversion hyp);fail).
- exists e; auto.
- induction H; auto.
- exists x; auto.
+ functional induction (mem x m); intros sorted hyp;try ((inversion hyp);fail).
+ exists _x; auto.
+ induction IHb; auto.
+ exists x0; auto.
inversion_clear sorted; auto.
Qed.
(** * [find] *)
-Fixpoint find (k:key) (s: t elt) {struct s} : option elt :=
+Function find (k:key) (s: t elt) {struct s} : option elt :=
match s with
| nil => None
| (k',x)::s' =>
@@ -150,31 +150,31 @@ Fixpoint find (k:key) (s: t elt) {struct s} : option elt :=
Lemma find_2 : forall m x e, find x m = Some e -> MapsTo x e m.
Proof.
intros m x. unfold PX.MapsTo.
- functional induction find x m;simpl;intros e' eqfind; inversion eqfind; auto.
+ functional induction (find x m);simpl;intros e' eqfind; inversion eqfind; auto.
Qed.
Lemma find_1 : forall m (Hm:Sort m) x e, MapsTo x e m -> find x m = Some e.
Proof.
intros m Hm x e; generalize Hm; clear Hm; unfold PX.MapsTo.
- functional induction find x m;simpl; subst; try clear H_eq_1.
+ functional induction (find x m);simpl; subst; try clear H_eq_1.
inversion 2.
inversion_clear 2.
- compute in H0; destruct H0; order.
- generalize (Sort_In_cons_1 Hm (InA_eqke_eqk H0)); compute; order.
+ clear H0;compute in H1; destruct H1;order.
+ clear H0;generalize (Sort_In_cons_1 Hm (InA_eqke_eqk H1)); compute; order.
- inversion_clear 2.
+ clear H0;inversion_clear 2.
compute in H0; destruct H0; intuition congruence.
generalize (Sort_In_cons_1 Hm (InA_eqke_eqk H0)); compute; order.
- do 2 inversion_clear 1; auto.
- compute in H3; destruct H3; order.
+ clear H0; do 2 inversion_clear 1; auto.
+ compute in H2; destruct H2; order.
Qed.
(** * [add] *)
-Fixpoint add (k : key) (x : elt) (s : t elt) {struct s} : t elt :=
+Function add (k : key) (x : elt) (s : t elt) {struct s} : t elt :=
match s with
| nil => (k,x) :: nil
| (k',y) :: l =>
@@ -189,7 +189,7 @@ Lemma add_1 : forall m x y e, X.eq x y -> MapsTo y e (add x e m).
Proof.
intros m x y e; generalize y; clear y.
unfold PX.MapsTo.
- functional induction add x e m;simpl;auto.
+ functional induction (add x e m);simpl;auto.
Qed.
Lemma add_2 : forall m x y e e',
@@ -197,25 +197,29 @@ Lemma add_2 : forall m x y e e',
Proof.
intros m x y e e'.
generalize y e; clear y e; unfold PX.MapsTo.
- functional induction add x e' m;simpl;auto; clear H_eq_1.
- intros y' e' eqky'; inversion_clear 1; destruct H0; simpl in *.
+ functional induction (add x e' m) ;simpl;auto; clear H0.
+ subst;auto.
+
+ intros y' e'' eqky'; inversion_clear 1; destruct H0; simpl in *.
order.
auto.
auto.
- intros y' e' eqky'; inversion_clear 1; intuition.
+ intros y' e'' eqky'; inversion_clear 1; intuition.
Qed.
+
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 m x y e e'. generalize y e; clear y e; unfold PX.MapsTo.
- functional induction add x e' m;simpl; intros.
- apply (In_inv_3 H0); compute; auto.
+ functional induction (add x e' m);simpl; intros.
apply (In_inv_3 H0); compute; auto.
- constructor 2; apply (In_inv_3 H0); compute; auto.
+ apply (In_inv_3 H1); compute; auto.
+ constructor 2; apply (In_inv_3 H1); compute; auto.
inversion_clear H1; auto.
Qed.
+
Lemma add_Inf : forall (m:t elt)(x x':key)(e e':elt),
Inf (x',e') m -> ltk (x',e') (x,e) -> Inf (x',e') (add x e m).
Proof.
@@ -242,7 +246,7 @@ Qed.
(** * [remove] *)
-Fixpoint remove (k : key) (s : t elt) {struct s} : t elt :=
+Function remove (k : key) (s : t elt) {struct s} : t elt :=
match s with
| nil => nil
| (k',x) :: l =>
@@ -256,30 +260,36 @@ Fixpoint remove (k : key) (s : t elt) {struct s} : t elt :=
Lemma remove_1 : forall m (Hm:Sort m) x y, X.eq x y -> ~ In y (remove x m).
Proof.
intros m Hm x y; generalize Hm; clear Hm.
- functional induction remove x m;simpl;intros;subst;try clear H_eq_1.
+ functional induction (remove x m);simpl;intros;subst.
red; inversion 1; inversion H1.
- apply Sort_Inf_NotIn with x; auto.
- constructor; compute; order.
+ apply Sort_Inf_NotIn with x0; auto.
+ clear H0;constructor; compute; order.
- inversion_clear Hm.
- apply Sort_Inf_NotIn with x; auto.
- apply Inf_eq with (k',x);auto; compute; apply X.eq_trans with k; auto.
+ clear H0;inversion_clear Hm.
+ apply Sort_Inf_NotIn with x0; auto.
+ apply Inf_eq with (k',x0);auto; compute; apply X.eq_trans with x; auto.
- inversion_clear Hm.
- assert (notin:~ In y (remove k l)) by auto.
- intros (x0,abs).
+ clear H0;inversion_clear Hm.
+ assert (notin:~ In y (remove x l)) by auto.
+ intros (x1,abs).
inversion_clear abs.
- compute in H3; destruct H3; order.
- apply notin; exists x0; auto.
+ compute in H2; destruct H2; order.
+ apply notin; exists x1; auto.
Qed.
+
Lemma remove_2 : forall m (Hm:Sort m) x y e,
~ X.eq x y -> MapsTo y e m -> MapsTo y e (remove x m).
Proof.
intros m Hm x y e; generalize Hm; clear Hm; unfold PX.MapsTo.
- functional induction remove x m;auto; try clear H_eq_1.
+ functional induction (remove x m);subst;auto;
+ match goal with
+ | [H: X.compare _ _ = _ |- _ ] => clear H
+ | _ => idtac
+ end.
+
inversion_clear 3; auto.
compute in H1; destruct H1; order.
@@ -290,7 +300,7 @@ Lemma remove_3 : forall m (Hm:Sort m) x y e,
MapsTo y e (remove x m) -> MapsTo y e m.
Proof.
intros m Hm x y e; generalize Hm; clear Hm; unfold PX.MapsTo.
- functional induction remove x m;auto.
+ functional induction (remove x m);subst;auto.
inversion_clear 1; inversion_clear 1; auto.
Qed.
@@ -341,8 +351,7 @@ Qed.
(** * [fold] *)
-Fixpoint fold (A:Set)(f:key->elt->A->A)(m:t elt) {struct m} : A -> A :=
- fun acc =>
+Function fold (A:Set)(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)
@@ -351,12 +360,12 @@ Fixpoint fold (A:Set)(f:key->elt->A->A)(m:t elt) {struct m} : A -> A :=
Lemma fold_1 : forall m (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.
Proof.
- intros; functional induction fold A f m i; auto.
+ intros; functional induction (fold f m i); auto.
Qed.
(** * [equal] *)
-Fixpoint equal (cmp:elt->elt->bool)(m m' : t elt) { struct m } : bool :=
+Function equal (cmp:elt->elt->bool)(m m' : t elt) { struct m } : bool :=
match m, m' with
| nil, nil => true
| (x,e)::l, (x',e')::l' =>
@@ -375,56 +384,52 @@ Lemma equal_1 : forall m (Hm:Sort m) m' (Hm': Sort m') cmp,
Equal 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; auto; unfold Equal;
- intuition; subst; try clear H_eq_3.
+ functional induction (equal cmp m m'); simpl; subst;auto; unfold Equal;
+ intuition; subst; match goal with
+ | [H: X.compare _ _ = _ |- _ ] => clear H
+ | _ => idtac
+ end.
- destruct p as (k,e).
- destruct (H0 k).
- destruct H2.
- exists e; auto.
- inversion H2.
- destruct (H0 x).
- destruct H.
- exists e; auto.
- inversion H.
- destruct (H0 x).
- assert (In x ((x',e')::l')).
- apply H; auto.
- exists e; auto.
- destruct (In_inv H3).
- order.
- inversion_clear Hm'.
- assert (Inf (x,e) l').
- apply Inf_lt with (x',e'); auto.
- elim (Sort_Inf_NotIn H5 H7 H4).
-
- assert (cmp e e' = true).
+ assert (cmp_e_e':cmp e e' = true).
apply H2 with x; auto.
- rewrite H0; simpl.
- apply H; auto.
+ rewrite cmp_e_e'; simpl.
+ apply IHb; auto.
inversion_clear Hm; auto.
inversion_clear Hm'; auto.
unfold Equal; intuition.
- destruct (H1 k).
+ destruct (H0 k).
assert (In k ((x,e) ::l)).
- destruct H3 as (e'', hyp); exists e''; auto.
- destruct (In_inv (H4 H6)); auto.
+ destruct H as (e'', hyp); exists e''; auto.
+ destruct (In_inv (H1 H4)); auto.
inversion_clear Hm.
- elim (Sort_Inf_NotIn H8 H9).
- destruct H3 as (e'', hyp); exists e''; auto.
+ elim (Sort_Inf_NotIn H6 H7).
+ destruct H as (e'', hyp); exists e''; auto.
apply MapsTo_eq with k; auto; order.
- destruct (H1 k).
+ destruct (H0 k).
assert (In k ((x',e') ::l')).
- destruct H3 as (e'', hyp); exists e''; auto.
- destruct (In_inv (H5 H6)); auto.
+ destruct H as (e'', hyp); exists e''; auto.
+ destruct (In_inv (H3 H4)); auto.
inversion_clear Hm'.
- elim (Sort_Inf_NotIn H8 H9).
- destruct H3 as (e'', hyp); exists e''; auto.
+ elim (Sort_Inf_NotIn H6 H7).
+ destruct H as (e'', hyp); exists e''; auto.
apply MapsTo_eq with k; auto; order.
apply H2 with k; destruct (eq_dec x k); auto.
+
+ destruct (X.compare x x'); try contradiction;clear H2.
+ destruct (H0 x).
+ assert (In x ((x',e')::l')).
+ apply H; auto.
+ exists e; auto.
+ destruct (In_inv H3).
+ order.
+ inversion_clear Hm'.
+ assert (Inf (x,e) l').
+ apply Inf_lt with (x',e'); auto.
+ elim (Sort_Inf_NotIn H5 H7 H4).
+
destruct (H0 x').
assert (In x' ((x,e)::l)).
apply H2; auto.
@@ -435,43 +440,70 @@ Proof.
assert (Inf (x',e') l).
apply Inf_lt with (x,e); auto.
elim (Sort_Inf_NotIn H5 H7 H4).
+
+ destruct m;
+ destruct m';try contradiction.
+
+ clear H1;destruct p as (k,e).
+ destruct (H0 k).
+ destruct H1.
+ exists e; auto.
+ inversion H1.
+
+ destruct p as (x,e).
+ destruct (H0 x).
+ destruct H.
+ exists e; auto.
+ inversion H.
+
+ destruct p;destruct p0;contradiction.
Qed.
+
Lemma equal_2 : forall m (Hm:Sort m) m' (Hm:Sort m') cmp,
equal cmp m m' = true -> Equal cmp m m'.
Proof.
intros m Hm m' Hm' cmp; generalize Hm Hm'; clear Hm Hm'.
- functional induction equal cmp m m'; simpl; auto; unfold Equal;
- intuition; try discriminate; subst; try clear H_eq_3;
- try solve [inversion H0]; destruct (andb_prop _ _ H0); clear H0;
- inversion_clear Hm; inversion_clear Hm'.
-
- destruct (H H0 H5 H3).
- destruct (In_inv H1).
+ 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.
+
+ inversion H0.
+
+ inversion_clear Hm;inversion_clear Hm'.
+ destruct (andb_prop _ _ H); clear H.
+ destruct (IHb H1 H3 H6).
+ destruct (In_inv H0).
exists e'; constructor; split; trivial; apply X.eq_trans with x; auto.
- destruct (H7 k).
- destruct (H10 H9) as (e'',hyp).
+ destruct (H k).
+ destruct (H9 H8) as (e'',hyp).
exists e''; auto.
- destruct (H H0 H5 H3).
- destruct (In_inv H1).
+ inversion_clear Hm;inversion_clear Hm'.
+ destruct (andb_prop _ _ H); clear H.
+ destruct (IHb H1 H3 H6).
+ destruct (In_inv H0).
exists e; constructor; split; trivial; apply X.eq_trans with x'; auto.
- destruct (H7 k).
- destruct (H11 H9) as (e'',hyp).
+ destruct (H k).
+ destruct (H10 H8) as (e'',hyp).
exists e''; auto.
- destruct (H H0 H6 H4).
- inversion_clear H1.
- destruct H10; simpl in *; subst.
+ inversion_clear Hm;inversion_clear Hm'.
+ destruct (andb_prop _ _ H); clear H.
+ destruct (IHb H1 H4 H7).
+ inversion_clear H0.
+ destruct H9; simpl in *; subst.
inversion_clear H2.
- destruct H10; simpl in *; subst; auto.
- elim (Sort_Inf_NotIn H6 H7).
+ destruct H9; simpl in *; subst; auto.
+ elim (Sort_Inf_NotIn H4 H5).
exists e'0; apply MapsTo_eq with k; auto; order.
inversion_clear H2.
- destruct H1; simpl in *; subst; auto.
- elim (Sort_Inf_NotIn H0 H5).
- exists e1; apply MapsTo_eq with k; auto; order.
- apply H9 with k; auto.
+ destruct H0; simpl in *; subst; auto.
+ elim (Sort_Inf_NotIn H1 H3).
+ exists e0; apply MapsTo_eq with k; auto; order.
+ apply H8 with k; auto.
Qed.
(** This lemma isn't part of the spec of [Equal], but is used in [FMapAVL] *)
@@ -791,7 +823,7 @@ Proof.
exact (combine_lelistA _ H0 H1).
inversion_clear Hm; inversion_clear Hm'.
constructor; auto.
- assert (lelistA (ltk (elt:=elt')) (k, e') m') by apply Inf_eq with (k',e'); auto.
+ assert (lelistA (ltk (elt:=elt')) (k, e') m') by (apply Inf_eq with (k',e'); auto).
exact (combine_lelistA _ H0 H3).
inversion_clear Hm; inversion_clear Hm'.
constructor; auto.
@@ -1006,84 +1038,126 @@ Module Make (X: OrderedType) <: S with Module E := X.
Module Raw := Raw X.
Module E := X.
-Definition key := X.t.
+Definition key := E.t.
Record slist (elt:Set) : Set :=
{this :> Raw.t elt; sorted : sort (@Raw.PX.ltk elt) this}.
-Definition t (elt:Set) := slist elt.
+Definition t (elt:Set) : Set := slist elt.
Section Elt.
Variable elt elt' elt'':Set.
Implicit Types m : t elt.
-
- Definition empty := Build_slist (Raw.empty_sorted elt).
- Definition is_empty m := Raw.is_empty m.(this).
- Definition add x e m := Build_slist (Raw.add_sorted m.(sorted) x e).
- Definition find x m := Raw.find x m.(this).
- Definition remove x m := Build_slist (Raw.remove_sorted m.(sorted) x).
- Definition mem x m := Raw.mem x m.(this).
+ Implicit Types x y : key.
+ Implicit Types e : elt.
+
+ Definition empty : t elt := Build_slist (Raw.empty_sorted elt).
+ Definition is_empty m : bool := Raw.is_empty m.(this).
+ Definition add x e m : t elt := Build_slist (Raw.add_sorted m.(sorted) x e).
+ Definition find x m : option elt := Raw.find x m.(this).
+ Definition remove x m : t elt := Build_slist (Raw.remove_sorted m.(sorted) x).
+ Definition mem x m : bool := Raw.mem x m.(this).
Definition map f m : t elt' := Build_slist (Raw.map_sorted m.(sorted) f).
- Definition mapi f m : t elt' := Build_slist (Raw.mapi_sorted m.(sorted) f).
+ Definition mapi (f:key->elt->elt') m : t elt' := Build_slist (Raw.mapi_sorted m.(sorted) f).
Definition map2 f m (m':t elt') : t elt'' :=
- Build_slist (Raw.map2_sorted f m.(sorted) m'.(sorted)).
- Definition elements m := @Raw.elements elt m.(this).
- Definition fold A f m i := @Raw.fold elt A f m.(this) i.
- Definition equal cmp m m' := @Raw.equal elt cmp m.(this) m'.(this).
-
- Definition MapsTo x e m := Raw.PX.MapsTo x e m.(this).
- Definition In x m := Raw.PX.In x m.(this).
- Definition Empty m := Raw.Empty m.(this).
- Definition Equal cmp m m' := @Raw.Equal elt cmp m.(this) m'.(this).
-
- Definition eq_key := Raw.PX.eqk.
- Definition eq_key_elt := Raw.PX.eqke.
- Definition lt_key := Raw.PX.ltk.
-
- Definition MapsTo_1 m := @Raw.PX.MapsTo_eq elt m.(this).
-
- Definition mem_1 m := @Raw.mem_1 elt m.(this) m.(sorted).
- Definition mem_2 m := @Raw.mem_2 elt m.(this) m.(sorted).
-
- Definition empty_1 := @Raw.empty_1.
-
- Definition is_empty_1 m := @Raw.is_empty_1 elt m.(this).
- Definition is_empty_2 m := @Raw.is_empty_2 elt m.(this).
-
- Definition add_1 m := @Raw.add_1 elt m.(this).
- Definition add_2 m := @Raw.add_2 elt m.(this).
- Definition add_3 m := @Raw.add_3 elt m.(this).
-
- Definition remove_1 m := @Raw.remove_1 elt m.(this) m.(sorted).
- Definition remove_2 m := @Raw.remove_2 elt m.(this) m.(sorted).
- Definition remove_3 m := @Raw.remove_3 elt m.(this) m.(sorted).
-
- Definition find_1 m := @Raw.find_1 elt m.(this) m.(sorted).
- Definition find_2 m := @Raw.find_2 elt m.(this).
-
- Definition elements_1 m := @Raw.elements_1 elt m.(this).
- Definition elements_2 m := @Raw.elements_2 elt m.(this).
- Definition elements_3 m := @Raw.elements_3 elt m.(this) m.(sorted).
-
- Definition fold_1 m := @Raw.fold_1 elt m.(this).
-
- Definition map_1 m := @Raw.map_1 elt elt' m.(this).
- Definition map_2 m := @Raw.map_2 elt elt' m.(this).
-
- Definition mapi_1 m := @Raw.mapi_1 elt elt' m.(this).
- Definition mapi_2 m := @Raw.mapi_2 elt elt' m.(this).
-
- Definition map2_1 m (m':t elt') x f :=
- @Raw.map2_1 elt elt' elt'' f m.(this) m.(sorted) m'.(this) m'.(sorted) x.
- Definition map2_2 m (m':t elt') x f :=
- @Raw.map2_2 elt elt' elt'' f m.(this) m.(sorted) m'.(this) m'.(sorted) x.
-
- Definition equal_1 m m' :=
- @Raw.equal_1 elt m.(this) m.(sorted) m'.(this) m'.(sorted).
- Definition equal_2 m m' :=
- @Raw.equal_2 elt m.(this) m.(sorted) m'.(this) m'.(sorted).
+ 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 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 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.
+
+ 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.PX.MapsTo_eq elt m.(this)). Qed.
+
+ Lemma mem_1 : forall m x, In x m -> mem x m = true.
+ Proof. intros m; exact (@Raw.mem_1 elt m.(this) m.(sorted)). Qed.
+ Lemma mem_2 : forall m x, mem x m = true -> In x m.
+ Proof. intros m; exact (@Raw.mem_2 elt m.(this) m.(sorted)). Qed.
+
+ Lemma empty_1 : Empty empty.
+ Proof. exact (@Raw.empty_1 elt). Qed.
+
+ Lemma is_empty_1 : forall m, Empty m -> is_empty m = true.
+ Proof. intros m; exact (@Raw.is_empty_1 elt m.(this)). Qed.
+ Lemma is_empty_2 : forall m, is_empty m = true -> Empty m.
+ Proof. intros m; exact (@Raw.is_empty_2 elt 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; exact (@Raw.add_1 elt m.(this)). 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; exact (@Raw.add_2 elt m.(this)). 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; exact (@Raw.add_3 elt m.(this)). Qed.
+
+ Lemma remove_1 : forall m x y, E.eq x y -> ~ In y (remove x m).
+ Proof. intros m; exact (@Raw.remove_1 elt m.(this) m.(sorted)). 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; exact (@Raw.remove_2 elt m.(this) m.(sorted)). Qed.
+ Lemma remove_3 : forall m x y e, MapsTo y e (remove x m) -> MapsTo y e m.
+ Proof. intros m; exact (@Raw.remove_3 elt m.(this) m.(sorted)). Qed.
+
+ Lemma find_1 : forall m x e, MapsTo x e m -> find x m = Some e.
+ Proof. intros m; exact (@Raw.find_1 elt m.(this) m.(sorted)). 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.
+
+ Lemma elements_1 : forall m x e, MapsTo x e m -> InA eq_key_elt (x,e) (elements m).
+ 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, sort lt_key (elements m).
+ Proof. intros m; exact (@Raw.elements_3 elt m.(this) m.(sorted)). Qed.
+
+ Lemma fold_1 : forall m (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.
+ 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.
+ 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'.
+ 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'),
+ 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'),
+ 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)
+ (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)
+ (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')
+ (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.
+ 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')
+ (x:key)(f:option elt->option elt'->option elt''),
+ In x (map2 f m m') -> In x m \/ In x m'.
+ Proof.
+ intros elt elt' elt'' m m' x f;
+ exact (@Raw.map2_2 elt elt' elt'' f m.(this) m.(sorted) m'.(this) m'.(sorted) x).
+ Qed.
+
End Make.
Module Make_ord (X: OrderedType)(D : OrderedType) <:
diff --git a/theories/FSets/FMapPositive.v b/theories/FSets/FMapPositive.v
new file mode 100644
index 00000000..dcb7fb49
--- /dev/null
+++ b/theories/FSets/FMapPositive.v
@@ -0,0 +1,1153 @@
+(***********************************************************************)
+(* 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: FMapPositive.v 8773 2006-04-29 14:31:32Z letouzey $ *)
+
+Require Import ZArith.
+Require Import OrderedType.
+Require Import FMapInterface.
+
+Set Implicit Arguments.
+
+Open Scope positive_scope.
+
+(** * An implementation of [FMapInterface.S] for positive keys. *)
+
+(** This file is an adaptation to the [FMap] framework of a work by
+ Xavier Leroy and Sandrine Blazy (used for building certified compilers).
+ Keys are of type [positive], and maps are binary trees: the sequence
+ of binary digits of a positive number corresponds to a path in such a tree.
+ This is quite similar to the [IntMap] library, except that no path compression
+ is implemented, and that the current file is simple enough to be
+ self-contained. *)
+
+(** Even if [positive] can be seen as an ordered type with respect to the
+ 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.
+ Definition t:=positive.
+ Definition eq:=@eq positive.
+
+ Fixpoint bits_lt (p q:positive) { struct p } : Prop :=
+ match p, q with
+ | xH, xI _ => True
+ | xH, _ => False
+ | xO p, xO q => bits_lt p q
+ | xO _, _ => True
+ | xI p, xI q => bits_lt p q
+ | xI _, _ => False
+ end.
+
+ Definition lt:=bits_lt.
+
+ Lemma 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.
+ induction y; destruct z; simpl; eauto; intuition.
+ induction y; destruct z; simpl; eauto; intuition.
+ induction y; destruct z; simpl; eauto; intuition.
+ Qed.
+
+ Lemma lt_trans : forall x y z : t, lt x y -> lt y z -> lt x z.
+ Proof.
+ exact bits_lt_trans.
+ Qed.
+
+ Lemma bits_lt_antirefl : forall x : positive, ~ bits_lt x x.
+ Proof.
+ induction x; simpl; auto.
+ Qed.
+
+ Lemma lt_not_eq : forall x y : t, lt x y -> ~ eq x y.
+ Proof.
+ intros; intro.
+ rewrite <- H0 in H; clear H0 y.
+ unfold lt in H.
+ exact (bits_lt_antirefl x H).
+ Qed.
+
+ Definition compare : forall x y : t, Compare lt eq x y.
+ Proof.
+ induction x; destruct y.
+ (* I I *)
+ destruct (IHx y).
+ apply LT; auto.
+ apply EQ; rewrite e; red; auto.
+ apply GT; auto.
+ (* I O *)
+ apply GT; simpl; auto.
+ (* I H *)
+ apply GT; simpl; auto.
+ (* O I *)
+ apply LT; simpl; auto.
+ (* O O *)
+ destruct (IHx y).
+ apply LT; auto.
+ apply EQ; rewrite e; red; auto.
+ apply GT; auto.
+ (* O H *)
+ apply LT; simpl; auto.
+ (* H I *)
+ apply LT; simpl; auto.
+ (* H O *)
+ apply GT; simpl; auto.
+ (* H H *)
+ apply EQ; red; auto.
+ Qed.
+
+End PositiveOrderedTypeBits.
+
+(** Other positive stuff *)
+
+Lemma peq_dec (x y: positive): {x = y} + {x <> y}.
+Proof.
+ intros. case_eq ((x ?= y) Eq); intros.
+ left. apply Pcompare_Eq_eq; auto.
+ right. red. intro. subst y. rewrite (Pcompare_refl x) in H. discriminate.
+ right. red. intro. subst y. rewrite (Pcompare_refl x) in H. discriminate.
+Qed.
+
+Fixpoint append (i j : positive) {struct i} : positive :=
+ match i with
+ | xH => j
+ | xI ii => xI (append ii j)
+ | xO ii => xO (append ii j)
+ end.
+
+Lemma append_assoc_0 :
+ forall (i j : positive), append i (xO j) = append (append i (xO xH)) j.
+Proof.
+ induction i; intros; destruct j; simpl;
+ try rewrite (IHi (xI j));
+ try rewrite (IHi (xO j));
+ try rewrite <- (IHi xH);
+ auto.
+Qed.
+
+Lemma append_assoc_1 :
+ forall (i j : positive), append i (xI j) = append (append i (xI xH)) j.
+Proof.
+ induction i; intros; destruct j; simpl;
+ try rewrite (IHi (xI j));
+ try rewrite (IHi (xO j));
+ try rewrite <- (IHi xH);
+ auto.
+Qed.
+
+Lemma append_neutral_r : forall (i : positive), append i xH = i.
+Proof.
+ induction i; simpl; congruence.
+Qed.
+
+Lemma append_neutral_l : forall (i : positive), append xH i = i.
+Proof.
+ simpl; auto.
+Qed.
+
+
+(** The module of maps over positive keys *)
+
+Module PositiveMap <: S with Module E:=PositiveOrderedTypeBits.
+
+ Module E:=PositiveOrderedTypeBits.
+
+ Definition key := positive.
+
+ Inductive tree (A : Set) : Set :=
+ | Leaf : tree A
+ | Node : tree A -> option A -> tree A -> tree A.
+
+ Definition t := tree.
+
+ Section A.
+ Variable A:Set.
+
+ Implicit Arguments Leaf [A].
+
+ Definition empty : t A := Leaf.
+
+ Fixpoint is_empty (m : t A) {struct m} : bool :=
+ match m with
+ | Leaf => true
+ | Node l None r => (is_empty l) && (is_empty r)
+ | _ => false
+ end.
+
+ Fixpoint find (i : positive) (m : t A) {struct i} : option A :=
+ match m with
+ | Leaf => None
+ | Node l o r =>
+ match i with
+ | xH => o
+ | xO ii => find ii l
+ | xI ii => find ii r
+ end
+ end.
+
+ Fixpoint mem (i : positive) (m : t A) {struct i} : bool :=
+ match m with
+ | Leaf => false
+ | Node l o r =>
+ match i with
+ | xH => match o with None => false | _ => true end
+ | xO ii => mem ii l
+ | xI ii => mem ii r
+ end
+ end.
+
+ Fixpoint add (i : positive) (v : A) (m : t A) {struct i} : t A :=
+ match m with
+ | Leaf =>
+ match i with
+ | xH => Node Leaf (Some v) Leaf
+ | xO ii => Node (add ii v Leaf) None Leaf
+ | xI ii => Node Leaf None (add ii v Leaf)
+ end
+ | Node l o r =>
+ match i with
+ | xH => Node l (Some v) r
+ | xO ii => Node (add ii v l) o r
+ | xI ii => Node l o (add ii v r)
+ end
+ end.
+
+ Fixpoint remove (i : positive) (m : t A) {struct i} : t A :=
+ match i with
+ | xH =>
+ match m with
+ | Leaf => Leaf
+ | Node Leaf o Leaf => Leaf
+ | Node l o r => Node l None r
+ end
+ | xO ii =>
+ match m with
+ | Leaf => Leaf
+ | Node l None Leaf =>
+ match remove ii l with
+ | Leaf => Leaf
+ | mm => Node mm None Leaf
+ end
+ | Node l o r => Node (remove ii l) o r
+ end
+ | xI ii =>
+ match m with
+ | Leaf => Leaf
+ | Node Leaf None r =>
+ match remove ii r with
+ | Leaf => Leaf
+ | mm => Node Leaf None mm
+ end
+ | Node l o r => Node l o (remove ii r)
+ end
+ end.
+
+ (** [elements] *)
+
+ Fixpoint xelements (m : t A) (i : positive) {struct m}
+ : list (positive * A) :=
+ match m with
+ | Leaf => nil
+ | Node l None r =>
+ (xelements l (append i (xO xH))) ++ (xelements r (append i (xI xH)))
+ | Node l (Some x) r =>
+ (xelements l (append i (xO xH)))
+ ++ ((i, x) :: xelements r (append i (xI xH)))
+ end.
+
+ (* Note: function [xelements] above is inefficient. We should apply
+ deforestation to it, but that makes the proofs even harder. *)
+
+ Definition elements (m : t A) := xelements m xH.
+
+ Section CompcertSpec.
+
+ Theorem gempty:
+ forall (i: positive), find i empty = None.
+ Proof.
+ destruct i; simpl; auto.
+ Qed.
+
+ Theorem gss:
+ forall (i: positive) (x: A) (m: t A), find i (add i x m) = Some x.
+ Proof.
+ induction i; destruct m; simpl; auto.
+ Qed.
+
+ Lemma gleaf : forall (i : positive), find i (Leaf : t A) = None.
+ Proof. exact gempty. Qed.
+
+ Theorem gso:
+ forall (i j: positive) (x: A) (m: t A),
+ i <> j -> find i (add j x m) = find i m.
+ Proof.
+ induction i; intros; destruct j; destruct m; simpl;
+ try rewrite <- (gleaf i); auto; try apply IHi; congruence.
+ Qed.
+
+ Lemma rleaf : forall (i : positive), remove i (Leaf : t A) = Leaf.
+ Proof. destruct i; simpl; auto. Qed.
+
+ Theorem grs:
+ forall (i: positive) (m: t A), find i (remove i m) = None.
+ Proof.
+ induction i; destruct m.
+ simpl; auto.
+ destruct m1; destruct o; destruct m2 as [ | ll oo rr]; simpl; auto.
+ rewrite (rleaf i); auto.
+ cut (find i (remove i (Node ll oo rr)) = None).
+ destruct (remove i (Node ll oo rr)); auto; apply IHi.
+ apply IHi.
+ simpl; auto.
+ destruct m1 as [ | ll oo rr]; destruct o; destruct m2; simpl; auto.
+ rewrite (rleaf i); auto.
+ cut (find i (remove i (Node ll oo rr)) = None).
+ destruct (remove i (Node ll oo rr)); auto; apply IHi.
+ apply IHi.
+ simpl; auto.
+ destruct m1; destruct m2; simpl; auto.
+ Qed.
+
+ Theorem gro:
+ forall (i j: positive) (m: t A),
+ i <> j -> find i (remove j m) = find i m.
+ Proof.
+ induction i; intros; destruct j; destruct m;
+ try rewrite (rleaf (xI j));
+ try rewrite (rleaf (xO j));
+ try rewrite (rleaf 1); auto;
+ destruct m1; destruct o; destruct m2;
+ simpl;
+ try apply IHi; try congruence;
+ try rewrite (rleaf j); auto;
+ try rewrite (gleaf i); auto.
+ cut (find i (remove j (Node m2_1 o m2_2)) = find i (Node m2_1 o m2_2));
+ [ destruct (remove j (Node m2_1 o m2_2)); try rewrite (gleaf i); auto
+ | apply IHi; congruence ].
+ destruct (remove j (Node m1_1 o0 m1_2)); simpl; try rewrite (gleaf i);
+ auto.
+ destruct (remove j (Node m2_1 o m2_2)); simpl; try rewrite (gleaf i);
+ auto.
+ cut (find i (remove j (Node m1_1 o0 m1_2)) = find i (Node m1_1 o0 m1_2));
+ [ destruct (remove j (Node m1_1 o0 m1_2)); try rewrite (gleaf i); auto
+ | apply IHi; congruence ].
+ destruct (remove j (Node m2_1 o m2_2)); simpl; try rewrite (gleaf i);
+ auto.
+ destruct (remove j (Node m1_1 o0 m1_2)); simpl; try rewrite (gleaf i);
+ auto.
+ Qed.
+
+ Lemma xelements_correct:
+ forall (m: t A) (i j : positive) (v: A),
+ find i m = Some v -> List.In (append j i, v) (xelements m j).
+ Proof.
+ induction m; intros.
+ rewrite (gleaf i) in H; congruence.
+ destruct o; destruct i; simpl; simpl in H.
+ rewrite append_assoc_1; apply in_or_app; right; apply in_cons;
+ apply IHm2; auto.
+ rewrite append_assoc_0; apply in_or_app; left; apply IHm1; auto.
+ rewrite append_neutral_r; apply in_or_app; injection H;
+ intro EQ; rewrite EQ; right; apply in_eq.
+ rewrite append_assoc_1; apply in_or_app; right; apply IHm2; auto.
+ rewrite append_assoc_0; apply in_or_app; left; apply IHm1; auto.
+ congruence.
+ Qed.
+
+ Theorem elements_correct:
+ forall (m: t A) (i: positive) (v: A),
+ find i m = Some v -> List.In (i, v) (elements m).
+ Proof.
+ intros m i v H.
+ exact (xelements_correct m i xH H).
+ Qed.
+
+ Fixpoint xfind (i j : positive) (m : t A) {struct j} : option A :=
+ match i, j with
+ | _, xH => find i m
+ | xO ii, xO jj => xfind ii jj m
+ | xI ii, xI jj => xfind ii jj m
+ | _, _ => None
+ end.
+
+ Lemma xfind_left :
+ forall (j i : positive) (m1 m2 : t A) (o : option A) (v : A),
+ xfind i (append j (xO xH)) m1 = Some v -> xfind i j (Node m1 o m2) = Some v.
+ Proof.
+ induction j; intros; destruct i; simpl; simpl in H; auto; try congruence.
+ destruct i; congruence.
+ Qed.
+
+ Lemma xelements_ii :
+ forall (m: t A) (i j : positive) (v: A),
+ List.In (xI i, v) (xelements m (xI j)) -> List.In (i, v) (xelements m j).
+ Proof.
+ induction m.
+ simpl; auto.
+ intros; destruct o; simpl; simpl in H; destruct (in_app_or _ _ _ H);
+ apply in_or_app.
+ left; apply IHm1; auto.
+ right; destruct (in_inv H0).
+ injection H1; intros Eq1 Eq2; rewrite Eq1; rewrite Eq2; apply in_eq.
+ apply in_cons; apply IHm2; auto.
+ left; apply IHm1; auto.
+ right; apply IHm2; auto.
+ Qed.
+
+ Lemma xelements_io :
+ forall (m: t A) (i j : positive) (v: A),
+ ~List.In (xI i, v) (xelements m (xO j)).
+ Proof.
+ induction m.
+ simpl; auto.
+ intros; destruct o; simpl; intro H; destruct (in_app_or _ _ _ H).
+ apply (IHm1 _ _ _ H0).
+ destruct (in_inv H0).
+ congruence.
+ apply (IHm2 _ _ _ H1).
+ apply (IHm1 _ _ _ H0).
+ apply (IHm2 _ _ _ H0).
+ Qed.
+
+ Lemma xelements_oo :
+ forall (m: t A) (i j : positive) (v: A),
+ List.In (xO i, v) (xelements m (xO j)) -> List.In (i, v) (xelements m j).
+ Proof.
+ induction m.
+ simpl; auto.
+ intros; destruct o; simpl; simpl in H; destruct (in_app_or _ _ _ H);
+ apply in_or_app.
+ left; apply IHm1; auto.
+ right; destruct (in_inv H0).
+ injection H1; intros Eq1 Eq2; rewrite Eq1; rewrite Eq2; apply in_eq.
+ apply in_cons; apply IHm2; auto.
+ left; apply IHm1; auto.
+ right; apply IHm2; auto.
+ Qed.
+
+ Lemma xelements_oi :
+ forall (m: t A) (i j : positive) (v: A),
+ ~List.In (xO i, v) (xelements m (xI j)).
+ Proof.
+ induction m.
+ simpl; auto.
+ intros; destruct o; simpl; intro H; destruct (in_app_or _ _ _ H).
+ apply (IHm1 _ _ _ H0).
+ destruct (in_inv H0).
+ congruence.
+ apply (IHm2 _ _ _ H1).
+ apply (IHm1 _ _ _ H0).
+ apply (IHm2 _ _ _ H0).
+ Qed.
+
+ Lemma xelements_ih :
+ forall (m1 m2: t A) (o: option A) (i : positive) (v: A),
+ List.In (xI i, v) (xelements (Node m1 o m2) xH) -> List.In (i, v) (xelements m2 xH).
+ Proof.
+ destruct o; simpl; intros; destruct (in_app_or _ _ _ H).
+ absurd (List.In (xI i, v) (xelements m1 2)); auto; apply xelements_io; auto.
+ destruct (in_inv H0).
+ congruence.
+ apply xelements_ii; auto.
+ absurd (List.In (xI i, v) (xelements m1 2)); auto; apply xelements_io; auto.
+ apply xelements_ii; auto.
+ Qed.
+
+ Lemma xelements_oh :
+ forall (m1 m2: t A) (o: option A) (i : positive) (v: A),
+ List.In (xO i, v) (xelements (Node m1 o m2) xH) -> List.In (i, v) (xelements m1 xH).
+ Proof.
+ destruct o; simpl; intros; destruct (in_app_or _ _ _ H).
+ apply xelements_oo; auto.
+ destruct (in_inv H0).
+ congruence.
+ absurd (List.In (xO i, v) (xelements m2 3)); auto; apply xelements_oi; auto.
+ apply xelements_oo; auto.
+ absurd (List.In (xO i, v) (xelements m2 3)); auto; apply xelements_oi; auto.
+ Qed.
+
+ Lemma xelements_hi :
+ forall (m: t A) (i : positive) (v: A),
+ ~List.In (xH, v) (xelements m (xI i)).
+ Proof.
+ induction m; intros.
+ simpl; auto.
+ destruct o; simpl; intro H; destruct (in_app_or _ _ _ H).
+ generalize H0; apply IHm1; auto.
+ destruct (in_inv H0).
+ congruence.
+ generalize H1; apply IHm2; auto.
+ generalize H0; apply IHm1; auto.
+ generalize H0; apply IHm2; auto.
+ Qed.
+
+ Lemma xelements_ho :
+ forall (m: t A) (i : positive) (v: A),
+ ~List.In (xH, v) (xelements m (xO i)).
+ Proof.
+ induction m; intros.
+ simpl; auto.
+ destruct o; simpl; intro H; destruct (in_app_or _ _ _ H).
+ generalize H0; apply IHm1; auto.
+ destruct (in_inv H0).
+ congruence.
+ generalize H1; apply IHm2; auto.
+ generalize H0; apply IHm1; auto.
+ generalize H0; apply IHm2; auto.
+ Qed.
+
+ Lemma find_xfind_h :
+ forall (m: t A) (i: positive), find i m = xfind i xH m.
+ Proof.
+ destruct i; simpl; auto.
+ Qed.
+
+ Lemma xelements_complete:
+ forall (i j : positive) (m: t A) (v: A),
+ List.In (i, v) (xelements m j) -> xfind i j m = Some v.
+ Proof.
+ induction i; simpl; intros; destruct j; simpl.
+ apply IHi; apply xelements_ii; auto.
+ absurd (List.In (xI i, v) (xelements m (xO j))); auto; apply xelements_io.
+ destruct m.
+ simpl in H; tauto.
+ rewrite find_xfind_h. apply IHi. apply (xelements_ih _ _ _ _ _ H).
+ absurd (List.In (xO i, v) (xelements m (xI j))); auto; apply xelements_oi.
+ apply IHi; apply xelements_oo; auto.
+ destruct m.
+ simpl in H; tauto.
+ rewrite find_xfind_h. apply IHi. apply (xelements_oh _ _ _ _ _ H).
+ absurd (List.In (xH, v) (xelements m (xI j))); auto; apply xelements_hi.
+ absurd (List.In (xH, v) (xelements m (xO j))); auto; apply xelements_ho.
+ destruct m.
+ simpl in H; tauto.
+ destruct o; simpl in H; destruct (in_app_or _ _ _ H).
+ absurd (List.In (xH, v) (xelements m1 (xO xH))); auto; apply xelements_ho.
+ destruct (in_inv H0).
+ congruence.
+ absurd (List.In (xH, v) (xelements m2 (xI xH))); auto; apply xelements_hi.
+ absurd (List.In (xH, v) (xelements m1 (xO xH))); auto; apply xelements_ho.
+ absurd (List.In (xH, v) (xelements m2 (xI xH))); auto; apply xelements_hi.
+ Qed.
+
+ Theorem elements_complete:
+ forall (m: t A) (i: positive) (v: A),
+ List.In (i, v) (elements m) -> find i m = Some v.
+ Proof.
+ intros m i v H.
+ unfold elements in H.
+ rewrite find_xfind_h.
+ exact (xelements_complete i xH m v H).
+ Qed.
+
+ End CompcertSpec.
+
+ Definition MapsTo (i:positive)(v:A)(m:t A) := find i m = Some v.
+
+ Definition In (i:positive)(m:t A) := exists e:A, MapsTo i e m.
+
+ Definition Empty m := forall (a : positive)(e:A) , ~ MapsTo a e m.
+
+ Definition eq_key (p p':positive*A) := E.eq (fst p) (fst p').
+
+ Definition eq_key_elt (p p':positive*A) :=
+ E.eq (fst p) (fst p') /\ (snd p) = (snd p').
+
+ Definition lt_key (p p':positive*A) := E.lt (fst p) (fst p').
+
+ Lemma mem_find :
+ forall m x, mem x m = match find x m with None => false | _ => true end.
+ Proof.
+ induction m; destruct x; simpl; auto.
+ Qed.
+
+ 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.
+
+ Lemma Empty_Node : forall l o r, Empty (Node l o r) <-> o=None /\ Empty l /\ Empty r.
+ Proof.
+ intros l o r.
+ split.
+ rewrite Empty_alt.
+ split.
+ destruct o; auto.
+ generalize (H 1); simpl; auto.
+ split; rewrite Empty_alt; intros.
+ generalize (H (xO a)); auto.
+ generalize (H (xI a)); auto.
+ intros (H,(H0,H1)).
+ subst.
+ rewrite Empty_alt; intros.
+ destruct a; auto.
+ simpl; generalize H1; rewrite Empty_alt; auto.
+ simpl; generalize H0; rewrite Empty_alt; auto.
+ Qed.
+
+ Section FMapSpec.
+
+ Lemma mem_1 : forall m x, In x m -> mem x m = true.
+ Proof.
+ unfold In, MapsTo; intros m x; rewrite mem_find.
+ destruct 1 as (e0,H0); rewrite H0; auto.
+ Qed.
+
+ Lemma mem_2 : forall m x, mem x m = true -> In x m.
+ Proof.
+ unfold In, MapsTo; intros m x; rewrite mem_find.
+ destruct (find x m).
+ exists a; auto.
+ intros; discriminate.
+ Qed.
+
+ 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; apply gempty.
+ Qed.
+
+ Lemma is_empty_1 : Empty m -> is_empty m = true.
+ Proof.
+ induction m; simpl; auto.
+ rewrite Empty_Node.
+ intros (H,(H0,H1)).
+ subst; simpl.
+ rewrite IHt0_1; simpl; auto.
+ Qed.
+
+ Lemma is_empty_2 : is_empty m = true -> Empty m.
+ Proof.
+ induction m; simpl; auto.
+ rewrite Empty_alt.
+ intros _; exact gempty.
+ rewrite Empty_Node.
+ destruct o.
+ intros; discriminate.
+ intro H; destruct (andb_prop _ _ H); intuition.
+ Qed.
+
+ Lemma add_1 : E.eq x y -> MapsTo y e (add x e m).
+ Proof.
+ unfold MapsTo.
+ intro H; rewrite H; clear H.
+ apply gss.
+ Qed.
+
+ Lemma add_2 : ~ E.eq x y -> MapsTo y e m -> MapsTo y e (add x e' m).
+ Proof.
+ unfold MapsTo.
+ intros; rewrite gso; auto.
+ Qed.
+
+ Lemma add_3 : ~ E.eq x y -> MapsTo y e (add x e' m) -> MapsTo y e m.
+ Proof.
+ unfold MapsTo.
+ intro H; rewrite gso; auto.
+ Qed.
+
+ Lemma remove_1 : E.eq x y -> ~ In y (remove x m).
+ Proof.
+ intros; intro.
+ generalize (mem_1 H0).
+ rewrite mem_find.
+ rewrite H.
+ rewrite grs.
+ intros; discriminate.
+ Qed.
+
+ Lemma remove_2 : ~ E.eq x y -> MapsTo y e m -> MapsTo y e (remove x m).
+ Proof.
+ unfold MapsTo.
+ intro H; rewrite gro; auto.
+ Qed.
+
+ Lemma remove_3 : MapsTo y e (remove x m) -> MapsTo y e m.
+ Proof.
+ unfold MapsTo.
+ destruct (peq_dec x y).
+ subst.
+ rewrite grs; intros; discriminate.
+ rewrite gro; auto.
+ Qed.
+
+ Lemma elements_1 :
+ MapsTo x e m -> InA eq_key_elt (x,e) (elements m).
+ Proof.
+ unfold MapsTo.
+ rewrite InA_alt.
+ intro H.
+ exists (x,e).
+ split.
+ red; simpl; unfold E.eq; auto.
+ apply elements_correct; auto.
+ Qed.
+
+ Lemma elements_2 :
+ InA eq_key_elt (x,e) (elements m) -> MapsTo x e m.
+ Proof.
+ unfold MapsTo.
+ rewrite InA_alt.
+ intros ((e0,a),(H,H0)).
+ red in H; simpl in H; unfold E.eq in H; destruct H; subst.
+ apply elements_complete; auto.
+ Qed.
+
+ Lemma xelements_bits_lt_1 : forall p p0 q m v,
+ List.In (p0,v) (xelements m (append p (xO q))) -> E.bits_lt p0 p.
+ Proof.
+ intros.
+ generalize (xelements_complete _ _ _ _ H); clear H; intros.
+ revert H; revert v; revert m; revert q; revert p0.
+ induction p; destruct p0; simpl; intros; eauto; try discriminate.
+ Qed.
+
+ Lemma xelements_bits_lt_2 : forall p p0 q m v,
+ List.In (p0,v) (xelements m (append p (xI q))) -> E.bits_lt p p0.
+ Proof.
+ intros.
+ generalize (xelements_complete _ _ _ _ H); clear H; intros.
+ revert H; revert v; revert m; revert q; revert p0.
+ induction p; destruct p0; simpl; intros; eauto; try discriminate.
+ Qed.
+
+ Lemma xelements_sort : forall p, sort lt_key (xelements m p).
+ Proof.
+ induction m.
+ simpl; auto.
+ destruct o; simpl; intros.
+ (* Some *)
+ apply (SortA_app (eqA:=eq_key_elt)); auto.
+ compute; intuition.
+ constructor; auto.
+ apply In_InfA; intros.
+ destruct y0.
+ red; red; simpl.
+ eapply xelements_bits_lt_2; eauto.
+ intros x0 y0.
+ do 2 rewrite InA_alt.
+ intros (y1,(Hy1,H)) (y2,(Hy2,H0)).
+ destruct y1; destruct x0; compute in Hy1; destruct Hy1; subst.
+ destruct y2; destruct y0; compute in Hy2; destruct Hy2; subst.
+ red; red; simpl.
+ destruct H0.
+ injection H0; clear H0; intros _ H0; subst.
+ eapply xelements_bits_lt_1; eauto.
+ apply E.bits_lt_trans with p.
+ eapply xelements_bits_lt_1; eauto.
+ eapply xelements_bits_lt_2; eauto.
+ (* None *)
+ apply (SortA_app (eqA:=eq_key_elt)); auto.
+ compute; intuition.
+ intros x0 y0.
+ do 2 rewrite InA_alt.
+ intros (y1,(Hy1,H)) (y2,(Hy2,H0)).
+ destruct y1; destruct x0; compute in Hy1; destruct Hy1; subst.
+ destruct y2; destruct y0; compute in Hy2; destruct Hy2; subst.
+ red; red; simpl.
+ apply E.bits_lt_trans with p.
+ eapply xelements_bits_lt_1; eauto.
+ eapply xelements_bits_lt_2; eauto.
+ Qed.
+
+ Lemma elements_3 : sort lt_key (elements m).
+ Proof.
+ unfold elements.
+ apply xelements_sort; auto.
+ Qed.
+
+ End FMapSpec.
+
+ (** [map] and [mapi] *)
+
+ Variable B : Set.
+
+ Fixpoint xmapi (f : positive -> A -> B) (m : t A) (i : positive)
+ {struct m} : t B :=
+ match m with
+ | Leaf => @Leaf B
+ | Node l o r => Node (xmapi f l (append i (xO xH)))
+ (option_map (f i) o)
+ (xmapi f r (append i (xI xH)))
+ end.
+
+ Definition mapi (f : positive -> A -> B) m := xmapi f m xH.
+
+ Definition map (f : A -> B) m := mapi (fun _ => f) m.
+
+ End A.
+
+ Lemma xgmapi:
+ forall (A B: Set) (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.
+ rewrite (append_assoc_1 j i); apply IHi.
+ rewrite (append_assoc_0 j i); apply IHi.
+ rewrite (append_neutral_r j); auto.
+ Qed.
+
+ Theorem gmapi:
+ forall (A B: Set) (f: positive -> A -> B) (i: positive) (m: t A),
+ find i (mapi f m) = option_map (f i) (find i m).
+ Proof.
+ intros.
+ unfold mapi.
+ replace (f i) with (f (append xH i)).
+ apply xgmapi.
+ rewrite append_neutral_l; 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.
+ exists x.
+ split; [red; auto|].
+ apply find_2.
+ generalize (find_1 H); clear H; intros.
+ rewrite gmapi.
+ rewrite H.
+ simpl; auto.
+ 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.
+ apply mem_2.
+ rewrite mem_find.
+ destruct H as (v,H).
+ generalize (find_1 H); clear H; intros.
+ rewrite gmapi in H.
+ destruct (find x m); auto.
+ simpl in *; discriminate.
+ 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.
+ 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'),
+ 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 f : option A -> option B -> option C.
+
+ Implicit Arguments Leaf [A].
+
+ Fixpoint xmap2_l (m : t A) {struct m} : t C :=
+ match m with
+ | Leaf => Leaf
+ | Node l o r => Node (xmap2_l l) (f o None) (xmap2_l r)
+ end.
+
+ Lemma xgmap2_l : forall (i : positive) (m : t A),
+ f None None = None -> find i (xmap2_l m) = f (find i m) None.
+ Proof.
+ induction i; intros; destruct m; simpl; auto.
+ Qed.
+
+ Fixpoint xmap2_r (m : t B) {struct m} : t C :=
+ match m with
+ | Leaf => Leaf
+ | Node l o r => Node (xmap2_r l) (f None o) (xmap2_r r)
+ end.
+
+ Lemma xgmap2_r : forall (i : positive) (m : t B),
+ f None None = None -> find i (xmap2_r m) = f None (find i m).
+ Proof.
+ induction i; intros; destruct m; simpl; auto.
+ Qed.
+
+ Fixpoint _map2 (m1 : t A)(m2 : t B) {struct m1} : t C :=
+ match m1 with
+ | Leaf => xmap2_r m2
+ | Node l1 o1 r1 =>
+ match m2 with
+ | Leaf => xmap2_l m1
+ | Node l2 o2 r2 => Node (_map2 l1 l2) (f o1 o2) (_map2 r1 r2)
+ end
+ end.
+
+ Lemma gmap2: forall (i: positive)(m1:t A)(m2: t B),
+ f None None = None ->
+ find i (_map2 m1 m2) = f (find i m1) (find i m2).
+ Proof.
+ induction i; intros; destruct m1; destruct m2; simpl; auto;
+ try apply xgmap2_r; try apply xgmap2_l; auto.
+ Qed.
+
+ End map2.
+
+ Definition map2 (elt elt' elt'':Set)(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')
+ (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.
+ intros.
+ unfold map2.
+ rewrite gmap2; auto.
+ generalize (@mem_1 _ m x) (@mem_1 _ m' x).
+ do 2 rewrite mem_find.
+ destruct (find x m); simpl; auto.
+ destruct (find x m'); simpl; auto.
+ intros.
+ destruct H; intuition; try discriminate.
+ Qed.
+
+ Lemma 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'.
+ Proof.
+ intros.
+ generalize (mem_1 H); clear H; intros.
+ rewrite mem_find in H.
+ unfold map2 in H.
+ rewrite gmap2 in H; auto.
+ generalize (@mem_2 _ m x) (@mem_2 _ m' x).
+ do 2 rewrite mem_find.
+ destruct (find x m); simpl in *; auto.
+ destruct (find x m'); simpl in *; auto.
+ Qed.
+
+
+ Definition fold (A B : Set) (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),
+ 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 :=
+ match m1, m2 with
+ | Leaf, _ => is_empty m2
+ | _, Leaf => is_empty m1
+ | Node l1 o1 r1, Node l2 o2 r2 =>
+ (match o1, o2 with
+ | None, None => true
+ | Some v1, Some v2 => cmp v1 v2
+ | _, _ => false
+ end)
+ && equal cmp l1 l2 && equal cmp r1 r2
+ end.
+
+ Definition Equal (A:Set)(cmp:A->A->bool)(m m':t A) :=
+ (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 (A:Set)(m m':t A)(cmp:A->A->bool),
+ Equal cmp m m' -> equal cmp m m' = true.
+ Proof.
+ induction m.
+ (* m = Leaf *)
+ destruct 1.
+ simpl.
+ apply is_empty_1.
+ red; red; intros.
+ assert (In a (Leaf A)).
+ rewrite H.
+ exists e; auto.
+ destruct H2; red in H2.
+ destruct a; simpl in *; discriminate.
+ (* m = Node *)
+ destruct m'.
+ (* m' = Leaf *)
+ destruct 1.
+ simpl.
+ destruct o.
+ assert (In xH (Leaf A)).
+ rewrite <- H.
+ exists a; red; auto.
+ destruct H1; red in H1; simpl in H1; discriminate.
+ apply andb_true_intro; split; apply is_empty_1; red; red; intros.
+ assert (In (xO a) (Leaf A)).
+ rewrite <- H.
+ exists e; auto.
+ destruct H2; red in H2; simpl in H2; discriminate.
+ assert (In (xI a) (Leaf A)).
+ rewrite <- H.
+ exists e; auto.
+ destruct H2; red in H2; simpl in H2; discriminate.
+ (* m' = Node *)
+ destruct 1.
+ assert (Equal 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).
+ 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.
+ simpl.
+ destruct o; destruct o0; simpl.
+ repeat (apply andb_true_intro; split); auto.
+ apply (H0 xH); red; auto.
+ generalize (H xH); unfold In, MapsTo; simpl; intuition.
+ destruct H4; try discriminate; eauto.
+ generalize (H xH); unfold In, MapsTo; simpl; intuition.
+ destruct H5; try discriminate; eauto.
+ 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'.
+ Proof.
+ induction m.
+ (* m = Leaf *)
+ simpl.
+ split; intros.
+ split.
+ destruct 1; red in H0; destruct k; discriminate.
+ destruct 1; elim (is_empty_2 H H0).
+ red in H0; destruct k; discriminate.
+ (* m = Node *)
+ destruct m'.
+ (* m' = Leaf *)
+ simpl.
+ destruct o; intros; try discriminate.
+ destruct (andb_prop _ _ H); clear H.
+ split; intros.
+ split; unfold In, MapsTo; destruct 1.
+ destruct k; simpl in *; try discriminate.
+ destruct (is_empty_2 H1 (find_2 _ _ H)).
+ destruct (is_empty_2 H0 (find_2 _ _ H)).
+ destruct k; simpl in *; discriminate.
+ unfold In, MapsTo; destruct k; simpl in *; discriminate.
+ (* m' = Node *)
+ destruct o; destruct o0; simpl; intros; try discriminate.
+ destruct (andb_prop _ _ H); clear H.
+ destruct (andb_prop _ _ H0); clear H0.
+ destruct (IHm1 _ _ H2); clear H2 IHm1.
+ destruct (IHm2 _ _ H1); clear H1 IHm2.
+ split; intros.
+ destruct k; unfold In, MapsTo in *; simpl; auto.
+ split; eauto.
+ destruct k; unfold In, MapsTo in *; simpl in *.
+ eapply H4; eauto.
+ eapply H3; eauto.
+ congruence.
+ destruct (andb_prop _ _ H); clear H.
+ destruct (IHm1 _ _ H0); clear H0 IHm1.
+ destruct (IHm2 _ _ H1); clear H1 IHm2.
+ split; intros.
+ destruct k; unfold In, MapsTo in *; simpl; auto.
+ split; eauto.
+ destruct k; unfold In, MapsTo in *; simpl in *.
+ eapply H3; eauto.
+ eapply H2; eauto.
+ try discriminate.
+ Qed.
+
+End PositiveMap.
+
+(** Here come some additionnal facts about this implementation.
+ Most are facts that cannot be derivable from the general interface. *)
+
+
+Module PositiveMapAdditionalFacts.
+ Import PositiveMap.
+
+ (* Derivable from the Map interface *)
+ Theorem gsspec:
+ forall (A:Set)(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.
+ destruct (peq_dec i j); [ rewrite e; apply gss | apply gso; auto ].
+ Qed.
+
+ (* Not derivable from the Map interface *)
+ Theorem gsident:
+ forall (A:Set)(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.
+ rewrite (IHi m2 v H); congruence.
+ rewrite (IHi m1 v H); congruence.
+ Qed.
+
+ Lemma xmap2_lr :
+ forall (A B : Set)(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.
+ induction m; intros; simpl; auto.
+ rewrite IHm1; auto.
+ rewrite IHm2; auto.
+ rewrite H; auto.
+ Qed.
+
+ Theorem map2_commut:
+ forall (A B: Set) (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.
+ Proof.
+ intros A B f g Eq1.
+ assert (Eq2: forall (i j: option A), g i j = f j i).
+ intros; auto.
+ induction m1; intros; destruct m2; simpl;
+ try rewrite Eq1;
+ repeat rewrite (xmap2_lr f g);
+ repeat rewrite (xmap2_lr g f);
+ auto.
+ rewrite IHm1_1.
+ rewrite IHm1_2.
+ auto.
+ Qed.
+
+End PositiveMapAdditionalFacts.
+
diff --git a/theories/FSets/FMapWeak.v b/theories/FSets/FMapWeak.v
index 90ebeffc..1ad190a4 100644
--- a/theories/FSets/FMapWeak.v
+++ b/theories/FSets/FMapWeak.v
@@ -6,7 +6,10 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: FMapWeak.v 8639 2006-03-16 19:21:55Z letouzey $ *)
+(* $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
new file mode 100644
index 00000000..18f73a3f
--- /dev/null
+++ b/theories/FSets/FMapWeakFacts.v
@@ -0,0 +1,599 @@
+(***********************************************************************)
+(* 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/FMapWeakList.v b/theories/FSets/FMapWeakList.v
index ce3893e0..3a91b868 100644
--- a/theories/FSets/FMapWeakList.v
+++ b/theories/FSets/FMapWeakList.v
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: FMapWeakList.v 8639 2006-03-16 19:21:55Z letouzey $ *)
+(* $Id: FMapWeakList.v 8899 2006-06-06 11:09:43Z jforest $ *)
(** * Finite map library *)
@@ -24,7 +24,7 @@ Arguments Scope list [type_scope].
Module Raw (X:DecidableType).
-Module PX := PairDecidableType X.
+Module PX := KeyDecidableType X.
Import PX.
Definition key := X.t.
@@ -34,7 +34,7 @@ Section Elt.
Variable elt : Set.
-(* now in PairDecidableType:
+(* 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').
@@ -91,7 +91,7 @@ Qed.
(** * [mem] *)
-Fixpoint mem (k : key) (s : t elt) {struct s} : bool :=
+Function mem (k : key) (s : t elt) {struct s} : bool :=
match s with
| nil => false
| (k',_) :: l => if X.eq_dec k k' then true else mem k l
@@ -100,30 +100,30 @@ Fixpoint mem (k : key) (s : t elt) {struct s} : bool :=
Lemma mem_1 : forall m (Hm:NoDupA m) x, In x m -> mem x m = true.
Proof.
intros m Hm x; generalize Hm; clear Hm.
- functional induction mem x m;intros NoDup belong1;trivial.
+ functional induction (mem x m);intros NoDup belong1;trivial.
inversion belong1. inversion H.
inversion_clear NoDup.
inversion_clear belong1.
- inversion_clear H3.
- compute in H4; destruct H4.
- elim H; auto.
- apply H0; auto.
- exists x; auto.
+ inversion_clear H2.
+ compute in H3; destruct H3.
+ contradiction.
+ apply IHb; auto.
+ exists x0; auto.
Qed.
Lemma mem_2 : forall m (Hm:NoDupA m) x, mem x m = true -> In x m.
Proof.
intros m Hm x; generalize Hm; clear Hm; unfold PX.In,PX.MapsTo.
- functional induction mem x m; intros NoDup hyp; try discriminate.
- exists e; auto.
+ functional induction (mem x m); intros NoDup hyp; try discriminate.
+ exists _x; auto.
inversion_clear NoDup.
- destruct H0; auto.
- exists x; auto.
+ destruct IHb; auto.
+ exists x0; auto.
Qed.
(** * [find] *)
-Fixpoint find (k:key) (s: t elt) {struct s} : option elt :=
+Function find (k:key) (s: t elt) {struct s} : option elt :=
match s with
| nil => None
| (k',x)::s' => if X.eq_dec k k' then Some x else find k s'
@@ -132,23 +132,23 @@ Fixpoint find (k:key) (s: t elt) {struct s} : option elt :=
Lemma find_2 : forall m x e, find x m = Some e -> MapsTo x e m.
Proof.
intros m x. unfold PX.MapsTo.
- functional induction find x m;simpl;intros e' eqfind; inversion eqfind; auto.
+ functional induction (find x m);simpl;intros e' eqfind; inversion eqfind; auto.
Qed.
Lemma find_1 : forall m (Hm:NoDupA m) x e,
MapsTo x e m -> find x m = Some e.
Proof.
intros m Hm x e; generalize Hm; clear Hm; unfold PX.MapsTo.
- functional induction find x m;simpl; subst; try clear H_eq_1.
+ functional induction (find x m);simpl; subst; try clear H_eq_1.
inversion 2.
do 2 inversion_clear 1.
compute in H3; destruct H3; subst; trivial.
- elim H0; apply InA_eqk with (k,e); auto.
+ elim H; apply InA_eqk with (x,e); auto.
do 2 inversion_clear 1; auto.
- compute in H4; destruct H4; elim H; auto.
+ compute in H3; destruct H3; elim _x; auto.
Qed.
(* Not part of the exported specifications, used later for [combine]. *)
@@ -166,7 +166,7 @@ Qed.
(** * [add] *)
-Fixpoint add (k : key) (x : elt) (s : t elt) {struct s} : t elt :=
+Function add (k : key) (x : elt) (s : t elt) {struct s} : t elt :=
match s with
| nil => (k,x) :: nil
| (k',y) :: l => if X.eq_dec k k' then (k,x)::l else (k',y)::add k x l
@@ -175,26 +175,26 @@ Fixpoint add (k : key) (x : elt) (s : t elt) {struct s} : t elt :=
Lemma add_1 : forall m x y e, X.eq x y -> MapsTo y e (add x e m).
Proof.
intros m x y e; generalize y; clear y; unfold PX.MapsTo.
- functional induction add x e m;simpl;auto.
+ functional induction (add x e m);simpl;auto.
Qed.
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 m x y e e'; generalize y e; clear y e; unfold PX.MapsTo.
- functional induction add x e' m;simpl;auto.
- intros y' e' eqky'; inversion_clear 1.
+ functional induction (add x e' m);simpl;auto.
+ intros y' e'' eqky'; inversion_clear 1.
destruct H1; simpl in *.
elim eqky'; apply X.eq_trans with k'; auto.
auto.
- intros y' e' eqky'; inversion_clear 1; intuition.
+ intros y' e'' eqky'; inversion_clear 1; intuition.
Qed.
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 m x y e e'. generalize y e; clear y e; unfold PX.MapsTo.
- functional induction add x e' m;simpl;auto.
+ functional induction (add x e' m);simpl;auto.
intros; apply (In_inv_3 H0); auto.
constructor 2; apply (In_inv_3 H1); auto.
inversion_clear 2; auto.
@@ -204,12 +204,12 @@ Lemma add_3' : forall m x y e e',
~ X.eq x y -> InA eqk (y,e) (add x e' m) -> InA eqk (y,e) m.
Proof.
intros m x y e e'. generalize y e; clear y e.
- functional induction add x e' m;simpl;auto.
+ functional induction (add x e' m);simpl;auto.
inversion_clear 2.
compute in H1; elim H; auto.
inversion H1.
constructor 2; inversion_clear H1; auto.
- compute in H2; elim H0; auto.
+ compute in H2; elim H; auto.
inversion_clear 2; auto.
Qed.
@@ -257,7 +257,7 @@ Qed.
(** * [remove] *)
-Fixpoint remove (k : key) (s : t elt) {struct s} : t elt :=
+Function remove (k : key) (s : t elt) {struct s} : t elt :=
match s with
| nil => nil
| (k',x) :: l => if X.eq_dec k k' then l else (k',x) :: remove k l
@@ -266,7 +266,7 @@ Fixpoint remove (k : key) (s : t elt) {struct s} : t elt :=
Lemma remove_1 : forall m (Hm:NoDupA m) x y, X.eq x y -> ~ In y (remove x m).
Proof.
intros m Hm x y; generalize Hm; clear Hm.
- functional induction remove x m;simpl;intros;auto.
+ functional induction (remove x m);simpl;intros;auto.
red; inversion 1; inversion H1.
@@ -275,14 +275,14 @@ Proof.
swap H1.
destruct H3 as (e,H3); unfold PX.MapsTo in H3.
apply InA_eqk with (y,e); auto.
- compute; apply X.eq_trans with k; auto.
+ compute; apply X.eq_trans with x; auto.
intro H2.
destruct H2 as (e,H2); inversion_clear H2.
- compute in H3; destruct H3.
- elim H; apply X.eq_trans with y; auto.
+ compute in H1; destruct H1.
+ elim _x; apply X.eq_trans with y; auto.
inversion_clear Hm.
- elim (H0 H4 H1).
+ elim (IHt0 H3 H).
exists e; auto.
Qed.
@@ -290,10 +290,10 @@ Lemma remove_2 : forall m (Hm:NoDupA m) x y e,
~ X.eq x y -> MapsTo y e m -> MapsTo y e (remove x m).
Proof.
intros m Hm x y e; generalize Hm; clear Hm; unfold PX.MapsTo.
- functional induction remove x m;auto.
+ functional induction (remove x m);auto.
inversion_clear 3; auto.
compute in H2; destruct H2.
- elim H0; apply X.eq_trans with k'; auto.
+ elim H; apply X.eq_trans with k'; auto.
inversion_clear 1; inversion_clear 2; auto.
Qed.
@@ -302,7 +302,7 @@ Lemma remove_3 : forall m (Hm:NoDupA m) x y e,
MapsTo y e (remove x m) -> MapsTo y e m.
Proof.
intros m Hm x y e; generalize Hm; clear Hm; unfold PX.MapsTo.
- functional induction remove x m;auto.
+ functional induction (remove x m);auto.
do 2 inversion_clear 1; auto.
Qed.
@@ -310,7 +310,7 @@ Lemma remove_3' : forall m (Hm:NoDupA m) x y e,
InA eqk (y,e) (remove x m) -> InA eqk (y,e) m.
Proof.
intros m Hm x y e; generalize Hm; clear Hm; unfold PX.MapsTo.
- functional induction remove x m;auto.
+ functional induction (remove x m);auto.
do 2 inversion_clear 1; auto.
Qed.
@@ -347,8 +347,7 @@ Qed.
(** * [fold] *)
-Fixpoint fold (A:Set)(f:key->elt->A->A)(m:t elt) {struct m} : A -> A :=
- fun acc =>
+Function fold (A:Set)(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)
@@ -357,7 +356,7 @@ Fixpoint fold (A:Set)(f:key->elt->A->A)(m:t elt) {struct m} : A -> A :=
Lemma fold_1 : forall m (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.
Proof.
- intros; functional induction fold A f m i; auto.
+ intros; functional induction (@fold A f m i); auto.
Qed.
(** * [equal] *)
@@ -878,83 +877,124 @@ Module Make (X: DecidableType) <: S with Module E:=X.
Module Raw := Raw X.
Module E := X.
- Definition key := X.t.
+ Definition key := E.t.
Record slist (elt:Set) : Set :=
{this :> Raw.t elt; NoDup : NoDupA (@Raw.PX.eqk elt) this}.
Definition t (elt:Set) := slist elt.
- Section Elt.
+Section Elt.
Variable elt elt' elt'':Set.
Implicit Types m : t elt.
-
- Definition empty := Build_slist (Raw.empty_NoDup elt).
- Definition is_empty m := Raw.is_empty m.(this).
- Definition add x e m := Build_slist (Raw.add_NoDup m.(NoDup) x e).
- Definition find x m := Raw.find x m.(this).
- Definition remove x m := Build_slist (Raw.remove_NoDup m.(NoDup) x).
- Definition mem x m := Raw.mem x m.(this).
+ Implicit Types x y : key.
+ Implicit Types e : elt.
+
+ Definition empty : t elt := Build_slist (Raw.empty_NoDup elt).
+ Definition is_empty m : bool := Raw.is_empty m.(this).
+ Definition add x e m : t elt := Build_slist (Raw.add_NoDup m.(NoDup) x e).
+ Definition find x m : option elt := Raw.find x m.(this).
+ Definition remove x m : t elt := Build_slist (Raw.remove_NoDup m.(NoDup) x).
+ Definition mem x m : bool := Raw.mem x m.(this).
Definition map f m : t elt' := Build_slist (Raw.map_NoDup m.(NoDup) f).
- Definition mapi f m : t elt' := Build_slist (Raw.mapi_NoDup m.(NoDup) f).
+ Definition mapi (f:key->elt->elt') m : t elt' := Build_slist (Raw.mapi_NoDup m.(NoDup) f).
Definition map2 f m (m':t elt') : t elt'' :=
- Build_slist (Raw.map2_NoDup f m.(NoDup) m'.(NoDup)).
- Definition elements m := @Raw.elements elt m.(this).
- Definition fold A f m i := @Raw.fold elt A f m.(this) i.
- Definition equal cmp m m' := @Raw.equal elt cmp m.(this) m'.(this).
-
- Definition MapsTo x e m := Raw.PX.MapsTo x e m.(this).
- Definition In x m := Raw.PX.In x m.(this).
- Definition Empty m := Raw.Empty m.(this).
- Definition Equal cmp m m' := @Raw.Equal elt cmp m.(this) m'.(this).
+ 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 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 eq_key : (key*elt) -> (key*elt) -> Prop := @Raw.PX.eqk elt.
+ Definition eq_key_elt : (key*elt) -> (key*elt) -> Prop:= @Raw.PX.eqke 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.PX.MapsTo_eq elt m.(this)). Qed.
+
+ Lemma mem_1 : forall m x, In x m -> mem x m = true.
+ Proof. intros m; exact (@Raw.mem_1 elt m.(this) m.(NoDup)). Qed.
+ Lemma mem_2 : forall m x, mem x m = true -> In x m.
+ Proof. intros m; exact (@Raw.mem_2 elt m.(this) m.(NoDup)). Qed.
+
+ Lemma empty_1 : Empty empty.
+ Proof. exact (@Raw.empty_1 elt). Qed.
+
+ Lemma is_empty_1 : forall m, Empty m -> is_empty m = true.
+ Proof. intros m; exact (@Raw.is_empty_1 elt m.(this)). Qed.
+ Lemma is_empty_2 : forall m, is_empty m = true -> Empty m.
+ Proof. intros m; exact (@Raw.is_empty_2 elt 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; exact (@Raw.add_1 elt m.(this)). 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; exact (@Raw.add_2 elt m.(this)). 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; exact (@Raw.add_3 elt m.(this)). Qed.
+
+ Lemma remove_1 : forall m x y, E.eq x y -> ~ In y (remove x m).
+ Proof. intros m; exact (@Raw.remove_1 elt m.(this) m.(NoDup)). 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; exact (@Raw.remove_2 elt m.(this) m.(NoDup)). Qed.
+ Lemma remove_3 : forall m x y e, MapsTo y e (remove x m) -> MapsTo y e m.
+ Proof. intros m; exact (@Raw.remove_3 elt m.(this) m.(NoDup)). Qed.
+
+ Lemma find_1 : forall m x e, MapsTo x e m -> find x m = Some e.
+ Proof. intros m; exact (@Raw.find_1 elt m.(this) m.(NoDup)). 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.
+
+ Lemma elements_1 : forall m x e, MapsTo x e m -> InA eq_key_elt (x,e) (elements m).
+ 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 fold_1 : forall m (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.
+ 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.
+ 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'.
+ Proof. intros m m'; exact (@Raw.equal_2 elt m.(this) m.(NoDup) m'.(this) m'.(NoDup)). Qed.
- Definition eq_key (p p':key*elt) := X.eq (fst p) (fst p').
+ End Elt.
- Definition eq_key_elt (p p':key*elt) :=
- X.eq (fst p) (fst p') /\ (snd p) = (snd p').
-
- Definition MapsTo_1 m := @Raw.PX.MapsTo_eq elt m.(this).
-
- Definition mem_1 m := @Raw.mem_1 elt m.(this) m.(NoDup).
- Definition mem_2 m := @Raw.mem_2 elt m.(this) m.(NoDup).
-
- Definition empty_1 := @Raw.empty_1.
-
- Definition is_empty_1 m := @Raw.is_empty_1 elt m.(this).
- Definition is_empty_2 m := @Raw.is_empty_2 elt m.(this).
-
- Definition add_1 m := @Raw.add_1 elt m.(this).
- Definition add_2 m := @Raw.add_2 elt m.(this).
- Definition add_3 m := @Raw.add_3 elt m.(this).
+ 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. 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'),
+ 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)
+ (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)
+ (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')
+ (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.
+ 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')
+ (x:key)(f:option elt->option elt'->option elt''),
+ In x (map2 f m m') -> In x m \/ In x m'.
+ Proof.
+ intros elt elt' elt'' m m' x f;
+ exact (@Raw.map2_2 elt elt' elt'' f m.(this) m.(NoDup) m'.(this) m'.(NoDup) x).
+ Qed.
- Definition remove_1 m := @Raw.remove_1 elt m.(this) m.(NoDup).
- Definition remove_2 m := @Raw.remove_2 elt m.(this) m.(NoDup).
- Definition remove_3 m := @Raw.remove_3 elt m.(this) m.(NoDup).
-
- Definition find_1 m := @Raw.find_1 elt m.(this) m.(NoDup).
- Definition find_2 m := @Raw.find_2 elt m.(this).
-
- Definition elements_1 m := @Raw.elements_1 elt m.(this).
- Definition elements_2 m := @Raw.elements_2 elt m.(this).
- Definition elements_3 m := @Raw.elements_3 elt m.(this) m.(NoDup).
-
- Definition fold_1 m := @Raw.fold_1 elt m.(this).
-
- Definition map_1 m := @Raw.map_1 elt elt' m.(this).
- Definition map_2 m := @Raw.map_2 elt elt' m.(this).
-
- Definition mapi_1 m := @Raw.mapi_1 elt elt' m.(this).
- Definition mapi_2 m := @Raw.mapi_2 elt elt' m.(this).
-
- Definition map2_1 m (m':t elt') x f :=
- @Raw.map2_1 elt elt' elt'' f m.(this) m.(NoDup) m'.(this) m'.(NoDup) x.
- Definition map2_2 m (m':t elt') x f :=
- @Raw.map2_2 elt elt' elt'' f m.(this) m.(NoDup) m'.(this) m'.(NoDup) x.
-
- Definition equal_1 m m' := @Raw.equal_1 elt m.(this) m.(NoDup) m'.(this) m'.(NoDup).
- Definition equal_2 m m' := @Raw.equal_2 elt m.(this) m.(NoDup) m'.(this) m'.(NoDup).
-
- End Elt.
End Make.
-
diff --git a/theories/FSets/FMaps.v b/theories/FSets/FMaps.v
index ae5b86c9..72ccad3f 100644
--- a/theories/FSets/FMaps.v
+++ b/theories/FSets/FMaps.v
@@ -6,7 +6,13 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: FMaps.v 8667 2006-03-28 11:59:44Z letouzey $ *)
+(* $Id: FMaps.v 8844 2006-05-22 17:22:36Z letouzey $ *)
+Require Export OrderedType.
+Require Export OrderedTypeEx.
+Require Export OrderedTypeAlt.
Require Export FMapInterface.
Require Export FMapList.
+Require Export FMapPositive.
+Require Export FMapIntMap.
+Require Export FMapFacts. \ No newline at end of file
diff --git a/theories/FSets/FSetAVL.v b/theories/FSets/FSetAVL.v
new file mode 100644
index 00000000..b385f50e
--- /dev/null
+++ b/theories/FSets/FSetAVL.v
@@ -0,0 +1,2900 @@
+
+(***********************************************************************)
+(* 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: FSetAVL.v 8899 2006-06-06 11:09:43Z jforest $ *)
+
+(** This module implements sets using AVL trees.
+ It follows the implementation from Ocaml's standard library. *)
+
+Require Import FSetInterface.
+Require Import FSetList.
+Require Import ZArith.
+Require Import Int.
+
+Set Firstorder Depth 3.
+
+Module Raw (I:Int)(X:OrderedType).
+Import I.
+Module II:=MoreInt(I).
+Import II.
+Open Scope Int_scope.
+
+Module E := X.
+Module MX := OrderedTypeFacts X.
+
+Definition elt := X.t.
+
+(** * Trees *)
+
+Inductive tree : Set :=
+ | Leaf : tree
+ | Node : tree -> X.t -> tree -> int -> tree.
+
+Notation t := tree.
+
+(** The fourth field of [Node] is the height of the tree *)
+
+(** 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
+ | H:f _ Leaf |- _ => inversion_clear H; inv f
+ | H:f (Node _ _ _ _) |- _ => inversion_clear H; inv f
+ | H:f _ (Node _ _ _ _) |- _ => inversion_clear H; inv f
+ | _ => idtac
+ end.
+
+(** Same, but with a backup of the original hypothesis. *)
+
+Ltac safe_inv f := match goal with
+ | H:f (Node _ _ _ _) |- _ =>
+ generalize H; inversion_clear H; safe_inv f
+ | _ => intros
+ end.
+
+(** * Occurrence in a tree *)
+
+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).
+
+(** [In] is compatible with [X.eq] *)
+
+Lemma In_1 :
+ forall s x y, X.eq x y -> In x s -> In y s.
+Proof.
+ induction s; simpl; intuition_in; eauto.
+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.
+
+(** 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.
+Qed.
+
+Lemma gt_leaf : forall x : elt, gt_tree x Leaf.
+Proof.
+ unfold gt_tree in |- *; intros; inversion H.
+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.
+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.
+Qed.
+
+Hint Resolve lt_leaf gt_leaf lt_tree_node gt_tree_node.
+
+Lemma lt_tree_not_in :
+ forall (x : elt) (t : tree), lt_tree x t -> ~ In x t.
+Proof.
+ intros; intro; order.
+Qed.
+
+Lemma lt_tree_trans :
+ forall x y, X.lt x y -> forall t, lt_tree x t -> lt_tree y t.
+Proof.
+ firstorder eauto.
+Qed.
+
+Lemma gt_tree_not_in :
+ forall (x : elt) (t : tree), gt_tree x t -> ~ In x t.
+Proof.
+ intros; intro; order.
+Qed.
+
+Lemma gt_tree_trans :
+ forall x y, X.lt y x -> forall t, gt_tree x t -> gt_tree y t.
+Proof.
+ firstorder 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).
+
+Hint Constructors bst.
+
+(** * 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 *)
+
+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.
+
+(** * 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.
+ intro H; elim (H x); auto.
+Qed.
+
+Lemma is_empty_2 : forall s, is_empty s = true -> Empty s.
+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.
+
+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.
+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.
+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.
+
+Lemma singleton_1 : forall x y, In y (singleton x) -> X.eq x y.
+Proof.
+ unfold singleton; inversion_clear 1; auto; inversion_clear H0.
+Qed.
+
+Lemma singleton_2 : forall x y, X.eq x y -> In y (singleton x).
+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).
+Proof.
+ unfold create; 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.
+
+Lemma create_in :
+ forall l x r y, In y (create l x r) <-> X.eq y x \/ In y l \/ In y r.
+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).
+Proof.
+ bal_tac; inv avl; repeat apply create_avl; simpl in *; auto; 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.
+ bal_tac; 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.
+ bal_tac; inv avl; simpl in *; omega_max.
+Qed.
+
+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).
+Proof.
+ intros s x; functional induction (add x s); auto; intros.
+ intuition_in.
+ (* LT *)
+ inv avl.
+ rewrite bal_in; auto.
+ rewrite (IHt y0 H1); intuition_in.
+ (* EQ *)
+ inv avl.
+ intuition.
+ eapply In_1; eauto.
+ (* GT *)
+ inv avl.
+ rewrite bal_in; auto.
+ rewrite (IHt y0 H2); intuition_in.
+Qed.
+
+Lemma add_bst : forall s x, bst s -> avl s -> bst (add x s).
+Proof.
+ intros s x; functional induction (add x s); auto; intros.
+ inv bst; inv avl; apply bal_bst; auto.
+ (* lt_tree -> lt_tree (add ...) *)
+ red; red in H5.
+ intros.
+ rewrite (add_in l x y0 H) in H1.
+ intuition.
+ eauto.
+ inv bst; inv avl; apply bal_bst; auto.
+ (* gt_tree -> gt_tree (add ...) *)
+ red; red in H5.
+ intros.
+ rewrite (add_in r x y0 H6) in H1.
+ intuition.
+ apply MX.lt_eq with x; auto.
+Qed.
+
+(** * 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.
+
+Ltac join_tac :=
+ intro l; induction l as [| ll _ lx lr Hlr lh];
+ [ | intros x r; induction r as [| rl Hrl rx rr _ rh]; unfold join;
+ [ | destruct (gt_le_dec lh (rh+2));
+ [ match goal with |- context b [ bal ?a ?b ?c] =>
+ replace (bal a b c)
+ with (bal ll lx (join lr x (Node rl rx rr rh))); [ | auto]
+ end
+ | destruct (gt_le_dec rh (lh+2));
+ [ match goal with |- context b [ bal ?a ?b ?c] =>
+ replace (bal a b c)
+ with (bal (join (Node ll lx lr lh) x rl) rx rr); [ | auto]
+ end
+ | ] ] ] ]; intros.
+
+Lemma join_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).
+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.
+
+ apply create_in.
+Qed.
+
+Lemma join_bst : forall l x r, bst l -> avl l -> bst r -> avl 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.
+Qed.
+
+(** * 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 H0 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.
+
+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))).
+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 H1).
+ rewrite H0; simpl; intros.
+ rewrite bal_in; auto.
+ rewrite H0 in IHp;generalize (IHp lh y H1).
+ intuition.
+ inversion_clear H8; 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)).
+Proof.
+ intros l x r; functional induction (remove_min l x r); subst;simpl in *; intros.
+ inv bst; auto.
+ inversion_clear H; inversion_clear H1.
+ rewrite_all H0;simpl in *.
+ apply bal_bst; auto.
+ firstorder.
+ intro; intros.
+ generalize (remove_min_in ll lx lr lh y H).
+ rewrite H0; simpl.
+ destruct 1.
+ apply H4; 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)).
+Proof.
+ intros l x r; functional induction (remove_min l x r); subst;simpl in *; intros.
+ inv bst; auto.
+ inversion_clear H; inversion_clear H1.
+ intro; intro.
+ generalize (IHp lh H2 H); clear H8 H7 IHp.
+ generalize (remove_min_avl ll lx lr lh H).
+ generalize (remove_min_in ll lx lr lh m H).
+ rewrite H0; simpl; intros.
+ rewrite (bal_in l' x r y H8 H6) in H1.
+ destruct H7.
+ firstorder.
+ apply MX.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 (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 H1.
+ generalize (remove_min_avl_1 l2 x2 r2 h2 H0).
+ rewrite H2; 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.
+
+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.
+ intuition_in.
+ intuition_in.
+ destruct s1;try contradiction;clear H1.
+ replace s2' with (fst (remove_min l2 x2 r2)); [|rewrite H2; auto].
+ rewrite bal_in; auto.
+ generalize (remove_min_avl l2 x2 r2 h2); rewrite H2; simpl; auto.
+ generalize (remove_min_in l2 x2 r2 h2 y); rewrite H2; simpl; intro.
+ rewrite H1; intuition.
+Qed.
+
+Lemma merge_bst : forall s1 s2, bst s1 -> avl s1 -> bst s2 -> avl s2 ->
+ (forall y1 y2 : elt, In y1 s1 -> In y2 s2 -> X.lt y1 y2) ->
+ bst (merge s1 s2).
+Proof.
+ intros s1 s2; functional induction (merge s1 s2); subst;simpl in *; intros; auto.
+ destruct s1;try contradiction;clear H1.
+ apply bal_bst; auto.
+ generalize (remove_min_bst l2 x2 r2 h2); rewrite H2; simpl in *; auto.
+ intro; intro.
+ apply H5; auto.
+ generalize (remove_min_in l2 x2 r2 h2 m); rewrite H2; simpl; intuition.
+ generalize (remove_min_gt_tree l2 x2 r2 h2); rewrite H2; simpl; auto.
+Qed.
+
+(** * 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 H1).
+ split.
+ apply bal_avl; auto.
+ omega_max.
+ omega_bal.
+ (* EQ *)
+ inv avl.
+ generalize (merge_avl_1 l r H1 H2 H3).
+ intuition omega_max.
+ (* GT *)
+ inv avl.
+ destruct (IHt H2).
+ 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.
+
+Lemma remove_in : forall s x y, bst s -> avl 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.
+ intuition_in.
+ (* LT *)
+ inv avl; inv bst; clear H0.
+ rewrite bal_in; auto.
+ generalize (IHt y0 H1); intuition; [ order | order | intuition_in ].
+ (* EQ *)
+ inv avl; inv bst; clear H0.
+ rewrite merge_in; intuition; [ order | order | intuition_in ].
+ elim H9; eauto.
+ (* GT *)
+ inv avl; inv bst; clear H0.
+ rewrite bal_in; auto.
+ generalize (IHt y0 H6); intuition; [ order | order | intuition_in ].
+Qed.
+
+Lemma remove_bst : forall s x, bst s -> avl s -> bst (remove x s).
+Proof.
+ intros s x; functional induction (remove x s); simpl; intros.
+ auto.
+ (* LT *)
+ inv avl; inv bst.
+ apply bal_bst; auto.
+ intro; intro.
+ rewrite (remove_in l x y0) in H; auto.
+ destruct H; eauto.
+ (* EQ *)
+ inv avl; inv bst.
+ apply merge_bst; eauto.
+ (* GT *)
+ inv avl; inv bst.
+ apply bal_bst; auto.
+ intro; intro.
+ rewrite (remove_in r x y0) in H; auto.
+ destruct H; eauto.
+Qed.
+
+ (** * 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.
+
+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.
+Qed.
+
+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.
+ 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 H3; auto).
+ inversion_clear 1; auto; order.
+ assert (X.lt t _x) by auto.
+ inversion_clear 2; auto;
+ (assert (~ X.lt t 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.
+ inversion 1.
+ destruct l;try contradiction.
+ clear H0;intro H0.
+ destruct (IHo H0 t); 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.
+
+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.
+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.
+ 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.
+ inversion_clear 2; auto;
+ (assert (~ X.lt x t) 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.
+ red; auto.
+ inversion 1.
+ destruct r;try contradiction.
+ clear H0;intros H0; destruct (IHo H0 t); auto.
+Qed.
+
+(** * Any element *)
+
+Definition choose := min_elt.
+
+Lemma choose_1 : forall s x, choose s = Some x -> In x s.
+Proof.
+ exact min_elt_1.
+Qed.
+
+Lemma choose_2 : forall s, choose s = None -> Empty s.
+Proof.
+ exact min_elt_3.
+Qed.
+
+(** * Concatenation
+
+ 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.
+
+Lemma concat_avl : forall s1 s2, avl s1 -> avl s2 -> avl (concat s1 s2).
+Proof.
+ intros s1 s2; functional induction (concat s1 s2); subst;auto.
+ destruct s1;try contradiction;clear H1.
+ intros; apply join_avl; auto.
+ generalize (remove_min_avl l2 x2 r2 h2 H0); rewrite H2; simpl; auto.
+Qed.
+
+Lemma concat_bst : forall s1 s2, bst s1 -> avl s1 -> bst s2 -> avl 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 H1.
+ intros; apply join_bst; auto.
+ generalize (remove_min_bst l2 x2 r2 h2 H1 H3); rewrite H2; simpl; auto.
+ generalize (remove_min_avl l2 x2 r2 h2 H3); rewrite H2; simpl; auto.
+ generalize (remove_min_in l2 x2 r2 h2 m H3); rewrite H2; simpl; auto.
+ destruct 1; intuition.
+ generalize (remove_min_gt_tree l2 x2 r2 h2 H1 H3); rewrite H2; 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 H1;intuition.
+ inversion_clear H5.
+ destruct s1;try contradiction;clear H1; intros.
+ rewrite (join_in (Node s1_1 t s1_2 i) m s2' y H0).
+ generalize (remove_min_avl l2 x2 r2 h2 H3); rewrite H2; simpl; auto.
+ generalize (remove_min_in l2 x2 r2 h2 y H3); rewrite H2; simpl.
+ intro EQ; rewrite EQ; intuition.
+Qed.
+
+(** * 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].
+*)
+
+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 H1 in IHp;simpl in IHp;inversion_clear 1; intuition.
+ simpl; inversion_clear 1; auto.
+ rewrite H1 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 H1 in IHp;simpl in *; inversion_clear 1; inversion_clear 1; clear H8 H9.
+ rewrite (IHp y0 H2 H6); clear IHp H0.
+ intuition.
+ inversion_clear H0; auto; order.
+ (* EQ *)
+ simpl in *; inversion_clear 1; inversion_clear 1; clear H8 H7 H0.
+ intuition.
+ order.
+ intuition_in; order.
+ (* GT *)
+ rewrite H1 in IHp;simpl in *; inversion_clear 1; inversion_clear 1; clear H9 H8.
+ rewrite join_in; auto.
+ generalize (split_avl r x H7); rewrite H1; simpl; intuition.
+ rewrite (IHp y0 H3 H7); clear H1.
+ intuition; [ eauto | eauto | intuition_in ].
+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).
+Proof.
+ intros s x; functional induction (split x s);subst;simpl in *.
+ intuition; try inversion_clear H1.
+ (* LT *)
+ rewrite H1 in IHp; simpl in *; inversion_clear 1; inversion_clear 1; clear H9 H8.
+ rewrite join_in; auto.
+ generalize (split_avl l x H6); rewrite H1; simpl; intuition.
+ rewrite (IHp y0 H2 H6); clear IHp H0.
+ intuition; [ order | order | intuition_in ].
+ (* EQ *)
+ simpl in *; inversion_clear 1; inversion_clear 1; clear H8 H7 H0.
+ intuition; [ order | intuition_in; order ].
+ (* GT *)
+ rewrite H1 in IHp; simpl in *; inversion_clear 1; inversion_clear 1; clear H9 H8.
+ rewrite (IHp y0 H3 H7); clear IHp H0.
+ intuition; intuition_in; order.
+Qed.
+
+Lemma split_in_3 : forall s x, bst s -> avl s ->
+ (fst (snd (split x s)) = true <-> In x s).
+Proof.
+ intros s x; functional induction (split x s);subst;simpl in *.
+ intuition; try inversion_clear H1.
+ (* LT *)
+ rewrite H1 in IHp; simpl in *; inversion_clear 1; inversion_clear 1; clear H9 H8.
+ rewrite IHp; auto.
+ intuition_in; absurd (X.lt x y); eauto.
+ (* EQ *)
+ simpl in *; inversion_clear 1; inversion_clear 1; intuition.
+ (* GT *)
+ rewrite H1 in IHp; simpl in *; inversion_clear 1; inversion_clear 1; clear H9 H8.
+ rewrite IHp; auto.
+ intuition_in; absurd (X.lt y x); eauto.
+Qed.
+
+Lemma split_bst : forall s x, bst s -> avl s ->
+ bst (fst (split x s)) /\ bst (snd (snd (split x s))).
+Proof.
+ intros s x; functional induction (split x s);subst;simpl in *.
+ intuition.
+ (* LT *)
+ rewrite H1 in IHp; simpl in *; inversion_clear 1; inversion_clear 1.
+ intuition.
+ apply join_bst; auto.
+ generalize (split_avl l x H6); rewrite H1; simpl; intuition.
+ intro; intro.
+ generalize (split_in_2 l x y0 H2 H6); rewrite H1; simpl; intuition.
+ (* EQ *)
+ simpl in *; inversion_clear 1; inversion_clear 1; intuition.
+ (* GT *)
+ rewrite H1 in IHp; simpl in *; inversion_clear 1; inversion_clear 1.
+ intuition.
+ apply join_bst; auto.
+ generalize (split_avl r x H7); rewrite H1; simpl; intuition.
+ intro; intro.
+ generalize (split_in_1 r x y0 H3 H7); rewrite H1; 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.
+
+Lemma inter_bst_in : forall s1 s2, bst s1 -> avl s1 -> bst s2 -> avl 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.
+ (* 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 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.
+ intuition_in.
+ generalize (H26 (In_1 _ _ _ H22 H35)); intro; discriminate.
+Qed.
+
+Lemma inter_bst : forall s1 s2, bst s1 -> avl s1 -> bst s2 -> avl s2 ->
+ bst (inter s1 s2).
+Proof.
+ intros; generalize (inter_bst_in s1 s2); intuition.
+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).
+Proof.
+ intros; generalize (inter_bst_in s1 s2); firstorder.
+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.
+
+Lemma diff_bst_in : forall s1 s2, bst s1 -> avl s1 -> bst s2 -> avl 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.
+ (* bst concat *)
+ apply concat_bst; try apply diff_avl; auto.
+ intros; generalize (H22 y1) (H24 y2); intuition eauto.
+ (* 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.
+ 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.
+Qed.
+
+Lemma diff_bst : forall s1 s2, bst s1 -> avl s1 -> bst s2 -> avl s2 ->
+ bst (diff s1 s2).
+Proof.
+ intros; generalize (diff_bst_in s1 s2); intuition.
+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).
+Proof.
+ intros; generalize (diff_bst_in s1 s2); firstorder.
+Qed.
+
+(** * 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) {struct t} : 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.
+
+Lemma elements_aux_in : forall s acc x,
+ InA X.eq x (elements_aux acc s) <-> In x s \/ InA X.eq x acc.
+Proof.
+ induction s as [ | l Hl x r Hr h ]; simpl; auto.
+ intuition.
+ inversion H0.
+ intros.
+ rewrite Hl.
+ destruct (Hr acc x0); clear Hl Hr.
+ intuition; inversion_clear H3; intuition.
+Qed.
+
+Lemma elements_in : forall s x, InA X.eq x (elements s) <-> In x s.
+Proof.
+ intros; generalize (elements_aux_in s nil x); intuition.
+ inversion_clear H0.
+Qed.
+
+Lemma elements_aux_sort : forall s acc, bst s -> sort X.lt acc ->
+ (forall x y : elt, InA X.eq x acc -> In y s -> X.lt y x) ->
+ sort X.lt (elements_aux acc s).
+Proof.
+ induction s as [ | l Hl y r Hr h]; simpl; intuition.
+ inv bst.
+ apply Hl; auto.
+ constructor.
+ apply Hr; auto.
+ apply MX.In_Inf; intros.
+ destruct (elements_aux_in r acc y0); intuition.
+ intros.
+ inversion_clear H.
+ order.
+ destruct (elements_aux_in r acc x); intuition eauto.
+Qed.
+
+Lemma elements_sort : forall s : tree, bst s -> sort X.lt (elements s).
+Proof.
+ intros; unfold elements; apply elements_aux_sort; auto.
+ intros; inversion H0.
+Qed.
+Hint Resolve elements_sort.
+
+(** * Filter *)
+
+Section F.
+Variable f : elt -> bool.
+
+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.
+
+Definition filter := filter_acc Leaf.
+
+Lemma filter_acc_avl : forall s acc, avl s -> avl acc ->
+ avl (filter_acc acc s).
+Proof.
+ induction s; simpl; auto.
+ intros.
+ inv avl.
+ apply IHs2; auto.
+ apply IHs1; auto.
+ destruct (f t); auto.
+Qed.
+Hint Resolve filter_acc_avl.
+
+Lemma filter_acc_bst : forall s acc, bst s -> avl s -> bst acc -> avl acc ->
+ bst (filter_acc acc s).
+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.
+
+Lemma filter_acc_in : forall s acc, avl s -> avl acc ->
+ compat_bool X.eq f -> forall x : elt,
+ In x (filter_acc 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.
+ case_eq (f t); intros.
+ rewrite (add_in); auto.
+ intuition_in.
+ rewrite (H1 _ _ H8).
+ 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.
+Qed.
+
+Lemma filter_bst : forall s, bst s -> avl s -> bst (filter s).
+Proof.
+ unfold filter; intros; apply filter_acc_bst; auto.
+Qed.
+
+Lemma filter_in : forall s, avl s ->
+ compat_bool X.eq f -> forall x : elt,
+ In x (filter 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).
+
+Lemma partition_acc_avl_1 : forall s acc, avl s ->
+ avl (fst acc) -> avl (fst (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.
+
+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.
+
+Lemma partition_acc_in_1 : forall s acc, avl s -> avl (fst 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.
+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.
+ case_eq (f t); simpl; intros.
+ rewrite (add_in); auto.
+ intuition_in.
+ rewrite (H1 _ _ H8).
+ intuition.
+ intuition_in.
+ rewrite (H1 _ _ H8) in H9.
+ rewrite H in H9; discriminate.
+Qed.
+
+Lemma partition_acc_in_2 : forall s acc, avl s -> avl (snd 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.
+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.
+ case_eq (f t); simpl; intros.
+ intuition.
+ intuition_in.
+ rewrite (H1 _ _ H8) in H9.
+ rewrite H in H9; discriminate.
+ rewrite (add_in); auto.
+ intuition_in.
+ rewrite (H1 _ _ H8).
+ intuition.
+Qed.
+
+Lemma partition_avl_1 : forall s, avl s -> avl (fst (partition s)).
+Proof.
+ unfold partition; intros; apply partition_acc_avl_1; auto.
+Qed.
+
+Lemma partition_avl_2 : forall s, avl s -> avl (snd (partition s)).
+Proof.
+ unfold partition; intros; apply partition_acc_avl_2; auto.
+Qed.
+
+Lemma partition_bst_1 : forall s, bst s -> avl s ->
+ bst (fst (partition s)).
+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)).
+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.
+
+Lemma for_all_1 : forall s, compat_bool E.eq f ->
+ For_all (fun x => f x = true) s -> for_all s = true.
+Proof.
+ induction s; simpl; auto.
+ intros.
+ rewrite IHs1; try red; auto.
+ rewrite IHs2; try red; auto.
+ generalize (H0 t).
+ 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.
+Proof.
+ induction s; simpl; auto; intros; red; intros; inv In.
+ destruct (andb_prop _ _ H0); auto.
+ destruct (andb_prop _ _ H1); eauto.
+ apply IHs1; auto.
+ destruct (andb_prop _ _ H0); auto.
+ destruct (andb_prop _ _ H1); auto.
+ apply IHs2; auto.
+ destruct (andb_prop _ _ H0); auto.
+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.
+Proof.
+ induction s; simpl; destruct 2 as (x,(U,V)); inv In.
+ 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.
+Qed.
+
+Lemma exists_2 : forall s, compat_bool E.eq f ->
+ exists_ s = true -> Exists (fun x => f x = true) s.
+Proof.
+ induction s; simpl; intros.
+ 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.
+
+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].
+
+Definition fold' (A : Set) (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),
+ L.fold f (elements_aux acc s) a = L.fold f acc (fold f s a).
+Proof.
+ simple induction s.
+ simpl in |- *; intuition.
+ simpl in |- *; intros.
+ rewrite H.
+ simpl.
+ apply H0.
+Qed.
+
+Lemma fold_equiv :
+ forall (A : Set) (s : tree) (f : elt -> A -> A) (a : A),
+ fold f s a = fold' f s a.
+Proof.
+ unfold fold', elements in |- *.
+ simple induction s; simpl in |- *; auto; intros.
+ rewrite fold_equiv_aux.
+ rewrite H0.
+ simpl in |- *; auto.
+Qed.
+
+Lemma fold_1 :
+ forall (s:t)(Hs:bst s)(A : Set)(f : elt -> A -> A)(i : A),
+ fold f s i = fold_left (fun a e => f e a) (elements s) i.
+Proof.
+ intros.
+ rewrite fold_equiv.
+ unfold fold'.
+ rewrite L.fold_1.
+ unfold L.elements; auto.
+ 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.
+
+Lemma cardinal_elements_aux_1 :
+ forall s acc, (length acc + cardinal s)%nat = length (elements_aux acc s).
+Proof.
+ simple induction s; simpl in |- *; intuition.
+ rewrite <- H.
+ simpl in |- *.
+ rewrite <- H0; omega.
+Qed.
+
+Lemma cardinal_elements_1 : forall s : tree, cardinal s = length (elements s).
+Proof.
+ exact (fun s => cardinal_elements_aux_1 s nil).
+Qed.
+
+(** NB: the remaining functions (union, subset, compare) are still defined
+ in a dependent style, due to the use of well-founded induction. *)
+
+(** 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.
+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.
+ 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.
+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.
+Proof.
+ simpl in |- *; intuition.
+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.
+
+Lemma eq_refl : forall s : t, eq s s.
+Proof.
+ unfold eq, Equal in |- *; intuition.
+Qed.
+
+Lemma eq_sym : forall s s' : t, eq s s' -> eq s' s.
+Proof.
+ unfold eq, Equal in |- *; firstorder.
+Qed.
+
+Lemma eq_trans : forall s s' s'' : t, eq s s' -> eq s' s'' -> eq s s''.
+Proof.
+ unfold eq, Equal in |- *; firstorder.
+Qed.
+
+Lemma eq_L_eq :
+ forall s s' : t, eq 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.
+Qed.
+
+Lemma L_eq_eq :
+ forall s s' : t, L.eq (elements s) (elements s') -> eq s s'.
+Proof.
+ unfold eq, Equal, L.eq, L.Equal in |- *; intros.
+ generalize (elements_in s a) (elements_in s' a).
+ firstorder.
+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'.
+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)
+*)
+
+(** ** Enumeration of the elements of a tree *)
+
+Inductive enumeration : Set :=
+ | End : enumeration
+ | More : elt -> tree -> enumeration -> enumeration.
+
+(** [flatten_e e] returns the list of elements of [e] i.e. the list
+ of elements actually compared *)
+
+Fixpoint flatten_e (e : enumeration) : list elt := match e with
+ | End => nil
+ | More x t r => x :: elements t ++ flatten_e r
+ end.
+
+(** [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.
+Proof.
+ simpl in |- *; unfold elements in |- *; simpl in |- *; intuition.
+ repeat rewrite elements_app.
+ repeat rewrite <- app_nil_end.
+ repeat rewrite app_ass; auto.
+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.
+Proof.
+ intros; simpl.
+ apply compare_flatten_1.
+Qed.
+
+(** termination of [compare_aux] *)
+
+Open 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
+ 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.
+Proof.
+ simple induction s.
+ simpl in |- *; omega.
+ intros.
+ Measure_e_t; omega. (* BUG Simpl! *)
+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.
+Proof.
+ simple induction e.
+ simpl in |- *; omega.
+ intros.
+ Measure_e; Measure_e_t_0 t; omega.
+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'.
+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.
+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).
+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.
+
+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.
+Qed.
+
+(** * Equality test *)
+
+Definition equal : forall s s' : t, bst s -> bst s' -> {Equal s s'} + {~ Equal s s'}.
+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.
+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').
+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.
+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.
+
+(** Alternative equal based on subset *)
+
+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. *)
+
+Module IntMake (I:Int)(X: OrderedType) <: S with Module E := X.
+
+ Module E := X.
+ Module Raw := Raw I X.
+
+ Record bbst : Set := Bbst {this :> Raw.t; is_bst : Raw.bst this; is_avl: Raw.avl this}.
+ Definition t := bbst.
+ 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.
+
+ 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.
+
+ Definition mem (x:elt)(s:t) : bool := Raw.mem x s.
+
+ Definition empty : t := Bbst _ Raw.empty_bst Raw.empty_avl.
+ 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 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 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)).
+ 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'.
+
+ 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 eq (s s':t) : Prop := Raw.eq s s'.
+ Definition lt (s s':t) : Prop := Raw.lt 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.
+ 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 (Raw.mem_1 s x (is_bst s)). Qed.
+ Lemma mem_2 : mem x s = true -> In x s.
+ Proof. exact (Raw.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.
+
+ Lemma equal_2 : equal s s' = true -> Equal s s'.
+ Proof.
+ unfold equal; destruct (Raw.equal s s'); simpl; intuition; discriminate.
+ 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.
+
+ Lemma subset_1 : Subset s s' -> subset s s' = true.
+ Proof.
+ unfold subset; destruct (Raw.subset s s'); simpl; intuition.
+ 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.
+
+ Lemma empty_1 : Empty empty.
+ Proof. exact Raw.empty_1. Qed.
+
+ Lemma is_empty_1 : Empty s -> is_empty s = true.
+ Proof. exact (Raw.is_empty_1 s). Qed.
+ Lemma is_empty_2 : is_empty s = true -> Empty s.
+ Proof. exact (Raw.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.
+
+ Lemma add_2 : In y s -> In y (add x s).
+ Proof.
+ unfold add, In; simpl; rewrite Raw.add_in; auto.
+ 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.
+
+ Lemma remove_1 : E.eq x y -> ~ In y (remove x s).
+ Proof.
+ unfold remove, In; simpl; rewrite Raw.remove_in; intuition.
+ 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.
+
+ Lemma remove_3 : In y (remove x s) -> In y s.
+ Proof.
+ unfold remove, In; simpl; rewrite Raw.remove_in; intuition.
+ Qed.
+
+ Lemma singleton_1 : In y (singleton x) -> E.eq x y.
+ Proof. exact (Raw.singleton_1 x y). Qed.
+ Lemma singleton_2 : E.eq x y -> In y (singleton x).
+ Proof. exact (Raw.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.
+
+ 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.
+
+ 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.
+
+ Lemma inter_1 : In x (inter s s') -> In x s.
+ Proof.
+ unfold inter, In; simpl; rewrite Raw.inter_in; intuition.
+ Qed.
+
+ Lemma inter_2 : In x (inter s s') -> In x s'.
+ Proof.
+ unfold inter, In; simpl; rewrite Raw.inter_in; intuition.
+ 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.
+
+ Lemma diff_1 : In x (diff s s') -> In x s.
+ Proof.
+ unfold diff, In; simpl; rewrite Raw.diff_in; intuition.
+ Qed.
+
+ Lemma diff_2 : In x (diff s s') -> ~ In x s'.
+ Proof.
+ unfold diff, In; simpl; rewrite Raw.diff_in; intuition.
+ 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.
+
+ 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 cardinal_1 : cardinal s = length (elements s).
+ Proof.
+ unfold cardinal, elements; intros; apply Raw.cardinal_elements_1; 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.
+
+ 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.
+
+ 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.
+
+ 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.
+ 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.
+
+ 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.
+ 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.
+
+ 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.
+ 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.
+ destruct (f a); 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.
+
+ Lemma elements_2 : InA E.eq x (elements s) -> In x s.
+ Proof.
+ unfold elements, In; rewrite Raw.elements_in; auto.
+ Qed.
+
+ Lemma elements_3 : sort E.lt (elements s).
+ Proof. exact (Raw.elements_sort _ (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.
+ 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.
+ Lemma min_elt_3 : min_elt s = None -> Empty s.
+ Proof. exact (Raw.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.
+ 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.
+ Lemma max_elt_3 : max_elt s = None -> Empty s.
+ Proof. exact (Raw.max_elt_3 s). Qed.
+
+ Lemma choose_1 : choose s = Some x -> In x s.
+ Proof. exact (Raw.choose_1 s x). Qed.
+ Lemma choose_2 : choose s = None -> Empty s.
+ Proof. exact (Raw.choose_2 s). Qed.
+
+ Lemma eq_refl : eq s s.
+ Proof. exact (Raw.eq_refl s). Qed.
+ Lemma eq_sym : eq s s' -> eq s' s.
+ Proof. exact (Raw.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.
+
+ Lemma lt_trans : lt s s' -> lt s' s'' -> lt s s''.
+ Proof. exact (Raw.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.
+
+ 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/FSetBridge.v b/theories/FSets/FSetBridge.v
index 3ea50df8..08985cfc 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 8639 2006-03-16 19:21:55Z letouzey $ *)
+(* $Id: FSetBridge.v 8834 2006-05-20 00:41:35Z letouzey $ *)
(** * Finite sets library *)
@@ -109,7 +109,7 @@ Module DepOfNodep (M: S) <: Sdep with Module E := M.E.
Definition elements :
forall s : t,
- {l : list elt | ME.Sort l /\ (forall x : elt, In x s <-> ME.In x l)}.
+ {l : list elt | sort E.lt l /\ (forall x : elt, In x s <-> InA E.eq x l)}.
Proof.
intros; exists (elements s); intuition.
Defined.
@@ -394,17 +394,17 @@ Module NodepOfDep (M: Sdep) <: S with Module E := M.E.
Definition elements (s : t) : list elt := let (l, _) := elements s in l.
- Lemma elements_1 : forall (s : t) (x : elt), In x s -> ME.In x (elements s).
+ Lemma elements_1 : forall (s : t) (x : elt), In x s -> InA E.eq x (elements s).
Proof.
intros; unfold elements in |- *; case (M.elements s); firstorder.
Qed.
- Lemma elements_2 : forall (s : t) (x : elt), ME.In x (elements s) -> In x s.
+ Lemma elements_2 : forall (s : t) (x : elt), InA E.eq x (elements s) -> In x s.
Proof.
intros s x; unfold elements in |- *; case (M.elements s); firstorder.
Qed.
- Lemma elements_3 : forall s : t, ME.Sort (elements s).
+ Lemma elements_3 : forall s : t, sort E.lt (elements s).
Proof.
intros; unfold elements in |- *; case (M.elements s); firstorder.
Qed.
diff --git a/theories/FSets/FSetEqProperties.v b/theories/FSets/FSetEqProperties.v
index 006d78c7..d7062d5a 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 8639 2006-03-16 19:21:55Z letouzey $ *)
+(* $Id: FSetEqProperties.v 8853 2006-05-23 18:17:38Z herbelin $ *)
(** * Finite sets library *)
@@ -276,7 +276,7 @@ Lemma is_empty_cardinal: is_empty s = zerob (cardinal s).
Proof.
intros; apply bool_1; split; intros.
rewrite cardinal_1; simpl; auto.
-assert (cardinal s = 0) by apply zerob_true_elim; auto.
+assert (cardinal s = 0) by (apply zerob_true_elim; auto).
auto.
Qed.
@@ -672,7 +672,7 @@ unfold Add, MP.Add; intros.
repeat rewrite filter_iff; auto.
rewrite H0; clear H0.
assert (E.eq x y -> f y = true) by
- intro H0; rewrite <- (Comp _ _ H0); auto.
+ (intro H0; rewrite <- (Comp _ _ H0); auto).
tauto.
Qed.
@@ -704,6 +704,11 @@ assert (f a || g a = true <-> f a = true \/ g a = true).
tauto.
Qed.
+Lemma filter_union: forall s s', filter f (union s s') [=] union (filter f s) (filter f s').
+Proof.
+unfold Equal; intros; set_iff; repeat rewrite filter_iff; auto; set_iff; tauto.
+Qed.
+
(** Properties of [for_all] *)
Lemma for_all_mem_1: forall s,
diff --git a/theories/FSets/FSetFacts.v b/theories/FSets/FSetFacts.v
index d8c0b802..aa57f066 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 8681 2006-04-05 11:56:14Z letouzey $ *)
+(* $Id: FSetFacts.v 8882 2006-05-31 21:55:30Z letouzey $ *)
(** * Finite sets library *)
@@ -131,7 +131,7 @@ Proof.
split; [apply exists_1 | apply exists_2]; auto.
Qed.
-Lemma elements_iff : In x s <-> ME.In x (elements s).
+Lemma elements_iff : In x s <-> InA E.eq x (elements s).
Proof.
split; [apply elements_1 | apply elements_2].
Qed.
@@ -159,6 +159,12 @@ 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.
diff --git a/theories/FSets/FSetInterface.v b/theories/FSets/FSetInterface.v
index c177abfe..64ad234b 100644
--- a/theories/FSets/FSetInterface.v
+++ b/theories/FSets/FSetInterface.v
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: FSetInterface.v 8671 2006-03-29 08:31:28Z letouzey $ *)
+(* $Id: FSetInterface.v 8820 2006-05-15 11:44:05Z letouzey $ *)
(** * Finite set library *)
@@ -153,7 +153,7 @@ Module Type S.
Section Spec.
Variable s s' s'' : t.
- Variable x y z : elt.
+ Variable x y : elt.
(** Specification of [In] *)
Parameter In_1 : E.eq x y -> In x s -> In y s.
@@ -254,6 +254,8 @@ Module Type S.
Parameter partition_2 : compat_bool E.eq f ->
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.
@@ -275,7 +277,6 @@ Module Type S.
(* Parameter choose_equal:
(equal s s')=true -> E.eq (choose s) (choose s'). *)
- End Filter.
End Spec.
(* begin hide *)
diff --git a/theories/FSets/FSetList.v b/theories/FSets/FSetList.v
index ca86ffcc..f6205542 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 8667 2006-03-28 11:59:44Z letouzey $ *)
+(* $Id: FSetList.v 8834 2006-05-20 00:41:35Z letouzey $ *)
(** * Finite sets library *)
@@ -199,6 +199,8 @@ Module Raw (X: OrderedType).
(** ** Proofs of set operation specifications. *)
+ Section ForNotations.
+
Notation Sort := (sort X.lt).
Notation Inf := (lelistA X.lt).
Notation In := (InA X.eq).
@@ -1020,6 +1022,9 @@ Module Raw (X: OrderedType).
destruct (e1 a0); auto.
Defined.
+ End ForNotations.
+ Hint Constructors lt.
+
End Raw.
(** * Encapsulation
@@ -1029,135 +1034,213 @@ End Raw.
Module Make (X: OrderedType) <: S with Module E := X.
- Module E := X.
Module Raw := Raw X.
+ Module E := X.
- Record slist : Set := {this :> Raw.t; sorted : sort X.lt this}.
+ Record slist : Set := {this :> Raw.t; sorted : sort E.lt this}.
Definition t := slist.
- Definition elt := X.t.
+ Definition elt := E.t.
- Definition In (x : elt) (s : t) := InA X.eq x s.(this).
- 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.
-
- Definition In_1 (s : t) := Raw.MX.In_eq (l:=s.(this)).
-
- Definition mem (x : elt) (s : t) := Raw.mem x s.
- Definition mem_1 (s : t) := Raw.mem_1 (sorted s).
- Definition mem_2 (s : t) := Raw.mem_2 (s:=s).
-
- Definition add x s := Build_slist (Raw.add_sort (sorted s) x).
- Definition add_1 (s : t) := Raw.add_1 (sorted s).
- Definition add_2 (s : t) := Raw.add_2 (sorted s).
- Definition add_3 (s : t) := Raw.add_3 (sorted s).
-
- Definition remove x s := Build_slist (Raw.remove_sort (sorted s) x).
- Definition remove_1 (s : t) := Raw.remove_1 (sorted s).
- Definition remove_2 (s : t) := Raw.remove_2 (sorted s).
- Definition remove_3 (s : t) := Raw.remove_3 (sorted s).
-
- Definition singleton x := Build_slist (Raw.singleton_sort x).
- Definition singleton_1 := Raw.singleton_1.
- Definition singleton_2 := Raw.singleton_2.
-
- Definition union (s s' : t) :=
+ Definition In (x : elt) (s : t) : Prop := InA E.eq x s.(this).
+ 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 mem (x : elt) (s : t) : bool := Raw.mem x s.
+ Definition add (x : elt)(s : t) : t := Build_slist (Raw.add_sort (sorted s) x).
+ Definition remove (x : elt)(s : t) : t := Build_slist (Raw.remove_sort (sorted s) x).
+ Definition singleton (x : elt) : t := Build_slist (Raw.singleton_sort x).
+ Definition union (s s' : t) : t :=
Build_slist (Raw.union_sort (sorted s) (sorted s')).
- Definition union_1 (s s' : t) := Raw.union_1 (sorted s) (sorted s').
- Definition union_2 (s s' : t) := Raw.union_2 (sorted s) (sorted s').
- Definition union_3 (s s' : t) := Raw.union_3 (sorted s) (sorted s').
-
- Definition inter (s s' : t) :=
+ Definition inter (s s' : t) : t :=
Build_slist (Raw.inter_sort (sorted s) (sorted s')).
- Definition inter_1 (s s' : t) := Raw.inter_1 (sorted s) (sorted s').
- Definition inter_2 (s s' : t) := Raw.inter_2 (sorted s) (sorted s').
- Definition inter_3 (s s' : t) := Raw.inter_3 (sorted s) (sorted s').
-
- Definition diff (s s' : t) :=
+ Definition diff (s s' : t) : t :=
Build_slist (Raw.diff_sort (sorted s) (sorted s')).
- Definition diff_1 (s s' : t) := Raw.diff_1 (sorted s) (sorted s').
- Definition diff_2 (s s' : t) := Raw.diff_2 (sorted s) (sorted s').
- Definition diff_3 (s s' : t) := Raw.diff_3 (sorted s) (sorted s').
-
- Definition equal (s s' : t) := Raw.equal s s'.
- Definition equal_1 (s s' : t) := Raw.equal_1 (sorted s) (sorted s').
- Definition equal_2 (s s' : t) := Raw.equal_2 (s:=s) (s':=s').
-
- Definition subset (s s' : t) := Raw.subset s s'.
- Definition subset_1 (s s' : t) := Raw.subset_1 (sorted s) (sorted s').
- Definition subset_2 (s s' : t) := Raw.subset_2 (s:=s) (s':=s').
+ Definition equal (s s' : t) : bool := Raw.equal s s'.
+ Definition subset (s s' : t) : bool := Raw.subset s s'.
+ Definition empty : t := Build_slist Raw.empty_sort.
+ Definition is_empty (s : t) : bool := Raw.is_empty 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 (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).
+ 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
+ (Build_slist (this:=fst p) (Raw.partition_sort_1 (sorted s) f),
+ Build_slist (this:=snd p) (Raw.partition_sort_2 (sorted s) f)).
+ Definition eq (s s' : t) : Prop := Raw.eq s s'.
+ Definition lt (s s' : t) : Prop := Raw.lt s s'.
- Definition empty := Build_slist Raw.empty_sort.
- Definition empty_1 := Raw.empty_1.
-
- Definition is_empty (s : t) := Raw.is_empty s.
- Definition is_empty_1 (s : t) := Raw.is_empty_1 (s:=s).
- Definition is_empty_2 (s : t) := Raw.is_empty_2 (s:=s).
-
- Definition elements (s : t) := Raw.elements s.
- Definition elements_1 (s : t) := Raw.elements_1 (s:=s).
- Definition elements_2 (s : t) := Raw.elements_2 (s:=s).
- Definition elements_3 (s : t) := Raw.elements_3 (sorted s).
-
- Definition min_elt (s : t) := Raw.min_elt s.
- Definition min_elt_1 (s : t) := Raw.min_elt_1 (s:=s).
- Definition min_elt_2 (s : t) := Raw.min_elt_2 (sorted s).
- Definition min_elt_3 (s : t) := Raw.min_elt_3 (s:=s).
-
- Definition max_elt (s : t) := Raw.max_elt s.
- Definition max_elt_1 (s : t) := Raw.max_elt_1 (s:=s).
- Definition max_elt_2 (s : t) := Raw.max_elt_2 (sorted s).
- Definition max_elt_3 (s : t) := Raw.max_elt_3 (s:=s).
-
- Definition choose := min_elt.
- Definition choose_1 := min_elt_1.
- Definition choose_2 := min_elt_3.
+ Section Spec.
+ Variable s s' s'': t.
+ Variable x y : elt.
+
+ Lemma In_1 : E.eq x y -> In x s -> In y s.
+ Proof. exact (fun H H' => Raw.MX.In_eq H H'). Qed.
- Definition fold (B : Set) (f : elt -> B -> B) (s : t) := Raw.fold (B:=B) f s.
- Definition fold_1 (s : t) := Raw.fold_1 (sorted s).
+ Lemma mem_1 : In x s -> mem x s = true.
+ Proof. exact (fun H => Raw.mem_1 s.(sorted) H). Qed.
+ Lemma mem_2 : mem x s = true -> In x s.
+ Proof. exact (fun H => Raw.mem_2 H). Qed.
- Definition cardinal (s : t) := Raw.cardinal s.
- Definition cardinal_1 (s : t) := Raw.cardinal_1 (sorted s).
+ Lemma equal_1 : Equal s s' -> equal s s' = true.
+ Proof. exact (Raw.equal_1 s.(sorted) s'.(sorted)). Qed.
+ Lemma equal_2 : equal s s' = true -> Equal s s'.
+ Proof. exact (fun H => Raw.equal_2 H). Qed.
+
+ Lemma subset_1 : Subset s s' -> subset s s' = true.
+ Proof. exact (Raw.subset_1 s.(sorted) s'.(sorted)). Qed.
+ Lemma subset_2 : subset s s' = true -> Subset s s'.
+ Proof. exact (fun H => Raw.subset_2 H). Qed.
+
+ Lemma empty_1 : Empty empty.
+ Proof. exact Raw.empty_1. Qed.
+
+ Lemma is_empty_1 : Empty s -> is_empty s = true.
+ Proof. exact (fun H => Raw.is_empty_1 H). Qed.
+ Lemma is_empty_2 : is_empty s = true -> Empty s.
+ Proof. exact (fun H => Raw.is_empty_2 H). Qed.
- Definition filter (f : elt -> bool) (s : t) :=
- Build_slist (Raw.filter_sort (sorted s) f).
- Definition filter_1 (s : t) := Raw.filter_1 (s:=s).
- Definition filter_2 (s : t) := Raw.filter_2 (s:=s).
- Definition filter_3 (s : t) := Raw.filter_3 (s:=s).
+ Lemma add_1 : E.eq x y -> In y (add x s).
+ Proof. exact (fun H => Raw.add_1 s.(sorted) H). Qed.
+ Lemma add_2 : In y s -> In y (add x s).
+ Proof. exact (fun H => Raw.add_2 s.(sorted) x H). Qed.
+ Lemma add_3 : ~ E.eq x y -> In y (add x s) -> In y s.
+ Proof. exact (fun H => Raw.add_3 s.(sorted) H). Qed.
+
+ Lemma remove_1 : E.eq x y -> ~ In y (remove x s).
+ Proof. exact (fun H => Raw.remove_1 s.(sorted) H). Qed.
+ Lemma remove_2 : ~ E.eq x y -> In y s -> In y (remove x s).
+ Proof. exact (fun H H' => Raw.remove_2 s.(sorted) H H'). Qed.
+ Lemma remove_3 : In y (remove x s) -> In y s.
+ Proof. exact (fun H => Raw.remove_3 s.(sorted) H). Qed.
+
+ Lemma singleton_1 : In y (singleton x) -> E.eq x y.
+ Proof. exact (fun H => Raw.singleton_1 H). Qed.
+ Lemma singleton_2 : E.eq x y -> In y (singleton x).
+ Proof. exact (fun H => Raw.singleton_2 H). Qed.
+
+ Lemma union_1 : In x (union s s') -> In x s \/ In x s'.
+ Proof. exact (fun H => Raw.union_1 s.(sorted) s'.(sorted) H). Qed.
+ Lemma union_2 : In x s -> In x (union s s').
+ Proof. exact (fun H => Raw.union_2 s.(sorted) s'.(sorted) H). Qed.
+ Lemma union_3 : In x s' -> In x (union s s').
+ Proof. exact (fun H => Raw.union_3 s.(sorted) s'.(sorted) H). Qed.
+
+ Lemma inter_1 : In x (inter s s') -> In x s.
+ Proof. exact (fun H => Raw.inter_1 s.(sorted) s'.(sorted) H). Qed.
+ Lemma inter_2 : In x (inter s s') -> In x s'.
+ Proof. exact (fun H => Raw.inter_2 s.(sorted) s'.(sorted) H). Qed.
+ Lemma inter_3 : In x s -> In x s' -> In x (inter s s').
+ Proof. exact (fun H => Raw.inter_3 s.(sorted) s'.(sorted) H). Qed.
+
+ Lemma diff_1 : In x (diff s s') -> In x s.
+ Proof. exact (fun H => Raw.diff_1 s.(sorted) s'.(sorted) H). Qed.
+ Lemma diff_2 : In x (diff s s') -> ~ In x s'.
+ Proof. exact (fun H => Raw.diff_2 s.(sorted) s'.(sorted) H). Qed.
+ 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.
- Definition for_all (f : elt -> bool) (s : t) := Raw.for_all f s.
- Definition for_all_1 (s : t) := Raw.for_all_1 (s:=s).
- Definition for_all_2 (s : t) := Raw.for_all_2 (s:=s).
+ Lemma 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.
+ Proof. exact (Raw.fold_1 s.(sorted)). Qed.
- Definition exists_ (f : elt -> bool) (s : t) := Raw.exists_ f s.
- Definition exists_1 (s : t) := Raw.exists_1 (s:=s).
- Definition exists_2 (s : t) := Raw.exists_2 (s:=s).
+ Lemma cardinal_1 : cardinal s = length (elements s).
+ Proof. exact (Raw.cardinal_1 s.(sorted)). Qed.
- Definition partition (f : elt -> bool) (s : t) :=
- let p := Raw.partition f s in
- (Build_slist (this:=fst p) (Raw.partition_sort_1 (sorted s) f),
- Build_slist (this:=snd p) (Raw.partition_sort_2 (sorted s) f)).
- Definition partition_1 (s : t) := Raw.partition_1 s.
- Definition partition_2 (s : t) := Raw.partition_2 s.
-
- Definition eq (s s' : t) := Raw.eq s s'.
- Definition eq_refl (s : t) := Raw.eq_refl s.
- Definition eq_sym (s s' : t) := Raw.eq_sym (s:=s) (s':=s').
- Definition eq_trans (s s' s'' : t) :=
- Raw.eq_trans (s:=s) (s':=s') (s'':=s'').
+ Section Filter.
- Definition lt (s s' : t) := Raw.lt s s'.
- Definition lt_trans (s s' s'' : t) :=
- Raw.lt_trans (s:=s) (s':=s') (s'':=s'').
- Definition lt_not_eq (s s' : t) := Raw.lt_not_eq (sorted s) (sorted s').
-
- Definition compare : forall s s' : t, Compare lt eq s s'.
- Proof.
- intros; elim (Raw.compare (sorted s) (sorted s'));
- [ constructor 1 | constructor 2 | constructor 3 ];
- auto.
- Defined.
+ Variable f : elt -> bool.
+
+ Lemma filter_1 : compat_bool E.eq f -> In x (filter f s) -> In x s.
+ Proof. exact (@Raw.filter_1 s x f). Qed.
+ Lemma filter_2 : compat_bool E.eq f -> In x (filter f s) -> f x = true.
+ Proof. exact (@Raw.filter_2 s x f). Qed.
+ Lemma filter_3 :
+ compat_bool E.eq f -> In x s -> f x = true -> In x (filter f s).
+ Proof. exact (@Raw.filter_3 s x f). 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 s f). 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 s f). 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 s f). 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 s f). Qed.
+
+ Lemma partition_1 :
+ compat_bool E.eq f -> Equal (fst (partition f s)) (filter f s).
+ Proof. exact (@Raw.partition_1 s f). Qed.
+ Lemma partition_2 :
+ compat_bool E.eq f ->
+ Equal (snd (partition f s)) (filter (fun x => negb (f x)) s).
+ Proof. exact (@Raw.partition_2 s f). Qed.
+
+ End Filter.
+
+ Lemma elements_1 : In x s -> InA E.eq x (elements s).
+ 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 : sort E.lt (elements s).
+ Proof. exact (Raw.elements_3 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.
+ Lemma min_elt_2 : min_elt s = Some x -> In y s -> ~ E.lt y x.
+ Proof. exact (fun H => Raw.min_elt_2 s.(sorted) H). Qed.
+ Lemma min_elt_3 : min_elt s = None -> Empty s.
+ Proof. exact (fun H => Raw.min_elt_3 H). Qed.
+
+ Lemma max_elt_1 : max_elt s = Some x -> In x s.
+ Proof. exact (fun H => Raw.max_elt_1 H). Qed.
+ Lemma max_elt_2 : max_elt s = Some x -> In y s -> ~ E.lt x y.
+ Proof. exact (fun H => Raw.max_elt_2 s.(sorted) H). Qed.
+ Lemma max_elt_3 : max_elt s = None -> Empty s.
+ Proof. exact (fun H => Raw.max_elt_3 H). Qed.
+
+ Lemma choose_1 : choose s = Some x -> In x s.
+ 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 eq_refl : eq s s.
+ Proof. exact (Raw.eq_refl s). Qed.
+ Lemma eq_sym : eq s s' -> eq s' s.
+ Proof. exact (@Raw.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.
+
+ Lemma lt_trans : lt s s' -> lt s' s'' -> lt s s''.
+ Proof. exact (@Raw.lt_trans s s' s''). Qed.
+ Lemma lt_not_eq : lt s s' -> ~ eq s s'.
+ Proof. exact (Raw.lt_not_eq s.(sorted) s'.(sorted)). Qed.
+
+ Definition compare : Compare lt eq s s'.
+ Proof.
+ elim (Raw.compare s.(sorted) s'.(sorted));
+ [ constructor 1 | constructor 2 | constructor 3 ];
+ auto.
+ Defined.
+
+ End Spec.
End Make.
diff --git a/theories/FSets/FSetProperties.v b/theories/FSets/FSetProperties.v
index 23843084..6e93a546 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 8639 2006-03-16 19:21:55Z letouzey $ *)
+(* $Id: FSetProperties.v 8853 2006-05-23 18:17:38Z herbelin $ *)
(** * Finite sets library *)
@@ -21,49 +21,13 @@ Require Import FSetFacts.
Set Implicit Arguments.
Unset Strict Implicit.
-Section Misc.
-Variable A B : Set.
-Variable eqA : A -> A -> Prop.
-Variable eqB : B -> B -> Prop.
-
-(** Two-argument functions that allow to reorder its arguments. *)
-Definition transpose (f : A -> B -> B) :=
- forall (x y : A) (z : B), eqB (f x (f y z)) (f y (f x z)).
-
-(** Compatibility of a two-argument function with respect to two equalities. *)
-Definition compat_op (f : A -> B -> B) :=
- forall (x x' : A) (y y' : B), eqA x x' -> eqB y y' -> eqB (f x y) (f x' y').
-
-(** Compatibility of a function upon natural numbers. *)
-Definition compat_nat (f : A -> nat) :=
- forall x x' : A, eqA x x' -> f x = f x'.
-
-End Misc.
Hint Unfold transpose compat_op compat_nat.
-
Hint Extern 1 (Setoid_Theory _ _) => constructor; congruence.
-Ltac trans_st x := match goal with
- | H : Setoid_Theory _ ?eqA |- ?eqA _ _ =>
- apply (Seq_trans _ _ H) with x; auto
- end.
-
-Ltac sym_st := match goal with
- | H : Setoid_Theory _ ?eqA |- ?eqA _ _ =>
- apply (Seq_sym _ _ H); auto
- end.
-
-Ltac refl_st := match goal with
- | H : Setoid_Theory _ ?eqA |- ?eqA _ _ =>
- apply (Seq_refl _ _ H); auto
- end.
-
-Definition gen_st : forall A : Set, Setoid_Theory _ (@eq A).
-Proof. auto. Qed.
-
Module Properties (M: S).
- Module ME := OrderedTypeFacts M.E.
- Import ME.
+ 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] *)
@@ -82,26 +46,29 @@ Module Properties (M: S).
Qed.
Section BasicProperties.
- Variable s s' s'' s1 s2 s3 : t.
- Variable x : elt.
(** properties of [Equal] *)
- Lemma equal_refl : s[=]s.
+ Lemma equal_refl : forall s, s[=]s.
Proof.
- apply eq_refl.
+ unfold Equal; intuition.
Qed.
- Lemma equal_sym : s[=]s' -> s'[=]s.
+ Lemma equal_sym : forall s s', s[=]s' -> s'[=]s.
Proof.
- apply eq_sym.
+ unfold Equal; intros.
+ rewrite H; intuition.
Qed.
- Lemma equal_trans : s1[=]s2 -> s2[=]s3 -> s1[=]s3.
+ Lemma equal_trans : forall s1 s2 s3, s1[=]s2 -> s2[=]s3 -> s1[=]s3.
Proof.
- intros; apply eq_trans with s2; auto.
+ 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.
@@ -154,6 +121,11 @@ Module Properties (M: S).
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] *)
@@ -174,6 +146,11 @@ Module Properties (M: S).
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] *)
@@ -185,7 +162,7 @@ Module Properties (M: S).
Lemma Equal_remove : s[=]s' -> remove x s [=] remove x s'.
Proof.
- intros; rewrite H; apply eq_refl.
+ intros; rewrite H; apply equal_refl.
Qed.
(** properties of [add] and [remove] *)
@@ -223,12 +200,12 @@ Module Properties (M: S).
Lemma union_equal_1 : s[=]s' -> union s s'' [=] union s' s''.
Proof.
- intros; rewrite H; apply eq_refl.
+ intros; rewrite H; apply equal_refl.
Qed.
Lemma union_equal_2 : s'[=]s'' -> union s s' [=] union s s''.
Proof.
- intros; rewrite H; apply eq_refl.
+ intros; rewrite H; apply equal_refl.
Qed.
Lemma union_assoc : union (union s s') s'' [=] union s (union s' s'').
@@ -261,6 +238,16 @@ Module Properties (M: S).
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.
@@ -290,12 +277,12 @@ Module Properties (M: S).
Lemma inter_equal_1 : s[=]s' -> inter s s'' [=] inter s' s''.
Proof.
- intros; rewrite H; apply eq_refl.
+ intros; rewrite H; apply equal_refl.
Qed.
Lemma inter_equal_2 : s'[=]s'' -> inter s s' [=] inter s s''.
Proof.
- intros; rewrite H; apply eq_refl.
+ intros; rewrite H; apply equal_refl.
Qed.
Lemma inter_assoc : inter (inter s s') s'' [=] inter s (inter s' s'').
@@ -447,140 +434,14 @@ Module Properties (M: S).
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 : set.
-
- Notation NoDup := (NoDupA E.eq).
- Notation EqList := (eqlistA E.eq).
-
- Section NoDupA_Remove.
-
- Let ListAdd x l l' := forall y : elt, ME.In y l' <-> E.eq x y \/ ME.In y l.
-
- Lemma removeA_add :
- forall s s' x x', NoDup s -> NoDup (x' :: s') ->
- ~ E.eq x x' -> ~ ME.In x s ->
- ListAdd x s (x' :: s') -> ListAdd x (removeA eq_dec x' s) s'.
- Proof.
- unfold ListAdd; intros.
- inversion_clear H0.
- rewrite removeA_InA; auto; [apply E.eq_trans|].
- split; intros.
- destruct (eq_dec x y); auto; intros.
- right; split; auto.
- destruct (H3 y); clear H3.
- destruct H6; intuition.
- swap H4; apply In_eq with y; auto.
- destruct H0.
- assert (ME.In y (x' :: s')) by rewrite H3; auto.
- inversion_clear H6; auto.
- elim H1; apply E.eq_trans with y; auto.
- destruct H0.
- assert (ME.In y (x' :: s')) by rewrite H3; auto.
- inversion_clear H7; auto.
- elim H6; auto.
- Qed.
-
- 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).
- Variables (i:A).
-
- Lemma removeA_fold_right_0 :
- forall s x, NoDup s -> ~ME.In x s ->
- eqA (fold_right f i s) (fold_right f i (removeA eq_dec x s)).
- Proof.
- simple induction s; simpl; intros.
- refl_st.
- inversion_clear H0.
- destruct (eq_dec x a); simpl; intros.
- absurd_hyp e; auto.
- apply Comp; auto.
- Qed.
-
- Lemma removeA_fold_right :
- forall s x, NoDup s -> ME.In x s ->
- eqA (fold_right f i s) (f x (fold_right f i (removeA eq_dec x s))).
- Proof.
- simple induction s; simpl.
- inversion_clear 2.
- intros.
- inversion_clear H0.
- destruct (eq_dec x a); simpl; intros.
- apply Comp; auto.
- apply removeA_fold_right_0; auto.
- swap H2; apply ME.In_eq with x; auto.
- inversion_clear H1.
- destruct n; auto.
- trans_st (f a (f x (fold_right f i (removeA eq_dec x l)))).
- Qed.
-
- Lemma fold_right_equal :
- forall s s', NoDup s -> NoDup s' ->
- EqList s s' -> eqA (fold_right f i s) (fold_right f i s').
- Proof.
- simple induction s.
- destruct s'; simpl.
- intros; refl_st; auto.
- unfold eqlistA; intros.
- destruct (H1 t0).
- assert (X : ME.In t0 nil); auto; inversion X.
- intros x l Hrec s' N N' E; simpl in *.
- trans_st (f x (fold_right f i (removeA eq_dec x s'))).
- apply Comp; auto.
- apply Hrec; auto.
- inversion N; auto.
- apply removeA_NoDupA; auto; apply E.eq_trans.
- apply removeA_eqlistA; auto; [apply E.eq_trans|].
- inversion_clear N; auto.
- sym_st.
- apply removeA_fold_right; auto.
- unfold eqlistA in E.
- rewrite <- E; auto.
- Qed.
-
- Lemma fold_right_add :
- forall s' s x, NoDup s -> NoDup s' -> ~ ME.In x s ->
- ListAdd x s s' -> eqA (fold_right f i s') (f x (fold_right f i s)).
- Proof.
- simple induction s'.
- unfold ListAdd; intros.
- destruct (H2 x); clear H2.
- assert (X : ME.In x nil); auto; inversion X.
- intros x' l' Hrec s x N N' IN EQ; simpl.
- (* if x=x' *)
- destruct (eq_dec x x').
- apply Comp; auto.
- apply fold_right_equal; auto.
- inversion_clear N'; trivial.
- unfold eqlistA; unfold ListAdd in EQ; intros.
- destruct (EQ x0); clear EQ.
- split; intros.
- destruct H; auto.
- inversion_clear N'.
- destruct H2; apply In_eq with x0; auto; order.
- assert (X:ME.In x0 (x' :: l')); auto; inversion_clear X; auto.
- destruct IN; apply In_eq with x0; auto; order.
- (* else x<>x' *)
- trans_st (f x' (f x (fold_right f i (removeA eq_dec x' s)))).
- apply Comp; auto.
- apply Hrec; auto.
- apply removeA_NoDupA; auto; apply E.eq_trans.
- inversion_clear N'; auto.
- rewrite removeA_InA; auto; [apply E.eq_trans|intuition].
- apply removeA_add; auto.
- trans_st (f x (f x' (fold_right f i (removeA eq_dec x' s)))).
- apply Comp; auto.
- sym_st.
- apply removeA_fold_right; auto.
- destruct (EQ x').
- destruct H; auto; destruct n; auto.
- Qed.
-
- End NoDupA_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:
*)
@@ -629,7 +490,9 @@ Module Properties (M: S).
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 := eqA); auto.
+ 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.
@@ -897,8 +760,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 (@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.
@@ -956,7 +819,23 @@ Module Properties (M: S).
rewrite (inter_subset_equal H); auto with arith.
Qed.
- Lemma union_inter_cardinal :
+ 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.
@@ -965,6 +844,15 @@ Module Properties (M: S).
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.
diff --git a/theories/FSets/FSetToFiniteSet.v b/theories/FSets/FSetToFiniteSet.v
new file mode 100644
index 00000000..8cf85efe
--- /dev/null
+++ b/theories/FSets/FSetToFiniteSet.v
@@ -0,0 +1,139 @@
+(***********************************************************************)
+(* 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: FSetToFiniteSet.v 8876 2006-05-30 13:43:15Z letouzey $ *)
+
+Require Import Ensembles Finite_sets.
+Require Import FSetInterface FSetProperties OrderedTypeEx.
+
+(** * Going from [FSets] with usual equality
+ to the old [Ensembles] and [Finite_sets] theory. *)
+
+Module S_to_Finite_set (U:UsualOrderedType)(M:S with Module E := U).
+ Module MP:= Properties(M).
+ Import M MP FM Ensembles Finite_sets.
+
+ Definition mkEns : M.t -> Ensemble M.elt :=
+ fun s x => M.In x s.
+
+ Notation " !! " := mkEns.
+
+ Lemma In_In : forall s x, M.In x s <-> In _ (!!s) x.
+ Proof.
+ unfold In; compute; auto.
+ Qed.
+
+ Lemma Subset_Included : forall s s', s[<=]s' <-> Included _ (!!s) (!!s').
+ Proof.
+ unfold Subset, Included, In, mkEns; intuition.
+ Qed.
+
+ Notation " a === b " := (Same_set M.elt a b) (at level 70, no associativity).
+
+ Lemma Equal_Same_set : forall s s', s[=]s' <-> !!s === !!s'.
+ Proof.
+ intros.
+ rewrite double_inclusion.
+ unfold Subset, Included, Same_set, In, mkEns; intuition.
+ Qed.
+
+ Lemma empty_Empty_Set : !!M.empty === Empty_set _.
+ Proof.
+ unfold Same_set, Included, mkEns, In.
+ split; intro; set_iff; inversion 1.
+ Qed.
+
+ Lemma Empty_Empty_set : forall s, Empty s -> !!s === Empty_set _.
+ Proof.
+ unfold Same_set, Included, mkEns, In.
+ split; intros.
+ destruct(H x H0).
+ inversion H0.
+ Qed.
+
+ Lemma singleton_Singleton : forall x, !!(M.singleton x) === Singleton _ x .
+ Proof.
+ unfold Same_set, Included, mkEns, In.
+ split; intro; set_iff; inversion 1; try constructor; auto.
+ Qed.
+
+ Lemma union_Union : forall s s', !!(union s s') === Union _ (!!s) (!!s').
+ Proof.
+ unfold Same_set, Included, mkEns, In.
+ split; intro; set_iff; inversion 1; [ constructor 1 | constructor 2 | | ]; auto.
+ Qed.
+
+ Lemma inter_Intersection : forall s s', !!(inter s s') === Intersection _ (!!s) (!!s').
+ Proof.
+ unfold Same_set, Included, mkEns, In.
+ split; intro; set_iff; inversion 1; try constructor; auto.
+ Qed.
+
+ Lemma add_Add : forall x s, !!(add x s) === Add _ (!!s) x.
+ Proof.
+ unfold Same_set, Included, mkEns, In.
+ split; intro; set_iff; inversion 1; unfold E.eq; auto with sets.
+ inversion H0.
+ constructor 2; constructor.
+ constructor 1; auto.
+ Qed.
+
+ Lemma Add_Add : forall x s s', MP.Add x s s' -> !!s' === Add _ (!!s) x.
+ Proof.
+ unfold Same_set, Included, mkEns, In.
+ split; intros.
+ red in H; rewrite H in H0.
+ destruct H0.
+ inversion H0.
+ constructor 2; constructor.
+ constructor 1; auto.
+ red in H; rewrite H; unfold E.eq in *.
+ inversion H0; auto.
+ inversion H1; auto.
+ Qed.
+
+ Lemma remove_Subtract : forall x s, !!(remove x s) === Subtract _ (!!s) x.
+ Proof.
+ unfold Same_set, Included, mkEns, In.
+ split; intro; set_iff; inversion 1; unfold E.eq in *; auto with sets.
+ split; auto.
+ swap H1.
+ inversion H2; auto.
+ Qed.
+
+ Lemma mkEns_Finite : forall s, Finite _ (!!s).
+ Proof.
+ intro s; pattern s; apply set_induction; clear s; intros.
+ intros; replace (!!s) with (Empty_set elt); auto with sets.
+ symmetry; apply Extensionality_Ensembles.
+ apply Empty_Empty_set; auto.
+ replace (!!s') with (Add _ (!!s) x).
+ constructor 2; auto.
+ symmetry; apply Extensionality_Ensembles.
+ apply Add_Add; auto.
+ Qed.
+
+ Lemma mkEns_cardinal : forall s, cardinal _ (!!s) (M.cardinal s).
+ Proof.
+ intro s; pattern s; apply set_induction; clear s; intros.
+ intros; replace (!!s) with (Empty_set elt); auto with sets.
+ rewrite cardinal_1; auto with sets.
+ symmetry; apply Extensionality_Ensembles.
+ apply Empty_Empty_set; auto.
+ replace (!!s') with (Add _ (!!s) x).
+ rewrite (cardinal_2 H0 H1); auto with sets.
+ symmetry; apply Extensionality_Ensembles.
+ apply Add_Add; auto.
+ Qed.
+
+End S_to_Finite_set.
diff --git a/theories/FSets/FSetWeak.v b/theories/FSets/FSetWeak.v
index 7ed61c9f..bfe34cd7 100644
--- a/theories/FSets/FSetWeak.v
+++ b/theories/FSets/FSetWeak.v
@@ -6,9 +6,11 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: FSetWeak.v 8641 2006-03-17 09:56:54Z letouzey $ *)
+(* $Id: FSetWeak.v 8819 2006-05-15 09:52:36Z letouzey $ *)
Require Export DecidableType.
+Require Export DecidableTypeEx.
Require Export FSetWeakInterface.
Require Export FSetFacts.
+Require Export FSetProperties.
Require Export FSetWeakList.
diff --git a/theories/FSets/FSetWeakFacts.v b/theories/FSets/FSetWeakFacts.v
index 46a73cc9..61797a95 100644
--- a/theories/FSets/FSetWeakFacts.v
+++ b/theories/FSets/FSetWeakFacts.v
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: FSetWeakFacts.v 8681 2006-04-05 11:56:14Z letouzey $ *)
+(* $Id: FSetWeakFacts.v 8882 2006-05-31 21:55:30Z letouzey $ *)
(** * Finite sets library *)
@@ -159,6 +159,12 @@ 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.
diff --git a/theories/FSets/FSetWeakInterface.v b/theories/FSets/FSetWeakInterface.v
index c1845494..a281ce22 100644
--- a/theories/FSets/FSetWeakInterface.v
+++ b/theories/FSets/FSetWeakInterface.v
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: FSetWeakInterface.v 8641 2006-03-17 09:56:54Z letouzey $ *)
+(* $Id: FSetWeakInterface.v 8820 2006-05-15 11:44:05Z letouzey $ *)
(** * Finite sets library *)
@@ -132,8 +132,8 @@ Module Type S.
Section Spec.
- Variable s s' s'' : t.
- Variable x y z : elt.
+ Variable s s' : t.
+ Variable x y : elt.
(** Specification of [In] *)
Parameter In_1 : E.eq x y -> In x s -> In y s.
@@ -226,15 +226,17 @@ Module Type S.
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 Filter.
End Spec.
Hint Immediate In_1.
@@ -243,6 +245,7 @@ Module Type S.
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.
+ 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 74c81f37..97080b7a 100644
--- a/theories/FSets/FSetWeakList.v
+++ b/theories/FSets/FSetWeakList.v
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: FSetWeakList.v 8639 2006-03-16 19:21:55Z letouzey $ *)
+(* $Id: FSetWeakList.v 8834 2006-05-20 00:41:35Z letouzey $ *)
(** * Finite sets library *)
@@ -114,7 +114,7 @@ Module Raw (X: DecidableType).
end.
(** ** Proofs of set operation specifications. *)
-
+ Section ForNotations.
Notation NoDup := (NoDupA X.eq).
Notation In := (InA X.eq).
@@ -750,6 +750,7 @@ Module Raw (X: DecidableType).
unfold eq, Equal; firstorder.
Qed.
+ End ForNotations.
End Raw.
(** * Encapsulation
@@ -759,115 +760,177 @@ End Raw.
Module Make (X: DecidableType) <: S with Module E := X.
- Module E := X.
Module Raw := Raw X.
+ Module E := X.
- Record slist : Set := {this :> Raw.t; unique : NoDupA X.eq this}.
+ Record slist : Set := {this :> Raw.t; unique : NoDupA E.eq this}.
Definition t := slist.
- Definition elt := X.t.
+ Definition elt := E.t.
- Definition In (x : elt) (s : t) := InA X.eq x s.(this).
- 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 : t) :=
+ Definition In (x : elt) (s : t) : Prop := InA E.eq x s.(this).
+ 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 : elt, In x s -> P x.
- Definition Exists (P : elt -> Prop) (s : t) := exists x : elt, In x s /\ P x.
-
- Definition In_1 (s : t) := Raw.In_eq (s:=s).
-
- Definition mem (x : elt) (s : t) := Raw.mem x s.
- Definition mem_1 (s : t) := Raw.mem_1 (s:=s).
- Definition mem_2 (s : t) := Raw.mem_2 (s:=s).
-
- Definition add x s := Build_slist (Raw.add_unique (unique s) x).
- Definition add_1 (s : t) := Raw.add_1 (unique s).
- Definition add_2 (s : t) := Raw.add_2 (unique s).
- Definition add_3 (s : t) := Raw.add_3 (unique s).
-
- Definition remove x s := Build_slist (Raw.remove_unique (unique s) x).
- Definition remove_1 (s : t) := Raw.remove_1 (unique s).
- Definition remove_2 (s : t) := Raw.remove_2 (unique s).
- Definition remove_3 (s : t) := Raw.remove_3 (unique s).
-
- Definition singleton x := Build_slist (Raw.singleton_unique x).
- Definition singleton_1 := Raw.singleton_1.
- Definition singleton_2 := Raw.singleton_2.
-
- Definition union (s s' : t) :=
- Build_slist (Raw.union_unique (unique s) (unique s')).
- Definition union_1 (s s' : t) := Raw.union_1 (unique s) (unique s').
- Definition union_2 (s s' : t) := Raw.union_2 (unique s) (unique s').
- Definition union_3 (s s' : t) := Raw.union_3 (unique s) (unique s').
+ Definition Exists (P : elt -> Prop) (s : t) : Prop := exists x : elt, In x s /\ P x.
- Definition inter (s s' : t) :=
+ Definition mem (x : elt) (s : t) : bool := Raw.mem x s.
+ Definition add (x : elt)(s : t) : t := Build_slist (Raw.add_unique (unique s) x).
+ Definition remove (x : elt)(s : t) : t := Build_slist (Raw.remove_unique (unique s) x).
+ Definition singleton (x : elt) : t := Build_slist (Raw.singleton_unique x).
+ Definition union (s s' : t) : t :=
+ Build_slist (Raw.union_unique (unique s) (unique s')).
+ Definition inter (s s' : t) : t :=
Build_slist (Raw.inter_unique (unique s) (unique s')).
- Definition inter_1 (s s' : t) := Raw.inter_1 (unique s) (unique s').
- Definition inter_2 (s s' : t) := Raw.inter_2 (unique s) (unique s').
- Definition inter_3 (s s' : t) := Raw.inter_3 (unique s) (unique s').
-
- Definition diff (s s' : t) :=
+ Definition diff (s s' : t) : t :=
Build_slist (Raw.diff_unique (unique s) (unique s')).
- Definition diff_1 (s s' : t) := Raw.diff_1 (unique s) (unique s').
- Definition diff_2 (s s' : t) := Raw.diff_2 (unique s) (unique s').
- Definition diff_3 (s s' : t) := Raw.diff_3 (unique s) (unique s').
-
- Definition equal (s s' : t) := Raw.equal s s'.
- Definition equal_1 (s s' : t) := Raw.equal_1 (unique s) (unique s').
- Definition equal_2 (s s' : t) := Raw.equal_2 (unique s) (unique s').
-
- Definition subset (s s' : t) := Raw.subset s s'.
- Definition subset_1 (s s' : t) := Raw.subset_1 (unique s) (unique s').
- Definition subset_2 (s s' : t) := Raw.subset_2 (unique s) (unique s').
+ Definition equal (s s' : t) : bool := Raw.equal s s'.
+ Definition subset (s s' : t) : bool := Raw.subset s s'.
+ Definition empty : t := Build_slist Raw.empty_unique.
+ 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 cardinal (s : t) : nat := Raw.cardinal s.
+ Definition filter (f : elt -> bool) (s : t) : t :=
+ Build_slist (Raw.filter_unique (unique s) f).
+ 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
+ (Build_slist (this:=fst p) (Raw.partition_unique_1 (unique s) f),
+ Build_slist (this:=snd p) (Raw.partition_unique_2 (unique s) f)).
+
+ Section Spec.
+ Variable s s' : t.
+ Variable x y : elt.
- Definition empty := Build_slist Raw.empty_unique.
- Definition empty_1 := Raw.empty_1.
+ Lemma In_1 : E.eq x y -> In x s -> In y s.
+ Proof. exact (fun H H' => Raw.In_eq H H'). Qed.
- Definition is_empty (s : t) := Raw.is_empty s.
- Definition is_empty_1 (s : t) := Raw.is_empty_1 (s:=s).
- Definition is_empty_2 (s : t) := Raw.is_empty_2 (s:=s).
-
- Definition elements (s : t) := Raw.elements s.
- Definition elements_1 (s : t) := Raw.elements_1 (s:=s).
- Definition elements_2 (s : t) := Raw.elements_2 (s:=s).
- Definition elements_3 (s : t) := Raw.elements_3 (unique s).
-
- Definition choose (s:t) := Raw.choose s.
- Definition choose_1 (s : t) := Raw.choose_1 (s:=s).
- Definition choose_2 (s : t) := Raw.choose_2 (s:=s).
-
- Definition fold (B : Set) (f : elt -> B -> B) (s : t) := Raw.fold (B:=B) f s.
- Definition fold_1 (s : t) := Raw.fold_1 (unique s).
+ Lemma mem_1 : In x s -> mem x s = true.
+ Proof. exact (fun H => Raw.mem_1 H). Qed.
+ Lemma mem_2 : mem x s = true -> In x s.
+ Proof. exact (fun H => Raw.mem_2 H). Qed.
- Definition cardinal (s : t) := Raw.cardinal s.
- Definition cardinal_1 (s : t) := Raw.cardinal_1 (unique s).
+ Lemma equal_1 : Equal s s' -> equal s s' = true.
+ Proof. exact (Raw.equal_1 s.(unique) s'.(unique)). Qed.
+ Lemma equal_2 : equal s s' = true -> Equal s s'.
+ Proof. exact (Raw.equal_2 s.(unique) s'.(unique)). Qed.
+
+ Lemma subset_1 : Subset s s' -> subset s s' = true.
+ Proof. exact (Raw.subset_1 s.(unique) s'.(unique)). Qed.
+ Lemma subset_2 : subset s s' = true -> Subset s s'.
+ Proof. exact (Raw.subset_2 s.(unique) s'.(unique)). Qed.
+
+ Lemma empty_1 : Empty empty.
+ Proof. exact Raw.empty_1. Qed.
+
+ Lemma is_empty_1 : Empty s -> is_empty s = true.
+ Proof. exact (fun H => Raw.is_empty_1 H). Qed.
+ Lemma is_empty_2 : is_empty s = true -> Empty s.
+ Proof. exact (fun H => Raw.is_empty_2 H). Qed.
- Definition filter (f : elt -> bool) (s : t) :=
- Build_slist (Raw.filter_unique (unique s) f).
- Definition filter_1 (s : t)(x:elt)(f: elt -> bool)(H:compat_bool X.eq f) :=
- @Raw.filter_1 s x f.
- Definition filter_2 (s : t) := Raw.filter_2 (s:=s).
- Definition filter_3 (s : t) := Raw.filter_3 (s:=s).
+ Lemma add_1 : E.eq x y -> In y (add x s).
+ Proof. exact (fun H => Raw.add_1 s.(unique) H). Qed.
+ Lemma add_2 : In y s -> In y (add x s).
+ Proof. exact (fun H => Raw.add_2 s.(unique) x H). Qed.
+ Lemma add_3 : ~ E.eq x y -> In y (add x s) -> In y s.
+ Proof. exact (fun H => Raw.add_3 s.(unique) H). Qed.
+
+ Lemma remove_1 : E.eq x y -> ~ In y (remove x s).
+ Proof. exact (fun H => Raw.remove_1 s.(unique) H). Qed.
+ Lemma remove_2 : ~ E.eq x y -> In y s -> In y (remove x s).
+ Proof. exact (fun H H' => Raw.remove_2 s.(unique) H H'). Qed.
+ Lemma remove_3 : In y (remove x s) -> In y s.
+ Proof. exact (fun H => Raw.remove_3 s.(unique) H). Qed.
+
+ Lemma singleton_1 : In y (singleton x) -> E.eq x y.
+ Proof. exact (fun H => Raw.singleton_1 H). Qed.
+ Lemma singleton_2 : E.eq x y -> In y (singleton x).
+ Proof. exact (fun H => Raw.singleton_2 H). Qed.
+
+ Lemma union_1 : In x (union s s') -> In x s \/ In x s'.
+ Proof. exact (fun H => Raw.union_1 s.(unique) s'.(unique) H). Qed.
+ Lemma union_2 : In x s -> In x (union s s').
+ Proof. exact (fun H => Raw.union_2 s.(unique) s'.(unique) H). Qed.
+ Lemma union_3 : In x s' -> In x (union s s').
+ Proof. exact (fun H => Raw.union_3 s.(unique) s'.(unique) H). Qed.
+
+ Lemma inter_1 : In x (inter s s') -> In x s.
+ Proof. exact (fun H => Raw.inter_1 s.(unique) s'.(unique) H). Qed.
+ Lemma inter_2 : In x (inter s s') -> In x s'.
+ Proof. exact (fun H => Raw.inter_2 s.(unique) s'.(unique) H). Qed.
+ Lemma inter_3 : In x s -> In x s' -> In x (inter s s').
+ Proof. exact (fun H => Raw.inter_3 s.(unique) s'.(unique) H). Qed.
+
+ Lemma diff_1 : In x (diff s s') -> In x s.
+ Proof. exact (fun H => Raw.diff_1 s.(unique) s'.(unique) H). Qed.
+ Lemma diff_2 : In x (diff s s') -> ~ In x s'.
+ Proof. exact (fun H => Raw.diff_2 s.(unique) s'.(unique) H). Qed.
+ 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.
- Definition for_all (f : elt -> bool) (s : t) := Raw.for_all f s.
- Definition for_all_1 (s : t) := Raw.for_all_1 (s:=s).
- Definition for_all_2 (s : t) := Raw.for_all_2 (s:=s).
+ Lemma 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.
+ Proof. exact (Raw.fold_1 s.(unique)). Qed.
- Definition exists_ (f : elt -> bool) (s : t) := Raw.exists_ f s.
- Definition exists_1 (s : t) := Raw.exists_1 (s:=s).
- Definition exists_2 (s : t) := Raw.exists_2 (s:=s).
+ Lemma cardinal_1 : cardinal s = length (elements s).
+ Proof. exact (Raw.cardinal_1 s.(unique)). Qed.
- Definition partition (f : elt -> bool) (s : t) :=
- let p := Raw.partition f s in
- (Build_slist (this:=fst p) (Raw.partition_unique_1 (unique s) f),
- Build_slist (this:=snd p) (Raw.partition_unique_2 (unique s) f)).
- Definition partition_1 (s : t) := Raw.partition_1 s.
- Definition partition_2 (s : t) := Raw.partition_2 s.
-
- Definition eq (s s' : t) := Raw.eq s s'.
- Definition eq_refl (s : t) := Raw.eq_refl s.
- Definition eq_sym (s s' : t) := Raw.eq_sym (s:=s) (s':=s').
- Definition eq_trans (s s' s'' : t) :=
- Raw.eq_trans (s:=s) (s':=s') (s'':=s'').
+ Section Filter.
+
+ Variable f : elt -> bool.
+
+ Lemma filter_1 : compat_bool E.eq f -> In x (filter f s) -> In x s.
+ Proof. exact (fun H => @Raw.filter_1 s x f). Qed.
+ Lemma filter_2 : compat_bool E.eq f -> In x (filter f s) -> f x = true.
+ Proof. exact (@Raw.filter_2 s x f). Qed.
+ Lemma filter_3 :
+ compat_bool E.eq f -> In x s -> f x = true -> In x (filter f s).
+ Proof. exact (@Raw.filter_3 s x f). 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 s f). 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 s f). 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 s f). 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 s f). Qed.
+
+ Lemma partition_1 :
+ compat_bool E.eq f -> Equal (fst (partition f s)) (filter f s).
+ Proof. exact (@Raw.partition_1 s f). Qed.
+ Lemma partition_2 :
+ compat_bool E.eq f ->
+ Equal (snd (partition f s)) (filter (fun x => negb (f x)) s).
+ Proof. exact (@Raw.partition_2 s f). Qed.
+
+ End Filter.
+
+ Lemma elements_1 : In x s -> InA E.eq x (elements s).
+ 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 choose_1 : choose s = Some x -> In x s.
+ 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.
+
+ End Spec.
End Make.
diff --git a/theories/FSets/FSetWeakProperties.v b/theories/FSets/FSetWeakProperties.v
new file mode 100644
index 00000000..a0054d36
--- /dev/null
+++ b/theories/FSets/FSetWeakProperties.v
@@ -0,0 +1,896 @@
+(***********************************************************************)
+(* 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 9dfcd51f..b0402db6 100644
--- a/theories/FSets/FSets.v
+++ b/theories/FSets/FSets.v
@@ -6,9 +6,11 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: FSets.v 8667 2006-03-28 11:59:44Z letouzey $ *)
+(* $Id: FSets.v 8897 2006-06-05 21:04:10Z letouzey $ *)
Require Export OrderedType.
+Require Export OrderedTypeEx.
+Require Export OrderedTypeAlt.
Require Export FSetInterface.
Require Export FSetBridge.
Require Export FSetProperties.
diff --git a/theories/FSets/OrderedType.v b/theories/FSets/OrderedType.v
index 2bf08dc7..f966cd4d 100644
--- a/theories/FSets/OrderedType.v
+++ b/theories/FSets/OrderedType.v
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(***********************************************************************)
-(* $Id: OrderedType.v 8667 2006-03-28 11:59:44Z letouzey $ *)
+(* $Id: OrderedType.v 8834 2006-05-20 00:41:35Z letouzey $ *)
Require Export SetoidList.
Set Implicit Arguments.
@@ -313,6 +313,8 @@ Ltac false_order := elimtype False; order.
(* Specialization of resuts about lists modulo. *)
+Section ForNotations.
+
Notation In:=(InA eq).
Notation Inf:=(lelistA lt).
Notation Sort:=(sort lt).
@@ -346,12 +348,14 @@ Proof. exact (InfA_alt eq_refl eq_sym lt_trans lt_eq eq_lt). Qed.
Lemma Sort_NoDup : forall l, Sort l -> NoDup l.
Proof. exact (SortA_NoDupA eq_refl eq_sym lt_trans lt_not_eq lt_eq eq_lt) . Qed.
+End ForNotations.
+
Hint Resolve ListIn_In Sort_NoDup Inf_lt.
Hint Immediate In_eq Inf_lt.
End OrderedTypeFacts.
-Module PairOrderedType(O:OrderedType).
+Module KeyOrderedType(O:OrderedType).
Import O.
Module MO:=OrderedTypeFacts(O).
Import MO.
@@ -561,6 +565,6 @@ Module PairOrderedType(O:OrderedType).
Hint Resolve Sort_Inf_NotIn.
Hint Resolve In_inv_2 In_inv_3.
-End PairOrderedType.
+End KeyOrderedType.
diff --git a/theories/FSets/OrderedTypeAlt.v b/theories/FSets/OrderedTypeAlt.v
new file mode 100644
index 00000000..9bcfbfc7
--- /dev/null
+++ b/theories/FSets/OrderedTypeAlt.v
@@ -0,0 +1,129 @@
+(***********************************************************************)
+(* 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: OrderedTypeAlt.v 8773 2006-04-29 14:31:32Z letouzey $ *)
+
+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 _ ]
+*)
+
+Module Type OrderedTypeAlt.
+
+ Parameter t : Set.
+
+ Parameter compare : t -> t -> comparison.
+
+ Infix "?=" := compare (at level 70, no associativity).
+
+ Parameter compare_sym :
+ forall x y, (y?=x) = CompOpp (x?=y).
+ Parameter compare_trans :
+ forall c x y z, (x?=y) = c -> (y?=z) = c -> (x?=z) = c.
+
+End OrderedTypeAlt.
+
+(** From this new presentation to the original one. *)
+
+Module OrderedType_from_Alt (O:OrderedTypeAlt) <: OrderedType.
+ Import O.
+
+ Definition t := t.
+
+ Definition eq x y := (x?=y) = Eq.
+ Definition lt x y := (x?=y) = Lt.
+
+ Lemma eq_refl : forall x, eq x x.
+ Proof.
+ intro x.
+ unfold eq.
+ assert (H:=compare_sym x x).
+ destruct (x ?= x); simpl in *; try discriminate; auto.
+ Qed.
+
+ Lemma eq_sym : forall x y, eq x y -> eq y x.
+ Proof.
+ unfold eq; intros.
+ rewrite compare_sym.
+ rewrite H; simpl; auto.
+ Qed.
+
+ Definition eq_trans := (compare_trans Eq).
+
+ Definition lt_trans := (compare_trans Lt).
+
+ Lemma lt_not_eq : forall x y, lt x y -> ~eq x y.
+ Proof.
+ unfold eq, lt; intros.
+ rewrite H; discriminate.
+ Qed.
+
+ Definition compare : forall x y, Compare lt eq x y.
+ Proof.
+ intros.
+ case_eq (x ?= y); intros.
+ apply EQ; auto.
+ apply LT; auto.
+ apply GT; red.
+ rewrite compare_sym; rewrite H; auto.
+ Defined.
+
+End OrderedType_from_Alt.
+
+(** From the original presentation to this alternative one. *)
+
+Module OrderedType_to_Alt (O:OrderedType) <: OrderedTypeAlt.
+ Import O.
+ Module MO:=OrderedTypeFacts(O).
+ Import MO.
+
+ Definition t := t.
+
+ Definition compare x y := match compare x y with
+ | LT _ => Lt
+ | EQ _ => Eq
+ | GT _ => Gt
+ end.
+
+ Infix "?=" := compare (at level 70, no associativity).
+
+ 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.
+ 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.
+ elim_comp; auto.
+ Qed.
+
+End OrderedType_to_Alt.
+
+
diff --git a/theories/FSets/OrderedTypeEx.v b/theories/FSets/OrderedTypeEx.v
new file mode 100644
index 00000000..1c5a4054
--- /dev/null
+++ b/theories/FSets/OrderedTypeEx.v
@@ -0,0 +1,248 @@
+(***********************************************************************)
+(* 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: OrderedTypeEx.v 8836 2006-05-20 21:34:27Z letouzey $ *)
+
+Require Import OrderedType.
+Require Import ZArith.
+Require Import Omega.
+Require Import NArith Ndec.
+Require Import Compare_dec.
+
+(** * Examples of Ordered Type structures. *)
+
+(** First, a particular case of [OrderedType] where
+ the equality is the usual one of Coq. *)
+
+Module Type UsualOrderedType.
+ Parameter t : Set.
+ Definition eq := @eq t.
+ Parameter lt : t -> t -> Prop.
+ Definition eq_refl := @refl_equal t.
+ Definition eq_sym := @sym_eq t.
+ Definition eq_trans := @trans_eq t.
+ Axiom lt_trans : forall x y z : t, lt x y -> lt y z -> lt x z.
+ Axiom lt_not_eq : forall x y : t, lt x y -> ~ eq x y.
+ Parameter compare : forall x y : t, Compare lt eq x y.
+End UsualOrderedType.
+
+(** a [UsualOrderedType] is in particular an [OrderedType]. *)
+
+Module UOT_to_OT (U:UsualOrderedType) <: OrderedType := U.
+
+(** [nat] is an ordered type with respect to the usual order on natural numbers. *)
+
+Module Nat_as_OT <: UsualOrderedType.
+
+ Definition t := nat.
+
+ Definition eq := @eq nat.
+ Definition eq_refl := @refl_equal t.
+ Definition eq_sym := @sym_eq t.
+ Definition eq_trans := @trans_eq t.
+
+ Definition lt := lt.
+
+ Lemma lt_trans : forall x y z : t, lt x y -> lt y z -> lt x z.
+ Proof. unfold lt in |- *; intros; apply lt_trans with y; auto. Qed.
+
+ Lemma lt_not_eq : forall x y : t, lt x y -> ~ eq x y.
+ Proof. unfold lt, eq in |- *; intros; omega. Qed.
+
+ Definition compare : forall x y : t, Compare lt eq x y.
+ Proof.
+ intros; case (lt_eq_lt_dec x y).
+ simple destruct 1; intro.
+ constructor 1; auto.
+ constructor 2; auto.
+ intro; constructor 3; auto.
+ Qed.
+
+End Nat_as_OT.
+
+
+(** [Z] is an ordered type with respect to the usual order on integers. *)
+
+Open Scope Z_scope.
+
+Module Z_as_OT <: UsualOrderedType.
+
+ Definition t := Z.
+ Definition eq := @eq Z.
+ Definition eq_refl := @refl_equal t.
+ Definition eq_sym := @sym_eq t.
+ Definition eq_trans := @trans_eq t.
+
+ Definition lt (x y:Z) := (x<y).
+
+ Lemma lt_trans : forall x y z, x<y -> y<z -> x<z.
+ Proof. intros; omega. Qed.
+
+ Lemma lt_not_eq : forall x y, x<y -> ~ x=y.
+ Proof. intros; omega. Qed.
+
+ Definition compare : forall x y, Compare lt eq x y.
+ Proof.
+ intros x y; case_eq (x ?= y); intros.
+ apply EQ; unfold eq; apply Zcompare_Eq_eq; auto.
+ apply LT; unfold lt, Zlt; auto.
+ apply GT; unfold lt, Zlt; rewrite <- Zcompare_Gt_Lt_antisym; auto.
+ Defined.
+
+End Z_as_OT.
+
+(** [positive] is an ordered type with respect to the usual order on natural numbers. *)
+
+Open Scope positive_scope.
+
+Module Positive_as_OT <: 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.
+
+ Definition lt p q:= (p ?= q) Eq = Lt.
+
+ Lemma lt_trans : forall x y z : t, lt x y -> lt y z -> lt x z.
+ Proof.
+ unfold lt; intros x y z.
+ change ((Zpos x < Zpos y)%Z -> (Zpos y < Zpos z)%Z -> (Zpos x < Zpos z)%Z).
+ omega.
+ Qed.
+
+ 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 Pcompare_refl in H; discriminate.
+ Qed.
+
+ Definition compare : forall x y : t, Compare lt eq x y.
+ Proof.
+ intros x y.
+ case_eq ((x ?= y) Eq); intros.
+ apply EQ; apply Pcompare_Eq_eq; auto.
+ apply LT; unfold lt; auto.
+ apply GT; unfold lt.
+ replace Eq with (CompOpp Eq); auto.
+ rewrite <- Pcompare_antisym; rewrite H; auto.
+ Qed.
+
+End Positive_as_OT.
+
+
+(** [N] is an ordered type with respect to the usual order on natural numbers. *)
+
+Open Scope positive_scope.
+
+Module N_as_OT <: 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:= Nle q p = false.
+
+ Definition lt_trans := Nlt_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.
+ Qed.
+
+ Definition compare : forall x y : t, Compare lt eq x y.
+ Proof.
+ intros x y.
+ 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.
+ destruct (x ?= y)%N; simpl; try discriminate.
+ intros (H0,_); elim H0; auto.
+ 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.
+ Qed.
+
+End N_as_OT.
+
+
+(** From two ordered types, we can build a new OrderedType
+ over their cartesian product, using the lexicographic order. *)
+
+Module PairOrderedType(O1 O2:OrderedType) <: OrderedType.
+ Module MO1:=OrderedTypeFacts(O1).
+ Module MO2:=OrderedTypeFacts(O2).
+
+ Definition t := prod O1.t O2.t.
+
+ Definition eq x y := O1.eq (fst x) (fst y) /\ O2.eq (snd x) (snd y).
+
+ Definition lt x y :=
+ O1.lt (fst x) (fst y) \/
+ (O1.eq (fst x) (fst y) /\ O2.lt (snd x) (snd y)).
+
+ Lemma eq_refl : forall x : t, eq x x.
+ Proof.
+ intros (x1,x2); red; simpl; auto.
+ Qed.
+
+ Lemma eq_sym : forall x y : t, eq x y -> eq y x.
+ Proof.
+ intros (x1,x2) (y1,y2); unfold eq; simpl; intuition.
+ Qed.
+
+ Lemma eq_trans : forall x y z : t, eq x y -> eq y z -> eq x z.
+ Proof.
+ intros (x1,x2) (y1,y2) (z1,z2); unfold eq; simpl; intuition eauto.
+ Qed.
+
+ Lemma lt_trans : forall x y z : t, lt x y -> lt y z -> lt x z.
+ Proof.
+ intros (x1,x2) (y1,y2) (z1,z2); unfold eq, lt; simpl; intuition.
+ left; eauto.
+ left; eapply MO1.lt_eq; eauto.
+ left; eapply MO1.eq_lt; eauto.
+ right; split; eauto.
+ Qed.
+
+ Lemma lt_not_eq : forall x y : t, lt x y -> ~ eq x y.
+ Proof.
+ intros (x1,x2) (y1,y2); unfold eq, lt; simpl; intuition.
+ apply (O1.lt_not_eq H0 H1).
+ apply (O2.lt_not_eq H3 H2).
+ Qed.
+
+ Definition compare : forall x y : t, Compare lt eq x y.
+ intros (x1,x2) (y1,y2).
+ destruct (O1.compare x1 y1).
+ apply LT; unfold lt; auto.
+ destruct (O2.compare x2 y2).
+ apply LT; unfold lt; auto.
+ apply EQ; unfold eq; auto.
+ apply GT; unfold lt; auto.
+ apply GT; unfold lt; auto.
+ Qed.
+
+End PairOrderedType.
+